├── Category.agda ├── Injections.agda ├── Injections ├── Sum.agda └── Type.agda ├── MetaRens.agda ├── NOTERM ├── Unification.agda └── Unification │ ├── Inversion.agda │ ├── Pruning.agda │ └── Specification.agda ├── README.md ├── Support ├── EqReasoning.agda ├── Equality.agda ├── List.agda ├── Nat.agda └── Product.agda ├── Syntax.agda ├── Syntax ├── Equality.agda ├── NbE.agda ├── NbE │ └── Properties.agda ├── No-Cycle.agda ├── OneHoleContext.agda ├── RenOrn.agda ├── Sub.agda └── Type.agda ├── Unification.agda ├── Unification ├── Injections.agda ├── Inversion.agda ├── MetaRens.agda ├── OccursCheck.agda ├── Pruning.agda ├── Pruning │ └── Epi-Decr.agda ├── Specification.agda └── Specification │ └── Decr-Sub.agda ├── Vars.agda └── Vars ├── MatchTwo.agda └── SumIso.agda /Category.agda: -------------------------------------------------------------------------------- 1 | module Category 2 | (Obj : Set) 3 | (_⇒_ : Obj -> Obj -> Set) 4 | (_∘_ : ∀ {a b c} -> b ⇒ c -> a ⇒ b -> a ⇒ c) 5 | (id : ∀ {a} -> a ⇒ a) 6 | (_==_ : ∀ {A B} (f g : A ⇒ B) -> Set) 7 | where 8 | 9 | private 10 | infix 4 _≡_ 11 | _≡_ : ∀ {A B} (f g : A ⇒ B) -> Set 12 | _≡_ = _==_ 13 | 14 | Monic : ∀ {A B} (f : A ⇒ B) -> Set 15 | Monic {A} {B} f = ∀ {C} {g1 g2 : C ⇒ A} -> (f ∘ g1) ≡ (f ∘ g2) -> g1 ≡ g2 16 | 17 | record IsPullback {X Y Z}(f : X ⇒ Z)(g : Y ⇒ Z)(P : Obj)(p₁ : P ⇒ X)(p₂ : P ⇒ Y) : Set where 18 | field 19 | commutes : f ∘ p₁ ≡ g ∘ p₂ 20 | 21 | universal : ∀ {Q}(q₁ : Q ⇒ X)(q₂ : Q ⇒ Y) 22 | → (commutes : f ∘ q₁ ≡ g ∘ q₂) → (Q ⇒ P) 23 | 24 | universal-unique : ∀ {Q} {q₁ : Q ⇒ X} {q₂ : Q ⇒ Y} 25 | {commutes : f ∘ q₁ ≡ g ∘ q₂} 26 | (u : Q ⇒ P) 27 | (p₁∘u≡q₁ : p₁ ∘ u ≡ q₁) 28 | (p₂∘u≡q₂ : p₂ ∘ u ≡ q₂) 29 | → u ≡ universal q₁ q₂ commutes 30 | 31 | p₁∘universal≡q₁ : ∀ {Q} {q₁ : Q ⇒ X} {q₂ : Q ⇒ Y} 32 | {commutes : f ∘ q₁ ≡ g ∘ q₂} 33 | → (p₁ ∘ universal q₁ q₂ commutes ≡ q₁) 34 | 35 | p₂∘universal≡q₂ : ∀ {Q} {q₁ : Q ⇒ X} {q₂ : Q ⇒ Y} 36 | {commutes : f ∘ q₁ ≡ g ∘ q₂} 37 | → (p₂ ∘ universal q₁ q₂ commutes ≡ q₂) 38 | 39 | 40 | record Pullback {X Y Z}(f : X ⇒ Z)(g : Y ⇒ Z) : Set where 41 | constructor Pull_,_,_,_ 42 | field 43 | P : Obj 44 | p₁ : P ⇒ X 45 | p₂ : P ⇒ Y 46 | isPullback : IsPullback f g P p₁ p₂ 47 | 48 | open IsPullback isPullback public 49 | 50 | 51 | record IsEqualizer {X Y}(f g : X ⇒ Y) (E : Obj) (e : E ⇒ X) : Set where 52 | field 53 | commutes : f ∘ e ≡ g ∘ e 54 | 55 | universal : ∀ {Q}(m : Q ⇒ X) 56 | → (commutes : f ∘ m ≡ g ∘ m) → (Q ⇒ E) 57 | 58 | universal-unique : ∀ {Q} {m : Q ⇒ X} 59 | {commutes : f ∘ m ≡ g ∘ m} 60 | (u : Q ⇒ E) 61 | (e∘u≡m : e ∘ u ≡ m) 62 | → u ≡ universal m commutes 63 | 64 | e∘universal≡m : ∀ {Q} {m : Q ⇒ X} 65 | {commutes : f ∘ m ≡ g ∘ m} 66 | → (e ∘ universal m commutes ≡ m) 67 | 68 | record Equalizer {X Y}(f g : X ⇒ Y) : Set where 69 | constructor Equ_,_,_ 70 | field 71 | E : Obj 72 | e : E ⇒ X 73 | isEqualizer : IsEqualizer f g E e 74 | 75 | open IsEqualizer isEqualizer public 76 | 77 | 78 | record Product (A B : Obj) : Set where 79 | constructor Prod 80 | field 81 | A×B : Obj 82 | π₁ : A×B ⇒ A 83 | π₂ : A×B ⇒ B 84 | ⟨_,_⟩ : ∀ {C} → (C ⇒ A) → (C ⇒ B) → (C ⇒ A×B) 85 | 86 | commute₁ : ∀ {C} {f : C ⇒ A} {g : C ⇒ B} → π₁ ∘ ⟨ f , g ⟩ ≡ f 87 | commute₂ : ∀ {C} {f : C ⇒ A} {g : C ⇒ B} → π₂ ∘ ⟨ f , g ⟩ ≡ g 88 | universal : ∀ {C} {f : C ⇒ A} {g : C ⇒ B} {i : C ⇒ A×B} 89 | → π₁ ∘ i ≡ f → π₂ ∘ i ≡ g → ⟨ f , g ⟩ ≡ i 90 | 91 | open import Relation.Binary 92 | module Props (assoc : ∀ {A B C D} {f : A ⇒ B} {g : B ⇒ C} {h : C ⇒ D} → (h ∘ g) ∘ f ≡ h ∘ (g ∘ f)) 93 | (identityˡ : ∀ {A B} {f : A ⇒ B} → id ∘ f ≡ f) 94 | (identityʳ : ∀ {A B} {f : A ⇒ B} → f ∘ id ≡ f) 95 | 96 | (equiv : ∀ {A B} → IsEquivalence (_≡_ {A} {B})) 97 | (∘-resp-≡ : ∀ {A B C} {f h : B ⇒ C} {g i : A ⇒ B} → f ≡ h → g ≡ i → f ∘ g ≡ h ∘ i) 98 | where 99 | module Dummy {A}{B} = IsEquivalence (equiv {A} {B}) 100 | open Dummy 101 | 102 | convert : {X Y Z : Obj} -> (f : X ⇒ Z)(g : Y ⇒ Z) -> (prod : Product X Y) -> let open Product prod in 103 | Equalizer (f ∘ π₁) (g ∘ π₂) -> Pullback f g 104 | convert f g prod equ = Pull E , (π₁ ∘ e) , (π₂ ∘ e) , 105 | (record { 106 | commutes = trans (sym assoc) (trans commutes assoc); 107 | universal = λ q₁ q₂ commutes₁ → universal ⟨ q₁ , q₂ ⟩ (trans assoc (trans (∘-resp-≡ refl commute₁) 108 | (trans commutes₁ (trans (sym (∘-resp-≡ refl commute₂)) (sym assoc))))); 109 | universal-unique = λ u p₁∘u≡q₁ p₂∘u≡q₂ → universal-unique u (sym (uniπ (trans (sym assoc) p₁∘u≡q₁) (trans (sym assoc) p₂∘u≡q₂))); 110 | p₁∘universal≡q₁ = trans assoc (trans (∘-resp-≡ refl e∘universal≡m) commute₁); 111 | p₂∘universal≡q₂ = trans assoc (trans (∘-resp-≡ refl e∘universal≡m) commute₂) }) 112 | where 113 | open Product prod renaming (universal to uniπ) 114 | open Equalizer equ 115 | 116 | Equalizer-ext : ∀ {X Z}{f1 f2 : X ⇒ Z} -> f1 ≡ f2 -> {g1 g2 : X ⇒ Z} -> g1 ≡ g2 -> Equalizer f1 g1 -> Equalizer f2 g2 117 | Equalizer-ext f1≡f2 g1≡g2 pull 118 | = Equ E , e , 119 | (record { 120 | commutes = trans (∘-resp-≡ (sym f1≡f2) refl) (trans commutes (∘-resp-≡ g1≡g2 refl)); 121 | universal = λ m commutes₁ → universal m (trans (∘-resp-≡ f1≡f2 refl) 122 | (trans commutes₁ (∘-resp-≡ (sym g1≡g2) refl))); 123 | universal-unique = universal-unique; 124 | e∘universal≡m = e∘universal≡m}) 125 | where 126 | open Equalizer pull 127 | 128 | under-assoc : ∀ {A B C D} {f : A ⇒ B} {g : B ⇒ C} {h : C ⇒ D} -> 129 | ∀ {C} {g1 : B ⇒ C} {h1 : C ⇒ D} -> h ∘ g ≡ h1 ∘ g1 -> h ∘ (g ∘ f) ≡ h1 ∘ (g1 ∘ f) 130 | under-assoc eq = trans (sym assoc) (trans (∘-resp-≡ eq refl) assoc) 131 | 132 | mono-pullback-stable : ∀ {X Y Z : Obj} -> (f : X ⇒ Z)(g : Y ⇒ Z) -> (pull : Pullback f g) -> Monic g -> Monic (Pullback.p₁ pull) 133 | mono-pullback-stable f g pull g-mono {C} {g1} {g2} p₁∘g1≡p₁∘g2 = 134 | trans 135 | (universal-unique {commutes = under-assoc commutes} g1 refl refl) 136 | (sym 137 | (universal-unique g2 (sym p₁∘g1≡p₁∘g2) 138 | (g-mono 139 | (trans (under-assoc (sym commutes)) 140 | (trans (∘-resp-≡ refl (sym p₁∘g1≡p₁∘g2)) 141 | (under-assoc commutes)))))) 142 | where 143 | open Pullback pull 144 | 145 | _∘mono_ : ∀ {A B C} {f : B ⇒ C}{g : A ⇒ B} -> Monic f -> Monic g -> Monic (f ∘ g) 146 | f-mono ∘mono g-mono = λ x → g-mono (f-mono (trans (sym assoc) (trans x assoc))) 147 | -------------------------------------------------------------------------------- /Injections.agda: -------------------------------------------------------------------------------- 1 | module Injections where 2 | 3 | open import Data.Product 4 | open import Relation.Nullary 5 | open import Data.Empty 6 | open import Data.Sum 7 | 8 | open import Support.Equality 9 | open ≡-Reasoning 10 | 11 | open import Vars public 12 | open import Injections.Type public 13 | 14 | invert : ∀ {A : Set} {xs ys : List A} (i : Inj xs ys) → ∀ {t} (y : ys ∋ t) → Dec (∃ \ x → i $ x ≡ y) 15 | invert [] y = no (λ { (() , _)}) 16 | invert ( z ∷ i [ pf ]) y with z ≅∋? y 17 | invert (.y ∷ i [ pf ]) y | yes refl` = yes (zero , refl) 18 | invert ( z ∷ i [ pf ]) y | no z≢y with invert i y 19 | invert ( z ∷ i [ pf ]) y | no z≢y | yes (x , i$x≡y) = yes (suc x , i$x≡y) 20 | invert ( z ∷ i [ pf ]) y | no z≢y | no ¬[i⁻¹y] = no (neither z≢y ¬[i⁻¹y]) 21 | where 22 | neither : ∀ {t} {y : _ ∋ t} → ¬ z ≅∋ y → ¬ (∃ \ x → i $ x ≡ y) → ¬ Σ (_ ∷ _ ∋ _) (λ x → (z ∷ i [ pf ]) $ x ≡ y) 23 | neither ¬1 ¬2 (zero , p) = ¬1 (refl , ≡-to-≅ p) 24 | neither ¬1 ¬2 (suc x , p) = ¬2 (x , p) 25 | 26 | abstract 27 | _∘i_ : ∀ {A : Set}{xs ys zs : List A} → Inj ys zs → Inj xs ys → Inj xs zs 28 | f ∘i g = quo (λ x x₁ → f $ (g $ x₁)) {λ x x₁ → injective g _ _ (injective f _ _ x₁)} 29 | 30 | mutual 31 | id-i : ∀ {A : Set}{xs : List A} → Inj xs xs 32 | id-i = quo (\ _ x → x) {\ _ e → e} 33 | 34 | Inj-thin : ∀ {A : Set}{x : A}{xs ys} → (v : ys ∋ x) -> Inj xs (ys - v) → Inj xs ys 35 | Inj-thin v f = quo (λ x x₁ → thin v x (f $ x₁)) {λ x x₁ → injective f _ _ (thin-inj v x₁)} 36 | 37 | abstract 38 | id-i$ : ∀ {A : Set}{xs : List A} -> ∀ {x}(v : xs ∋ x) -> id-i $ v ≡ v 39 | id-i$ v = iso2 _ _ v 40 | 41 | right-id : ∀ {A : Set}{xs ys : List A} → (i : Inj xs ys) → i ∘i id-i ≡ i 42 | right-id i = begin quo (λ x z → i $ (id-i $ z)) ≡⟨ quo-ext (λ x v → cong (_$_ i) (iso2 _ _ v)) ⟩ 43 | quo (λ x → _$_ i) ≡⟨ iso1 i (λ x eq → injective i _ _ eq) ⟩ 44 | i ∎ 45 | 46 | left-id : ∀ {A : Set}{xs ys : List A} → (i : Inj xs ys) → id-i ∘i i ≡ i 47 | left-id i = begin quo (λ x z → id-i $ (i $ z)) ≡⟨ quo-ext (λ x v → (iso2 _ _ (i $ v))) ⟩ 48 | quo (λ x → _$_ i) ≡⟨ iso1 i (λ x eq → injective i _ _ eq) ⟩ 49 | i ∎ 50 | 51 | apply-∘ : ∀ {A : Set}{xs ys zs : List A} → (j : Inj ys zs) → (i : Inj xs ys) → ∀ {x} {v : x ∈ xs} → (j ∘i i) $ v ≡ j $ (i $ v) 52 | apply-∘ j i {x}{v} = iso2 _ _ v 53 | 54 | assoc-∘i : ∀ {A : Set}{xs ys ws zs : List A} {f : Inj ws zs}{g : Inj _ ws}{h : Inj xs ys} → f ∘i (g ∘i h) ≡ (f ∘i g) ∘i h 55 | assoc-∘i {f = f}{g = g}{h = h} = quo-ext λ x v → 56 | begin f $ (quo (λ x₁ x₂ → g $ (h $ x₂)) $ v) ≡⟨ cong (_$_ f) (iso2 _ _ v) ⟩ 57 | f $ (g $ (h $ v)) ≡⟨ sym (iso2 _ _ (h $ v)) ⟩ 58 | quo (λ x₁ x₂ → f $ (g $ x₂)) $ (h $ v) ∎ 59 | 60 | cong-$ : ∀ {A : Set}{xs ys : List A} {f g : _} {inj1 inj2} → quo {_} {xs} {ys} f {inj1} ≡ quo g {inj2} → ∀ {s} (x : xs ∋ s) → f s x ≡ g s x 61 | cong-$ {A} {xs} {ys} {f} {g} eq x = begin f _ x ≡⟨ sym (iso2 f _ x) ⟩ 62 | quo f $ x ≡⟨ cong (λ f₁ → f₁ $ x) eq ⟩ 63 | quo g $ x ≡⟨ iso2 g _ x ⟩ 64 | g _ x ∎ 65 | 66 | ∘i-inj : ∀ {A : Set}{xs ys zs : List A} → (i : Inj ys zs) (j1 j2 : Inj xs ys) → (i ∘i j1) ≡ (i ∘i j2) → j1 ≡ j2 67 | ∘i-inj i j1 j2 eq = begin j1 ≡⟨ sym (iso1 j1 (λ x x₁ → injective j1 _ _ x₁)) ⟩ 68 | quo (λ x → _$_ j1) ≡⟨ quo-ext (λ x v → injective i _ _ (cong-$ eq v)) ⟩ 69 | quo (λ x → _$_ j2) ≡⟨ iso1 j2 (λ x x₁ → injective j2 _ _ x₁) ⟩ 70 | j2 ∎ 71 | 72 | 73 | Inj-thin-$ : ∀ {A : Set}{x : A}{xs ys} → (v : ys ∋ x) -> (f : Inj xs (ys - v)) -> ∀ {y}(u : xs ∋ y) -> Inj-thin v f $ u ≡ thin v _ (f $ u) 74 | Inj-thin-$ v f u = iso2 _ _ u 75 | 76 | Inj-thin-inj : ∀ {A : Set}{x : A}{xs ys} → (v : ys ∋ x) -> (f g : Inj xs (ys - v)) -> Inj-thin v f ≡ Inj-thin v g -> f ≡ g 77 | Inj-thin-inj v f g eq = ext-$ f g (λ x v₁ → thin-inj v (cong-$ {inj1 = λ x₁ x₂ → injective f _ _ (thin-inj v x₂)} {inj2 = λ x₁ x₂ → injective g _ _ (thin-inj v x₂)} eq v₁)) -- XXX regression 78 | 79 | Inj-thin-∘i : ∀ {A : Set}{x : A}{zs xs ys} → (v : ys ∋ x) -> (f : Inj xs (ys - v)) (m : Inj zs xs) -> Inj-thin v f ∘i m ≡ Inj-thin v (f ∘i m) 80 | Inj-thin-∘i v f m = quo-ext (λ x v₁ → trans (iso2 _ _ (m $ v₁)) (cong (thin v x) (sym (iso2 _ _ v₁)))) 81 | 82 | v∉Inj-thinv : ∀ {A : Set}{x : A}{xs ys} → (v : ys ∋ x) -> (f : Inj xs (ys - v)) -> v ∉ Inj-thin v f 83 | v∉Inj-thinv v f = ∉Im$-∉ (λ x x₁ → thin v _ (f $ x₁)) v (λ b → x∉thinx v (f $ b)) 84 | 85 | test : ∀ {A : Set} {xs ys} → (f : ∀ (x : A) → x ∈ xs → x ∈ ys){inj1 inj2 : ∀ x → {i j : x ∈ xs} → f x i ≡ f x j → i ≡ j} → quo f {inj1} ≡ quo f {inj2} 86 | test f = refl 87 | 88 | weak : ∀ {A : Set}{x : A}{xs ys} → Inj xs ys → Inj xs (x ∷ ys) 89 | weak f = Inj-thin zero f 90 | 91 | apply-weakid : ∀ {A : Set}{xs : List A}{x y}(i : xs ∋ x) -> weak {x = y} id-i $ i ≡ suc i 92 | apply-weakid i = begin 93 | quo (λ x v → thin zero x (id-i $ v)) $ i ≡⟨ iso2 _ _ i ⟩ 94 | suc (id-i $ i) ≡⟨ cong suc (id-i$ i) ⟩ 95 | suc i ∎ 96 | 97 | _∷[]_ : ∀ {A : Set}{x : A}{xs ys} → (v : ys ∋ x) -> Inj xs (ys - v) → Inj (x ∷ xs) ys 98 | v ∷[] f = v ∷ Inj-thin v f [ v∉Inj-thinv v f ] 99 | 100 | cons : ∀ {A : Set}{x : A}{xs ys} → Inj xs ys → Inj (x ∷ xs) (x ∷ ys) 101 | cons z = zero ∷[] z 102 | 103 | abstract 104 | 105 | cons-id : ∀ {A : Set}{x : A}{xs} -> cons id-i ≡ id-i {_} {x ∷ xs} 106 | cons-id = cong-∷[] refl (quo-ext (λ x v → cong suc (id-i$ v))) 107 | 108 | cons-∘i : ∀ {A : Set}{xs ys zs : List A}{x} → (j : Inj ys zs) → (i : Inj xs ys) → cons {A} {x} (j ∘i i) ≡ cons j ∘i cons i 109 | cons-∘i j i = cong-∷[] refl (begin 110 | quo (λ x z → suc (proj₁ (quo' (λ v v₁ → j $ (i $ v₁))) $ z)) 111 | ≡⟨ quo-ext {injg = λ x x₁ → injective i _ _ (injective (weak j) _ _ x₁)} (λ x v → 112 | begin suc (proj₁ (quo' (λ v₁ v₂ → j $ (i $ v₂))) $ v) ≡⟨ cong suc (iso2 _ _ v) ⟩ 113 | suc (j $ (i $ v)) ≡⟨ sym (iso2 _ _ (i $ v)) ⟩ 114 | quo (λ x₁ x₂ → suc (j $ x₂)) $ (i $ v) ∎) ⟩ 115 | quo (λ x v → cons j $ suc (i $ v)) ≡⟨ sym (quo-ext (λ x₁ v → cong (_$_ (cons j)) (iso2 (λ _ x → suc (i $ x)) _ v))) ⟩ 116 | quo (λ x v → cons j $ (quo (λ z x₁ → suc (i $ x₁)) $ v)) ∎) 117 | 118 | 119 | ∘-ext : ∀ {A : Set}{xs ys zs ws : List A} {f : Inj ys zs}{g : Inj xs ys}{f1 : Inj ws zs}{g1 : Inj xs ws} -> f ∘i g ≡ f1 ∘i g1 120 | -> (∀ x (v : xs ∋ x) -> f $ (g $ v) ≡ f1 $ (g1 $ v)) 121 | ∘-ext eq = (\ x v -> trans (sym (apply-∘ _ _)) (trans (cong (\ f -> f $ v) eq) ((apply-∘ _ _)))) 122 | 123 | ext-∘ : ∀ {A : Set}{xs ys zs ws : List A} {f : Inj ys zs}{g : Inj xs ys}{f1 : Inj ws zs}{g1 : Inj xs ws} -> 124 | (∀ x (v : xs ∋ x) -> f $ (g $ v) ≡ f1 $ (g1 $ v)) -> f ∘i g ≡ f1 ∘i g1 125 | ext-∘ eq = ext-$ _ _ (\ x v -> trans (apply-∘ _ _) (trans (eq x v) (sym (apply-∘ _ _)))) 126 | 127 | -- Transforming pointwise representations of universal morphisms into Inj ones. 128 | Equ-universal-quote : ∀ {A : Set} {xs ys : List A} → (i j : Inj xs ys) → ∀ {E} → (e : Inj E xs) -> 129 | (∀ a (y : xs ∋ a) -> i $ y ≡ j $ y -> ∃ \ z -> y ≡ e $ z) -> 130 | {as : List A} (h : Inj as xs) → i ∘i h ≡ j ∘i h → Σ (Inj as E) (λ z → e ∘i z ≡ h ) 131 | Equ-universal-quote {A} {xs} {ys} i j {E} e c {as} h eq = 132 | quo u {λ x {v} {w} eq1 → injective h v w (begin h $ v ≡⟨ proj₂ (f x v) ⟩ 133 | e $ u x v ≡⟨ cong (_$_ e) eq1 ⟩ 134 | e $ u x w ≡⟨ sym (proj₂ (f x w)) ⟩ 135 | h $ w ∎)} 136 | , ext-$ (e ∘i quo u) h (λ x v → begin 137 | (e ∘i quo u) $ v ≡⟨ apply-∘ _ _ ⟩ 138 | e $ (quo u $ v) ≡⟨ cong (_$_ e) (iso2 _ _ v) ⟩ 139 | e $ u x v ≡⟨ sym (proj₂ (f x v)) ⟩ 140 | h $ v ∎) 141 | where 142 | f : ∀ a (y : as ∋ a) -> ∃ \ z -> h $ y ≡ e $ z 143 | f a y = c a (h $ y) (∘-ext eq a y) 144 | u = (λ x v → proj₁ (f x v)) 145 | 146 | Pull-universal-quote : ∀ {A : Set} {X Y Z : List A} → (i : Inj X Z)(j : Inj Y Z) -> ∀ {P} -> (p₁ : Inj P X) (p₂ : Inj P Y) 147 | -> (∀ (a : A) (y : Y ∋ a)(x : X ∋ a) -> i $ x ≡ j $ y -> (∃ \ z -> p₁ $ z ≡ x × p₂ $ z ≡ y)) 148 | -> ∀ {Q} -> (q₁ : Inj Q X) (q₂ : Inj Q Y) -> i ∘i q₁ ≡ j ∘i q₂ -> ∃ \ u -> q₁ ≡ p₁ ∘i u × q₂ ≡ p₂ ∘i u 149 | Pull-universal-quote i j p₁ p₂ uni {Q} q₁ q₂ commutes = 150 | quo u {λ x {v} {w} eq → injective q₁ v w (begin q₁ $ v ≡⟨ sym (proj₁ (proj₂ (f x v))) ⟩ 151 | p₁ $ u x v ≡⟨ cong (_$_ p₁) eq ⟩ 152 | p₁ $ u x w ≡⟨ proj₁ (proj₂ (f x w)) ⟩ 153 | q₁ $ w ∎)} 154 | , ext-$ q₁ (p₁ ∘i quo u) (λ x v → begin 155 | q₁ $ v ≡⟨ sym (proj₁ (proj₂ (f x v))) ⟩ 156 | p₁ $ u x v ≡⟨ cong (_$_ p₁) (sym (iso2 _ _ v)) ⟩ 157 | p₁ $ (quo u $ v) ≡⟨ sym (apply-∘ _ _) ⟩ 158 | (p₁ ∘i quo u) $ v ∎) 159 | , ext-$ q₂ (p₂ ∘i quo u) (λ x v → begin 160 | q₂ $ v ≡⟨ sym (proj₂ (proj₂ (f x v))) ⟩ 161 | p₂ $ u x v ≡⟨ cong (_$_ p₂) (sym (iso2 _ _ v)) ⟩ 162 | p₂ $ (quo u $ v) ≡⟨ sym (apply-∘ _ _) ⟩ 163 | (p₂ ∘i quo u) $ v ∎) 164 | where 165 | f : ∀ a (v : Q ∋ a) -> (∃ \ z -> p₁ $ z ≡ q₁ $ v × p₂ $ z ≡ q₂ $ v) 166 | f a v = uni a (q₂ $ v) (q₁ $ v) (∘-ext commutes a v) 167 | u : ∀ a (v : Q ∋ a) -> _ 168 | u a v = proj₁ (f a v) 169 | 170 | -------------------------------------------------------------------------------- /Injections/Sum.agda: -------------------------------------------------------------------------------- 1 | module Injections.Sum where 2 | 3 | open import Relation.Binary.PropositionalEquality 4 | open ≡-Reasoning 5 | open import Data.Sum 6 | 7 | open import Injections 8 | open import Vars.SumIso 9 | 10 | 11 | inj₁-inj : ∀ {A B : Set} {x : A}{y} -> inj₁ {A = A} {B} x ≡ inj₁ y -> x ≡ y 12 | inj₁-inj refl = refl 13 | 14 | inj₂-inj : ∀ {A B : Set} {x}{y} -> inj₂ {A = A} {B} x ≡ inj₂ y -> x ≡ y 15 | inj₂-inj refl = refl 16 | 17 | left# : ∀ {A : Set} xs {ys : List A} -> Inj xs (xs ++ ys) 18 | left# xs {ys} = quo (λ x i → glue# xs (inj₁ i)) {λ x {i} {j} eq → 19 | inj₁-inj (trans (sym (split∘glue≡id (inj₁ i))) (trans (cong (split# xs) eq) (split∘glue≡id (inj₁ j))))} 20 | 21 | right# : ∀ {A : Set} xs {ys : List A} -> Inj ys (xs ++ ys) 22 | right# xs {ys} = quo (λ x i → glue# xs (inj₂ i)) {λ x {i} {j} eq → 23 | inj₂-inj (trans (sym (split∘glue≡id (inj₂ i))) (trans (cong (split# xs) eq) (split∘glue≡id (inj₂ j))))} 24 | 25 | left : ∀ {A : Set} {xs ys : List A} -> Inj xs (xs ++ ys) 26 | left = left# _ 27 | 28 | right : ∀ {A : Set} {xs ys : List A} -> Inj ys (xs ++ ys) 29 | right {xs = xs} = right# xs 30 | -------------------------------------------------------------------------------- /Injections/Type.agda: -------------------------------------------------------------------------------- 1 | module Injections.Type where 2 | 3 | open import Relation.Nullary 4 | open import Relation.Nullary.Decidable 5 | open import Data.Product 6 | open import Data.Empty 7 | open import Data.Unit 8 | open import Data.Sum 9 | 10 | open import Support.Equality 11 | open ≡-Reasoning 12 | 13 | open import Vars 14 | 15 | -- Inj X Y is a first-order representation for injective maps (∀ S → X ∋ S -> Y ∋ S). 16 | -- It's a little clunkier to work with but its propositional equality 17 | -- is extensional which saves the trouble of proving that 18 | -- the properties we'll care about respect the pointwise equality. 19 | -- 20 | -- The isomorphism with the functional representation is witnessed by quo and _$_ below, 21 | -- among some convenience lemmas. 22 | mutual 23 | data Inj {A : Set}: (List A) → List A → Set where 24 | [] : ∀ {xs} → Inj [] xs 25 | _∷_[_] : ∀ {xs} {y} {ys : List A} (i : y ∈ xs) (is : Inj ys xs) (pf : (i ∉ is)) → Inj (y ∷ ys) xs 26 | 27 | _∉_ : ∀ {A : Set}{y : A} {xs ys} → y ∈ xs → Inj ys xs → Set 28 | v ∉ [] = ⊤ 29 | v ∉ (u ∷ i [ _ ]) = False (v ≅∋? u) × v ∉ i 30 | 31 | proof-irr-False : ∀ {P : Set}{d : Dec P} -> (p q : False d) -> p ≡ q 32 | proof-irr-False {d = yes _} () _ 33 | proof-irr-False {d = no _} _ _ = refl 34 | 35 | proof-irr-∉ : ∀ {A : Set} {xs ys} {y : A} {i : y ∈ xs} (f : Inj ys xs) → (p q : i ∉ f) → p ≡ q 36 | proof-irr-∉ [] p q = refl 37 | proof-irr-∉ (v ∷ f [ pf ]) (Fp , p) (Fq , q) = cong₂ _,_ (proof-irr-False Fp Fq) (proof-irr-∉ f p q) 38 | 39 | cong-∷[] : ∀ {A : Set}{xs ys} {y : A} 40 | {i j : y ∈ xs} → i ≡ j → 41 | {is js : Inj ys xs} → is ≡ js → 42 | ∀ {pf1 pf2} → i ∷ is [ pf1 ] ≡ j ∷ js [ pf2 ] 43 | cong-∷[] {i = i} refl {is} refl {pf1} {pf2} = cong (λ pf → i ∷ is [ pf ]) (proof-irr-∉ is pf1 pf2) 44 | 45 | mkFalse : ∀ {P : Set} → (P → ⊥) → ∀ {d : Dec P} → False d 46 | mkFalse ¬p {yes p} = ¬p p 47 | mkFalse ¬p {no ¬p₁} = tt 48 | 49 | fromFalse : ∀ {P : Set} {d : Dec P} → False d → P → ⊥ 50 | fromFalse {P} {yes p} () 51 | fromFalse {P} {no ¬p} _ = ¬p 52 | 53 | quo' : ∀ {A : Set} {xs ys} → (f : ∀ (x : A) → x ∈ xs → x ∈ ys){inj : ∀ x → {i j : x ∈ xs} → f x i ≡ f x j → i ≡ j} → 54 | Σ (Inj xs ys) \ is → (∀ x (i : x ∈ ys) → (∀ y j → False (i ≅∋? (f y j))) → (i ∉ is)) 55 | quo' {_} {[]} f {inj} = [] , (λ x i x₁ → _) 56 | quo' {_} {x ∷ xs} f {inj} = is , proof 57 | where 58 | rec = (quo' {_} {xs} (λ x₁ x₂ → f x₁ (suc x₂)) {(λ x₁ x₂ → suc-inj1 (inj x₁ x₂))}) 59 | 60 | abstract 61 | pf : f x zero ∉ proj₁ (quo' (λ x₁ x₂ → f x₁ (suc x₂)) {(λ x₁ x₂ → suc-inj1 (inj x₁ x₂))}) 62 | pf = proj₂ rec x (f x zero) (λ y j → mkFalse (lemma y j)) 63 | where 64 | lemma : ∀ y j → f x zero ≅∋ f y (suc j) → ⊥ 65 | lemma .x j (refl , eq) with inj x (≅-to-≡ eq) 66 | ... | () 67 | 68 | is = f x zero ∷ proj₁ rec [ pf ] 69 | 70 | proof : ∀ x i → (∀ y j → False (i ≅∋? (f y j))) → i ∉ is 71 | proof z i e = e x zero , proj₂ rec z i (λ y j → e y (suc j)) 72 | 73 | quo : ∀ {A : Set} {xs ys} → (f : ∀ (x : A) → x ∈ xs → x ∈ ys){inj : ∀ x → {i j : x ∈ xs} → f x i ≡ f x j → i ≡ j} → (Inj xs ys) 74 | quo f {inj} = proj₁ (quo' f {inj}) 75 | 76 | quo-ext : ∀ {A : Set} {xs ys} → {f : ∀ (x : A) → x ∈ xs → x ∈ ys}{injf : ∀ x → {i j : x ∈ xs} → f x i ≡ f x j → i ≡ j} → 77 | {g : ∀ (x : A) → x ∈ xs → x ∈ ys}{injg : ∀ x → {i j : x ∈ xs} → g x i ≡ g x j → i ≡ j} → 78 | (∀ x v → f x v ≡ g x v) → quo f {injf} ≡ quo g {injg} 79 | quo-ext {A} {[]} eq = refl 80 | quo-ext {A} {x ∷ xs} {injf = injf} {injg = injg} eq = cong-∷[] (eq _ zero) (quo-ext {injf = λ x₁ x₂ → suc-inj1 (injf x₁ x₂)} 81 | {injg = λ x₁ x₂ → suc-inj1 (injg x₁ x₂)} (λ x₁ v → eq x₁ (suc v))) 82 | 83 | _$_ : ∀ {A : Set} {xs ys : List A} → Inj xs ys → ∀ {x} → x ∈ xs → x ∈ ys 84 | (i ∷ is [ pf ]) $ zero = i 85 | (i ∷ is [ pf ]) $ suc x₂ = is $ x₂ 86 | [] $ () 87 | 88 | _∉Im_ : ∀ {A : Set} {xs ys : List A} → ∀ {x} (i : x ∈ ys) → (f : Inj xs ys) → Set 89 | i ∉Im f = ∀ b → ¬ i ≡ f $ b 90 | 91 | ∉-∉Im : ∀ {A : Set} {xs ys : List A} → (f : Inj xs ys) → ∀ {x} (i : x ∈ ys) → i ∉ f → i ∉Im f 92 | ∉-∉Im (i₁ ∷ f [ pf ]) .i₁ i∉f zero refl = fromFalse (proj₁ i∉f) refl` 93 | ∉-∉Im (i₁ ∷ f [ pf ]) i i∉f (suc b) eq = ∉-∉Im f i (proj₂ i∉f) b eq 94 | ∉-∉Im [] _ _ () _ 95 | 96 | ∉Im-∉ : ∀ {A : Set} {xs ys : List A} → (f : Inj xs ys) → ∀ {x} (i : x ∈ ys) → i ∉Im f → i ∉ f 97 | ∉Im-∉ [] _ _ = _ 98 | ∉Im-∉ {_}{x ∷ _} {ys} (i ∷ f [ pf ]) {t} i₁ ¬p = mkFalse (aux i pf ¬p) , ∉Im-∉ f i₁ (λ b x → ¬p (suc b) x) 99 | where aux : ∀ {x} (i : ys ∋ x) pf → i₁ ∉Im (i ∷ f [ pf ]) → i₁ ≅∋ i → ⊥ 100 | aux .i₁ pf₁ ¬Im refl` = ¬Im zero refl 101 | 102 | injective : ∀ {A : Set} {xs ys : List A} → (f : Inj xs ys) → ∀ {x} → (a b : x ∈ xs) → f $ a ≡ f $ b → a ≡ b 103 | injective f zero zero eq = refl 104 | injective (i ∷ f [ pf ]) zero (suc b) eq = ⊥-elim (∉-∉Im f i pf b eq) 105 | injective (i ∷ f [ pf ]) (suc a₁) zero eq = ⊥-elim (∉-∉Im f i pf a₁ (sym eq)) 106 | injective (i ∷ f [ pf ]) (suc a₁) (suc b) eq = cong suc (injective f a₁ b eq) 107 | 108 | iso1 : ∀ {A : Set} {xs ys : List A} → (f : Inj xs ys) → (inj : _) → quo (\ x v → f $ v) {inj} ≡ f 109 | iso1 [] _ = refl 110 | iso1 (i ∷ f [ pf ]) inj = cong-∷[] refl (iso1 f (λ x₁ x₂ → suc-inj1 (inj x₁ x₂))) 111 | 112 | iso2 : ∀ {A : Set} {xs ys : List A} → (f : ∀ (x : A) → x ∈ xs → x ∈ ys)(inj : ∀ x → {i j : x ∈ xs} → f x i ≡ f x j → i ≡ j) 113 | → ∀ {x} (v : x ∈ xs) → quo f {inj} $ v ≡ f x v 114 | iso2 f inj zero = refl 115 | iso2 f inj (suc v) = iso2 (λ x₁ x₂ → f x₁ (suc x₂)) (λ x₁ x₂ → suc-inj1 (inj x₁ x₂)) v 116 | 117 | iso1- : ∀ {A : Set} {xs ys : List A} → (f : Inj xs ys) → quo (\ x v → f $ v) {\ x -> injective f _ _} ≡ f 118 | 119 | iso1- f = iso1 f _ 120 | 121 | ext-$ : ∀ {A : Set} {xs ys : List A} → (f g : Inj xs ys) → (∀ x (v : xs ∋ x) -> f $ v ≡ g $ v) -> f ≡ g 122 | ext-$ f g eq = trans (sym (iso1- f)) (trans (quo-ext {injf = λ x → injective f _ _} {injg = λ x → injective g _ _} eq) (iso1- g)) 123 | 124 | ∉Im$-∉ : ∀ {A : Set} {xs ys : List A} (f : ∀ x (v : x ∈ xs) → x ∈ ys){inj} 125 | → ∀ {x} (i : x ∈ ys) → (∀ (b : x ∈ xs) → i ≡ f x b → ⊥) → i ∉ (quo f {inj}) 126 | ∉Im$-∉ f {inj} i eq = ∉Im-∉ (quo f {inj}) i (λ b x → eq b (begin i ≡⟨ x ⟩ quo f $ b ≡⟨ iso2 f inj b ⟩ f _ b ∎)) 127 | -------------------------------------------------------------------------------- /MetaRens.agda: -------------------------------------------------------------------------------- 1 | module MetaRens where 2 | 3 | open import Data.Nat hiding (_≤_) 4 | open import Relation.Nullary 5 | import Relation.Nullary.Decidable as Dec 6 | open import Data.Empty 7 | open import Data.Unit hiding (_≤_) 8 | open import Data.Sum 9 | open import Data.List.All 10 | 11 | open import Support.Equality 12 | open ≡-Reasoning 13 | open import Support.Product 14 | 15 | import Category 16 | 17 | open import Injections 18 | 19 | open import Syntax 20 | 21 | open import Unification.Injections 22 | 23 | 24 | record VarClosure (D : MCtx) (S : MTy) : Set where 25 | constructor _/_ 26 | field 27 | {Ψ} : Ctx 28 | ρ-env : Inj Ψ (ctx S) 29 | body : D ∋ (type S <<- Ψ) 30 | 31 | open VarClosure public 32 | 33 | map-Vc : ∀ {τ D S0 S} → Inj S0 S → VarClosure D (τ <<- S0) → VarClosure D (τ <<- S) 34 | map-Vc j (i / u) = ((j ∘i i) / u) 35 | 36 | record _≈vc_ {D S}(x y : VarClosure D S) : Set where 37 | constructor vc 38 | field 39 | Ψeq : Ψ x ≡ Ψ y 40 | ρeq : ρ-env x ≅ ρ-env y 41 | beq : body x ≅ body y 42 | 43 | to-vc : ∀ {D S}{x y : VarClosure D S} → x ≡ y → x ≈vc y 44 | to-vc refl = vc refl refl refl 45 | 46 | promote : ∀ {D S}{x y : VarClosure D S} → x ≈vc y → x ≡ y 47 | promote (vc refl refl refl) = refl 48 | 49 | map-Vc-inj : ∀ {τ D S0 S} → (i : Inj S0 S) → {x y : VarClosure D (τ <<- S0)} → map-Vc i x ≈vc map-Vc i y -> x ≡ y 50 | map-Vc-inj i {ix / x} {jy / .x} (vc refl eq refl) = cong₂ _/_ (∘i-inj i _ _ (≅-to-≡ eq)) refl 51 | 52 | left-map[id] : ∀ {D S} -> {x : VarClosure D S} -> map-Vc id-i x ≡ x 53 | left-map[id] = cong₂ _/_ (left-id _) refl 54 | 55 | MetaRen : MCtx → MCtx → Set 56 | MetaRen G D = ∀ S → G ∋ S → VarClosure D S 57 | 58 | toSub : ∀ {Sg G D} → MetaRen G D → Sub Sg G D 59 | toSub r = λ S x → mvar (body (r S x)) (ρ-env (r S x)) 60 | 61 | idmr : ∀ {G} → MetaRen G G 62 | idmr = \ S x → id-i / x 63 | 64 | _∘mr_ : ∀ {G1 G2 G3} → MetaRen G2 G3 → MetaRen G1 G2 → MetaRen G1 G3 65 | f ∘mr g = λ S x → let gr = g S x in map-Vc (ρ-env gr) (f _ (body gr)) 66 | 67 | -- Kleisli star 68 | _⋆_ : ∀ {G D} → MetaRen G D → ∀ {S} → VarClosure G S → VarClosure D S 69 | f ⋆ (i / u) = map-Vc i (f _ u) 70 | 71 | bind : ∀ {G D} → MetaRen G D → ∀ {S} → VarClosure G S → VarClosure D S 72 | bind = _⋆_ 73 | 74 | singleton : ∀ {G S} → (u : G ∋ S) → ∀ {Ψ} → Inj Ψ (ctx S) → MetaRen G ((G - u) <: (type S <<- Ψ)) 75 | singleton u j T v with thick u v 76 | singleton u j T v | inj₁ x = id-i / suc (proj₁ x) 77 | singleton .v j ._ v | inj₂ refl` = j / zero 78 | 79 | singleton-refl : ∀ {G S} (u : G ∋ S) {Ψ} (i : Inj Ψ (ctx S)) → singleton u i _ u ≡ i / zero 80 | singleton-refl u i rewrite thick-refl u = refl 81 | 82 | singleton-thin : ∀ {G S} (u : G ∋ S) {Ψ} (i : Inj Ψ (ctx S)) → ∀ {T} (v : (G - u) ∋ T) → singleton u i _ (thin u T v) ≡ id-i / (suc v) 83 | singleton-thin u i v rewrite thick-thin u v = refl 84 | 85 | wk : ∀ {D S x} → VarClosure D S → VarClosure (x ∷ D) S 86 | wk (i / u) = i / (suc u) 87 | 88 | wk-inj : ∀ {G S x} {u v : VarClosure G S} -> wk {G} {S} {x} u ≡ wk v -> u ≡ v 89 | wk-inj refl = refl 90 | 91 | _≡mr_ : ∀ {G D} (f g : MetaRen G D) -> Set 92 | f ≡mr g = ∀ S x -> f S x ≡ g S x 93 | 94 | ∘mr-resp-≡ : ∀ {A B C} {f h : MetaRen B C} {g i : MetaRen A B} → f ≡mr h → g ≡mr i → (f ∘mr g) ≡mr (h ∘mr i) 95 | ∘mr-resp-≡ f≡h g≡i S x rewrite g≡i S x = cong (map-Vc _) (f≡h _ _) 96 | 97 | singleton-inv : ∀ {G S}(u : G ∋ S) {Ψ} (i : Inj Ψ (ctx S)) -> let f = singleton u i in 98 | ∀ S (x : _ ∋ S) -> ∃ \ Ψ -> ∃ \ y -> ∃ \ j -> f (type S <<- Ψ) y ≡ j / x 99 | singleton-inv u i ._ zero = _ , u , i , singleton-refl u i 100 | singleton-inv u i S (suc x) = _ , thin u S x , id-i , singleton-thin u i x 101 | 102 | module MRop = Category MCtx (λ X Y → MetaRen Y X) (λ f g → g ∘mr f) idmr (λ f g → ∀ S x → f S x ≡ g S x) 103 | module MRopProps = MRop.Props (λ S x → cong₂ _/_ (sym assoc-∘i) refl) (λ S x → left-map[id]) (λ S x → cong₂ _/_ (right-id _) refl) 104 | (λ {A} {B} → record { refl = λ S x₁ → refl; sym = λ x S x₁ → sym (x _ _); trans = λ x x₁ S x₂ → trans (x S x₂) (x₁ S x₂) }) 105 | (λ eq₁ eq₂ S x → ∘mr-resp-≡ eq₂ eq₁ S x) 106 | 107 | id-epic : ∀ {G} -> MRop.Monic (idmr {G}) 108 | id-epic eq S x = map-Vc-inj id-i (to-vc (eq S x)) 109 | 110 | singleton-epic : ∀ {G S}(u : G ∋ S) {Ψ} (i : Inj Ψ (ctx S)) -> let f = singleton u i in 111 | MRop.Monic f 112 | singleton-epic u i eq S x with singleton-inv u i S x 113 | ... | _ , y , j , eq' with eq _ y 114 | ... | eq'' rewrite eq' = map-Vc-inj j (to-vc eq'') 115 | 116 | eval : ∀ {p} {A : Set} {P : A → Set p} {xs : List A} → 117 | All P xs → (∀ (x : A) → xs ∋ x → P x) 118 | eval (px ∷ f) x zero = px 119 | eval (px ∷ f) x (suc u) = eval f x u 120 | eval [] _ () 121 | 122 | reify : ∀ {p} {A : Set} {P : A → Set p} {xs} → 123 | (f : ∀ x → xs ∋ x → P x) → Σ (All P xs) \ a -> (∀ x i -> eval a x i ≡ f x i) 124 | reify {p} {A} {P} {[]} f = [] , \ _ () 125 | reify {p} {A} {P} {x ∷ xs} f = (f x zero ∷ proj₁ rec) , (\ { ._ zero -> refl ; _ (suc i) -> proj₂ rec _ i}) 126 | where rec = reify (λ x u → f x (suc u)) 127 | 128 | _hasBody_ : ∀ {G}{T}(cl : VarClosure G T) {S}(x : G ∋ S) -> Set 129 | _hasBody_ {G} {T} cl {S} x = ∃ \ (j : Inj (ctx S) (ctx T)) -> type T ≡ type S × cl ≅ (j / x) 130 | 131 | dec-HasBody : ∀ {G}{T}(cl : VarClosure G T) {S}(x : G ∋ S) -> Dec (cl hasBody x) 132 | dec-HasBody (j / y) x with y ≅∋? x 133 | dec-HasBody {G} {.type₁ <<- ctx₁} (j / .x) {type₁ <<- ctx} x | yes refl` = yes (j , refl`) 134 | dec-HasBody {G} {type <<- ctx₁} (j / y) {type₁ <<- ctx} x | no ¬p = no (aux ¬p) 135 | where aux : ∀ {type}{y : _ ∋ (type <<- _)} -> (¬ (y ≅∋ x)) -> Σ (Inj ctx ctx₁) (λ j₁ → Σ (type ≡ type₁) (λ x₁ → j / y ≅ j₁ / x)) → ⊥ 136 | aux ¬p₁ (proj₁ , refl , eq) = ¬p₁ ((cong (_<<-_ _) (_≈vc_.Ψeq (to-vc (≅-to-≡ eq)))) , (_≈vc_.beq (to-vc (≅-to-≡ eq)))) 137 | 138 | Image : ∀ {G G1} (f : MetaRen G G1) {S} (x : G1 ∋ S) -> Set 139 | Image f x = ∃ \ Ψ -> ∃ \ y -> ∃ \ j -> f (_ <<- Ψ) y ≡ j / x 140 | 141 | thin1 : ∀ {S G T} -> VarClosure (S ∷ G) T -> VarClosure (S ∷ S ∷ G) T 142 | thin1 (i / zero) = i / zero 143 | thin1 (i / suc v) = i / (suc (suc v)) 144 | 145 | epic-inv : ∀ {G G1}(f : MetaRen G G1) -> MRop.Monic f -> ∀ S (x : _ ∋ S) -> Image f x 146 | epic-inv f f-epic S x with any? (\ v -> dec-HasBody (f _ v) x) 147 | epic-inv f f-epic (._ <<- _) x | yes (T , v , j , refl , eq) = _ , v , j , ≅-to-≡ eq 148 | epic-inv {G} {G1} f f-epic S x | no ¬p = ⊥-elim absurd where 149 | g1 = \ S v -> wk (singleton x id-i S v) 150 | g2 = \ S v -> thin1 (singleton x id-i S v) 151 | 152 | g1∘f≡g2∘f : (g1 ∘mr f) ≡mr (g2 ∘mr f) 153 | g1∘f≡g2∘f S v with f _ v | inspect (f _) v | thick x (body (f _ v)) 154 | g1∘f≡g2∘f S₁ v | jfv / fv | _ | inj₁ x₁ = refl 155 | g1∘f≡g2∘f (._ <<- _) v | jfv / .x | ⌞ eq ⌟ | inj₂ refl` 156 | = ⊥-elim (¬p (_ , v , jfv , refl , Het.≡-to-≅ eq)) 157 | 158 | absurd : ⊥ 159 | absurd with trans (f-epic {S ∷ S ∷ (G1 - x)} {g1} {g2} g1∘f≡g2∘f _ x) (cong thin1 (singleton-refl x id-i)) 160 | ... | () 161 | -------------------------------------------------------------------------------- /NOTERM/Unification.agda: -------------------------------------------------------------------------------- 1 | module NOTERM.Unification where 2 | 3 | open import Data.Nat hiding (_≤_) 4 | open import Relation.Binary 5 | open DecTotalOrder Data.Nat.decTotalOrder 6 | using () renaming (refl to ≤-refl; trans to ≤-trans) 7 | open import Data.Empty 8 | open import Data.Sum renaming (inj₁ to yes; inj₂ to no; map to map⊎) 9 | 10 | open import Support.Equality 11 | open ≡-Reasoning 12 | 13 | open import Injections 14 | open import MetaRens 15 | 16 | open import Syntax 17 | 18 | open import NOTERM.Unification.Specification 19 | open import Unification.MetaRens 20 | open import Unification.Injections 21 | open import Unification.OccursCheck 22 | open import NOTERM.Unification.Pruning 23 | open import NOTERM.Unification.Inversion 24 | 25 | 26 | flexSame : ∀ {Sg G D S} → (u : G ∋ S) → (i j : Inj (ctx S) D) → ∃σ-pat Max (Unifies {Sg} (Tm.mvar u i) (mvar u j)) 27 | flexSame {Sg} {G} {D} {B <<- Ss} u i j = _ , σ 28 | , [σ]Unifies[mvar[u,i],mvar[u,j]] 29 | , sup-σ 30 | where 31 | i,j⇒e = equalizer i j 32 | open Equalizer i,j⇒e 33 | 34 | σ = toSub (singleton u e) 35 | 36 | [σ]Unifies[mvar[u,i],mvar[u,j]] : ren i (σ _ u) ≡T ren j (σ _ u) 37 | [σ]Unifies[mvar[u,i],mvar[u,j]] rewrite thick-refl u = ≡-T (cong (mvar zero) commutes) 38 | 39 | sup-σ : Sup (Unifies (mvar u i) (mvar u j)) σ 40 | sup-σ {G'} ρ ρ-unifies = δ , ρ≡δ∘σ where 41 | 42 | ∃s[ren[e,s]≡ρ[u]] = forget (lift-equalizer i,j⇒e (ρ (B <<- Ss) u) ρ-unifies) 43 | 44 | δ : Sub< false > Sg (B <<- E ∷ G - u) G' 45 | δ ._ zero = proj₁ ∃s[ren[e,s]≡ρ[u]] 46 | δ S₁ (suc v) = ρ _ (thin u _ v) 47 | 48 | ρ≡δ∘σ : ρ ≡s (δ ∘s σ) 49 | ρ≡δ∘σ S v with thick u v 50 | ρ≡δ∘σ S .(thin u S w) | inj₁ (w , refl) = sym (ren-id (ρ S (thin u S w))) 51 | ρ≡δ∘σ .(B <<- Ss) .u | inj₂ refl` = sym (proj₂ ∃s[ren[e,s]≡ρ[u]]) 52 | 53 | 54 | flexRigid : ∀ {Sg G D S} (u : G ∋ S) (i : Inj (ctx S) D) (s : Tm Sg (G - u) D (! type S)) → Spec (mvar u i) (sub (thin-s u) s) 55 | flexRigid {Sg} {G} {S = S} u i s with prune i s 56 | ... | ((Pr ρ , m) , ρ-sup) 57 | with invertTm i s ρ m 58 | ... | no NotInv = no λ {(_ , σ , σ-unifies) → 59 | let 60 | eq = begin 61 | ren i (σ S u) ≡⟨ T-≡ σ-unifies ⟩ 62 | subT σ (subT (thin-s u) s) ≡⟨ Sub∘.subT-∘ s ⟩ 63 | subT (σ ∘s thin-s u) s ∎ 64 | σ≤ρ = ρ-sup (σ ∘s thin-s u) (σ S u) (≡-T eq) 65 | 66 | in NotInv (proj₁ σ≤ρ) (σ S u , ≡-T (begin ren i (σ S u) ≡⟨ eq ⟩ 67 | subT (σ ∘s thin-s u) s ≡⟨ subT-ext (proj₂ σ≤ρ) s ⟩ 68 | subT (proj₁ σ≤ρ ∘s ρ) s ≡⟨ sym (Sub∘.subT-∘ s) ⟩ 69 | subT (proj₁ σ≤ρ) (subT ρ s) ∎))} 70 | 71 | ... | yes (t , ren[i,t]≡sub[ρ,s]) = yes 72 | (_ , σ , 73 | ≡-T (begin 74 | ren i (σ _ u) ≡⟨ cong (ren i) σ[u]≡t ⟩ 75 | ren i t ≡⟨ ren[i,t]≡sub[ρ,s] ⟩ 76 | sub ρ s ≡⟨ sub-ext ρ≡σ∘thin[u] s ⟩ 77 | sub (σ ∘s thin-s u) s ≡⟨ sym (sub-∘ s) ⟩ 78 | sub σ (sub (thin-s u) s) ∎) , σ-sup ) 79 | where 80 | σ : Sub Sg G _ 81 | σ S v with thick u v 82 | σ S v | inj₁ (w , eq) = ρ _ w 83 | σ ._ .u | inj₂ refl` = t 84 | 85 | σ[u]≡t : σ _ u ≡ t 86 | σ[u]≡t rewrite thick-refl u = refl 87 | 88 | ρ≡σ∘thin[u] : ρ ≡s (σ ∘s thin-s u) 89 | ρ≡σ∘thin[u] S y rewrite thick-thin u y = sym (ren-id _) 90 | 91 | σ-sup : Sup (Unifies (mvar u i) (sub (thin-s u) s)) σ 92 | σ-sup ρ₁ ρ₁-unifies = δ , ρ₁≡δ∘σ where 93 | ren[i,ρ₁[u]]≡sub[ρ₁∘thin[u],s] = begin 94 | sub ρ₁ (mvar u i) ≡⟨ T-≡ ρ₁-unifies ⟩ 95 | sub ρ₁ (sub (thin-s u) s) ≡⟨ Sub∘.subT-∘ s ⟩ 96 | sub (ρ₁ ∘s thin-s u) s ∎ 97 | 98 | ρ₁∘thin[u]≤ρ = ρ-sup (ρ₁ ∘s thin-s u) (ρ₁ _ u) (≡-T ren[i,ρ₁[u]]≡sub[ρ₁∘thin[u],s]) 99 | δ = proj₁ ρ₁∘thin[u]≤ρ 100 | ρ₁∘thin[u]≡δ∘ρ = proj₂ ρ₁∘thin[u]≤ρ 101 | 102 | ρ₁≡δ∘σ : ρ₁ ≡s (δ ∘s σ) 103 | ρ₁≡δ∘σ S u₁ with thick u u₁ 104 | ρ₁≡δ∘σ S ._ | inj₁ (v , refl) = begin 105 | ρ₁ S (thin u S v) ≡⟨ sym (ren-id (ρ₁ S (thin u S v))) ⟩ 106 | sub ρ₁ (thin-s u S v) ≡⟨ ρ₁∘thin[u]≡δ∘ρ S v ⟩ 107 | sub δ (ρ S v) ∎ 108 | ρ₁≡δ∘σ ._ .u | inj₂ refl` = ren-inj i (ρ₁ _ u) (sub δ t) -- crucial use of injectivity to show 109 | (begin -- that we got the most general solution 110 | ren i (ρ₁ _ u) ≡⟨ ren[i,ρ₁[u]]≡sub[ρ₁∘thin[u],s] ⟩ 111 | sub (ρ₁ ∘s thin-s u) s ≡⟨ sub-ext ρ₁∘thin[u]≡δ∘ρ s ⟩ 112 | sub (δ ∘s ρ) s ≡⟨ sym (Sub∘.subT-∘ s) ⟩ 113 | sub δ (sub ρ s) ≡⟨ cong (sub δ) (sym ren[i,t]≡sub[ρ,s]) ⟩ 114 | sub δ (ren i t) ≡⟨ sub-nat t ⟩ 115 | ren i (sub δ t) ∎) 116 | 117 | 118 | flexAny : ∀ {Sg G D S} → (u : G ∋ S) → (i : Inj (ctx S) D) → (t : Tm Sg G D (! (type S))) → Spec (mvar u i) t 119 | flexAny u i t with check u t 120 | flexAny u i .(sub (thin-s u) s) | inj₁ (s , refl) = flexRigid u i s 121 | flexAny u i .(mvar u j) | inj₂ (G1 , j , [] , refl) = yes (flexSame u i j) 122 | flexAny u i .(wrap (d ∷ c) (mvar u j)) | inj₂ (G1 , j , d ∷ c , refl) = no λ {(D1 , s , eq) → 123 | No-Cycle (subD s d) (subC s c) (s _ u) i j 124 | (trans (T-≡ eq) (wrap-sub s (d ∷ c) (mvar u j)))} 125 | 126 | {-# NO_TERMINATION_CHECK #-} 127 | mutual 128 | unify : ∀ {Sg G D T} → (x y : Tm Sg G D T) → Spec x y 129 | -- congruence and directly failing cases 130 | unify (con c xs) (con c₁ ys) with c ≅∋? c₁ 131 | unify (con c xs) (con c₁ ys) | no c≢c₁ = no (λ {(_ , _ , eq) → c≢c₁ (con-inj₁ eq)}) 132 | unify (con c xs) (con .c ys) | yes refl` = cong-spec (con c) (unifyTms xs ys) 133 | unify (var x xs) (var y ys) with x ≅∋? y 134 | unify (var x xs) (var y ys) | no x≢y = no (λ {(_ , _ , eq) → x≢y (var-inj₁ eq)}) 135 | unify (var x xs) (var .x ys) | yes refl` = cong-spec (var x) (unifyTms xs ys) 136 | unify (lam x) (lam y) = cong-spec lam {x} {y} (unify x y) 137 | unify (con _ _) (var _ _) = no λ {(_ , _ , ())} 138 | unify (var _ _) (con _ _) = no λ {(_ , _ , ())} 139 | 140 | -- flexible cases 141 | unify (mvar u i) t = flexAny u i t 142 | unify t (mvar u i) = spec-comm (mvar u i) t (flexAny u i t) 143 | 144 | unifyTms : ∀ {Sg G D Ts} → (x y : Tms Sg G D Ts) → Spec x y 145 | unifyTms [] [] = yes (∃σMax[Unifies[x,x]] []) 146 | unifyTms (s ∷ xs) (t ∷ ys) 147 | with unify s t 148 | ... | no ¬unify[s,t] = no λ {(_ , ρ , eq ∷ _) → ¬unify[s,t] (_ , ρ , eq)} 149 | ... | yes (_ , σ , eq , sup) 150 | with unifyTms (subs σ xs) (subs σ ys) 151 | ... | no ¬unify[σxs,σys] = no λ {(_ , σ1 , eqt ∷ eqts) → ¬unify[σxs,σys] (shift eqts under σ by sup σ1 eqt) } 152 | ... | yes (_ , σ1 , eq1 , sup1) = yes (_ , (σ1 ∘s σ) , optimist s t xs ys σ σ1 (eq , sup) (eq1 , sup1)) 153 | -------------------------------------------------------------------------------- /NOTERM/Unification/Inversion.agda: -------------------------------------------------------------------------------- 1 | module NOTERM.Unification.Inversion where 2 | 3 | open import Relation.Binary.PropositionalEquality 4 | open import Data.Empty 5 | open import Data.Sum renaming (map to map⊎) 6 | open import Data.Sum renaming (inj₁ to yes; inj₂ to no) 7 | 8 | open import Injections 9 | 10 | open import Syntax 11 | open import NOTERM.Unification.Pruning 12 | 13 | 14 | NotInv : ∀ {Sg G D D' T} (i : Inj D D') (t : Term Sg G D' T) → Set 15 | NotInv {Sg} {G} i t = ∀ {G1} (σ : Sub< false > Sg G G1) -> ¬ ∃ \ s -> renT i s ≡T subT σ t 16 | 17 | map-NI : ∀ {Sg G DI D T D' T'}{i : Inj D' DI}{t : Term Sg G D T} 18 | (d : DTm Sg G (DI , T') (D , T)) → NotInv (wrapD-Inj d i) t → NotInv i (wrapD d t) 19 | map-NI lam notinv σ (lam s , lam eq) = notinv σ (s , eq) 20 | map-NI (head ts) notinv σ (s ∷ ss , eq ∷ eq') = notinv σ (s , eq) 21 | map-NI (tail t) notinv σ (s ∷ ss , eq ∷ eq') = notinv σ (ss , eq') 22 | map-NI (con c) notinv σ (con .c ss , con refl eq) = notinv σ (ss , eq) 23 | map-NI (var ._) notinv σ (var x ss , var refl eq) = notinv σ (ss , eq) 24 | map-NI (con c) notinv σ (var x ss , ()) 25 | map-NI (con c) notinv σ (mvar u j , ()) 26 | map-NI (var x) notinv σ (con c ss , ()) 27 | map-NI (var x) notinv σ (mvar u j , ()) 28 | 29 | NI-var : ∀ {Sg G D D1 Ss B} {i : Inj D D1} {ts : Tms Sg G D1 Ss} {x : D1 ∋ (Ss ->> B)} 30 | → ¬ ∃ (λ y → i $ y ≡ x) → NotInv {T = inj₁ _} i (var x ts) 31 | NI-var ¬∃y[iy≡x] σ (con c ts , ()) 32 | NI-var ¬∃y[iy≡x] σ (mvar u j , ()) 33 | NI-var ¬∃y[iy≡x] σ (var y _ , var iy≡x _) = ¬∃y[iy≡x] (y , iy≡x) 34 | 35 | mutual 36 | invertTm' : ∀ {Sg G Ss D T} (i : Inj Ss D) (t : Tm Sg G D T) → 37 | AllMV∈ i t → i ⁻¹ t ⊎ NotInv i t 38 | invertTm' i (con c ts) (con m) = map⊎ (con c) (map-NI (con c)) (invertTm's i ts m) 39 | invertTm' i (mvar v h) (mvar (k , comm)) = yes (mvar v (k , comm)) 40 | invertTm' i (var x ts) (var m) 41 | with invert i x | invertTm's i ts m 42 | ... | yes (y , iy≡x) | yes i⁻¹ts = yes (var y iy≡x i⁻¹ts) 43 | ... | _ | no NI[ts] = no (map-NI (var x) NI[ts]) 44 | ... | no ¬∃y[iy≡x] | _ = no (NI-var ¬∃y[iy≡x]) 45 | invertTm' i (lam t) (lam m) = map⊎ lam (map-NI {t = t} lam) (invertTm' (cons i) t m) 46 | 47 | invertTm's : ∀ {Sg G Ss D T} (i : Inj Ss D) (t : Tms Sg G D T) → 48 | AllMV∈ i t → i ⁻¹ t ⊎ NotInv i t 49 | invertTm's i [] m = yes [] 50 | invertTm's i (t ∷ ts) (mt ∷ mts) 51 | with invertTm' i t mt | invertTm's i ts mts 52 | ... | yes p | yes ps = yes (p ∷ ps) 53 | ... | yes p | no ¬ps = no (map-NI (tail t) ¬ps) 54 | ... | no ¬p | _ = no (map-NI {t = t} (head ts) ¬p) 55 | 56 | invertTm : ∀ {Sg G G1 Ss D T} (i : Inj Ss D) → (t : Tm Sg G D T) → (ρ : Sub Sg G G1) → ρ / t ∈ i 57 | → (∃ \ s → ren i s ≡ sub ρ t) ⊎ NotInv i (sub ρ t) 58 | invertTm i t ρ m = map⊎ forget (λ x → x) (invertTm' i (sub ρ t) m) 59 | 60 | -------------------------------------------------------------------------------- /NOTERM/Unification/Pruning.agda: -------------------------------------------------------------------------------- 1 | module NOTERM.Unification.Pruning where 2 | 3 | open import Data.Empty 4 | open import Data.Sum 5 | 6 | open import Support.Equality 7 | open ≡-Reasoning 8 | 9 | open import Injections 10 | open import MetaRens 11 | 12 | open import Syntax 13 | 14 | open import NOTERM.Unification.Specification 15 | open import Unification.MetaRens 16 | open import Unification.Injections 17 | open import Unification.Pruning.Epi-Decr 18 | 19 | data AllMV∈ {Sg : Ctx} {G : MCtx} {D0 D : Ctx} (i : Inj D0 D) : ∀ {T} → Term Sg G D T → Set where 20 | [] : AllMV∈ i {inj₂ []} [] 21 | _∷_ : ∀ {S Ss t ts} → (m : AllMV∈ i {inj₁ S} t) → (ms : AllMV∈ i {inj₂ Ss} ts) → AllMV∈ i {inj₂ (S ∷ Ss)} (t ∷ ts) 22 | con : ∀ {Ss B c ts} → (ms : AllMV∈ i {inj₂ Ss} ts) → AllMV∈ i {inj₁ (! B)} (con c ts) 23 | var : ∀ {Ss B x ts} → (ms : AllMV∈ i {inj₂ Ss} ts) → AllMV∈ i {inj₁ (! B)} (var x ts) 24 | lam : ∀ {S Ss B t} → (m : AllMV∈ (cons i) {inj₁ (Ss ->> B)} t) → AllMV∈ i {inj₁ ((S ∷ Ss) ->> B)} (lam t) 25 | mvar : ∀ {Ss B} {v : G ∋ (B <<- Ss)}{h} → (∃ \ k → i ∘i k ≡ h) → AllMV∈ i {inj₁ (! B)} (mvar v h) 26 | 27 | -- s / t ∈ i holds iff all the variables appearing as arguments to 28 | -- meta-vars in (sub s t) are present in the image of i 29 | _/_∈_ : ∀ {Sg G1 G2 D1 D2 T} → Sub Sg G1 G2 → Term Sg G1 D2 T → Inj D1 D2 → Set 30 | _/_∈_ s t i = AllMV∈ i (subT s t) 31 | 32 | -- (\ s -> s / t ∈ i) is closed under post-composition, we'll need this in the _∷_ case of prune. 33 | _/_∈_-∘Closed : ∀ {Sg G1 G2 G3 D1 D2 T} (f : Sub Sg G2 G3) {g : Sub Sg G1 G2} {i : Inj D1 D2} → 34 | ∀ {t : Term Sg G1 D2 T} → g / t ∈ i → (f ∘s g) / t ∈ i 35 | _/_∈_-∘Closed f m = subst (AllMV∈ _) (subT-∘ _) (go f m) where 36 | mutual 37 | go2 : ∀ {Sg G X Y Z T} {i : Inj Y Z} {h : Inj X Z} → (∃ \ k → i ∘i k ≡ h) → 38 | (t : Tm Sg G X T) → AllMV∈ i (renT h t) 39 | go2 le (con c ts) = con (go2s le ts) 40 | go2 le (var x ts) = var (go2s le ts) 41 | go2 (k , i∘k≡h) (lam t) = lam (go2 (cons k , trans (sym (cons-∘i _ _)) (cong cons i∘k≡h)) t) 42 | go2 (k , i∘k≡h) (mvar u j) = mvar (k ∘i j , trans assoc-∘i (cong (λ h → h ∘i j) i∘k≡h)) 43 | 44 | go2s : ∀ {Sg G X Y Z T} {i : Inj Y Z}{h : Inj X Z} → (∃ \ k → i ∘i k ≡ h) → 45 | (t : Tms Sg G X T) → AllMV∈ i (renT h t) 46 | go2s le [] = [] 47 | go2s le (t ∷ ts) = go2 le t ∷ go2s le ts 48 | 49 | go : ∀ {Sg G1 G3 D1 D2 T} (f : Sub Sg G1 G3) {i : Inj D1 D2} → 50 | {t : Term Sg G1 D2 T} → AllMV∈ i t → f / t ∈ i 51 | go f [] = [] 52 | go f (m ∷ ms) = go f m ∷ go f ms 53 | go f (con ms) = con (go f ms) 54 | go f (var ms) = var (go f ms) 55 | go f (lam m) = lam (go f m) 56 | go f {i = i} (mvar {v = v} {h = h} (k , i∘k≡h)) = go2 (k , i∘k≡h) (f _ v) 57 | 58 | 59 | -- A few properties of substitutions respected by _/_∈_ 60 | _/_∈_-∘ : ∀ {Sg G1 G2 G3 D1 D2 T} {f : Sub Sg G2 G3} (g : Sub Sg G1 G2) {i : Inj D1 D2} → 61 | ∀ (t : Term Sg G1 D2 T) → f / subT g t ∈ i → (f ∘s g) / t ∈ i 62 | _/_∈_-∘ g t m = subst (AllMV∈ _) (subT-∘ _) m 63 | 64 | _/_∈_-ext : ∀ {Sg G G1 D1 D2 T} {i : Inj D1 D2} {f g : Sub Sg G G1} → 65 | f ≡s g → ∀ {t : Term Sg G D2 T} → f / t ∈ i → g / t ∈ i 66 | _/_∈_-ext f≡g {t} m = subst (AllMV∈ _) (subT-ext f≡g t) m 67 | 68 | -- In the flexible-rigid case we'll need to find z and ρ such that ren i z ≡ sub ρ t, 69 | -- this module is about finding such a ρ, which we call the pruner. 70 | -- Its role is to handle occurrences in t like (mvar u j) where there are variables in j 71 | -- which are not in i: ρ will substitute u with a term that ignores them, 72 | -- since their presence would make finding z impossible. 73 | record Pruner {Sg G D1 D2 T} (i : Inj D1 D2) (t : Term Sg G D2 T) : Set where 74 | constructor Pr_,_ 75 | field 76 | {G1} : MCtx 77 | pruner : MetaRen G G1 78 | prunes : toSub pruner / t ∈ i 79 | 80 | record PrunerSub {Sg G D1 D2 T} (i : Inj D1 D2) (t : Term Sg G D2 T) : Set where 81 | constructor Pr_,_ 82 | field 83 | {G1} : MCtx 84 | pruner : Sub Sg G G1 85 | prunes : pruner / t ∈ i 86 | 87 | open Pruner using (pruner; prunes) 88 | 89 | _∙_ : ∀ {Sg G D1 D2 T} → {i : Inj D1 D2} {t : Term Sg G D2 T} → 90 | ∀ {D1 D2 T} → {j : Inj D1 D2} {s : Term Sg G D2 T} → 91 | (∀ {G1}{σ : Sub Sg G G1} → σ / t ∈ i → σ / s ∈ j) → Pruner i t → Pruner j s 92 | f ∙ (Pr ρ , m) = Pr ρ , f m 93 | 94 | mutual 95 | prune' : ∀ {Sg G D1 D2 T} {i : Inj D1 D2} (t : Tm Sg G D2 T) 96 | → Pruner i t 97 | -- congruence cases 98 | prune' (con c ts) = con ∙ prune's ts 99 | prune' (var x ts) = var ∙ prune's ts 100 | prune' (lam t) = lam ∙ prune' t 101 | 102 | prune' {i = i} (mvar u j) = Pr singleton u p₂ , mvar aux where 103 | open Pullback (pullback i j) 104 | -- (toSub (singleton u p₂)) (mvar u j) = mvar zero (j ∘i p₂), so we 105 | -- need aux to show that (j ∘i p₂) only contains variables in i 106 | aux : ∃ \ k → i ∘i k ≡ j ∘i ρ-env (singleton u p₂ _ u) 107 | aux rewrite thick-refl u = p₁ , commutes 108 | 109 | prune's : ∀ {Sg G D1 D2 T} {i : Inj D1 D2} (t : Tms Sg G D2 T) 110 | → Pruner i t 111 | prune's [] = Pr idmr , [] 112 | prune's {i = i} (t ∷ ts) = Pr p₁ ∘mr pruner pr-t , (prunes[t] ∷ prunes[ts]) 113 | where 114 | pr-t = prune' {i = i} t 115 | pr-ts = prune's {i = i} ts 116 | push = (pushout (pruner pr-t) (pruner pr-ts)) 117 | open Mixed.Pushout push 118 | prunes[t] = _/_∈_-∘Closed (toSub p₁) {toSub (pruner pr-t)} {i} {t} (prunes pr-t) 119 | prunes[ts] = _/_∈_-ext {f = toSub (p₂ ∘mr pruner pr-ts)} (λ S u → ↓↓-inj (sym ([]eq commutes S u))) 120 | (_/_∈_-∘Closed (toSub p₂) {_} {i} {ts} (prunes pr-ts)) 121 | 122 | 123 | -- prune-sup makes use of the universal property of pullbacks to prove 124 | -- that the pruner computed above is more general than any possible 125 | -- solution to the equation runT i z ≡ sub s t from which we started. 126 | mutual 127 | prune-sup : ∀ {Sg G D1 D2 T} (i : Inj D1 D2) (t : Tm Sg G D2 T) → 128 | ∀ {G1} (s : Sub< false > Sg G G1) z → ren i z ≡T sub s t → s ≤ toSub (pruner (prune' {i = i} t)) 129 | prune-sup i (con c ts) s (con c₁ ts₁) (con _ eq) = prune-sups i ts s ts₁ eq 130 | prune-sup i (var x ts) s (var x₁ ts₁) (var eqv eq) = prune-sups i ts s ts₁ eq 131 | prune-sup i (lam t) s (lam z) (lam eq) = prune-sup (cons i) t s z eq 132 | 133 | prune-sup i (con c ts) s (mvar u j) () 134 | prune-sup i (con c ts) s (var x ts₁) () 135 | prune-sup i (var x ts) s (con c ts₁) () 136 | prune-sup i (var x ts) s (mvar u j) () 137 | 138 | prune-sup {Sg} {G} i (mvar {Ss = Ss} {B} u j) {G2} s z eq = δ , s≡δ∘pruner 139 | where 140 | pull = pullback i j 141 | open Pullback pull 142 | open Σ (forget (lift-pullback pull z (s (B <<- Ss) u) eq)) renaming 143 | (proj₁ to x; proj₂ to ren[p₂,x]≡s[u]) 144 | δ : (S : MTy) → B <<- P ∷ G - u ∋ S → Tm< false > Sg G2 (ctx S) ([] ->> type S) 145 | δ .(B <<- P) zero = x 146 | δ S (suc v) = s S (thin u S v) 147 | s≡δ∘pruner : (S : MTy) (v : G ∋ S) → s S v ≡ sub δ (toSub (singleton u p₂) S v) 148 | s≡δ∘pruner S v with thick u v 149 | s≡δ∘pruner S .(thin u S v) | inj₁ (v , refl) = sym (ren-id _) 150 | s≡δ∘pruner .(B <<- Ss) .u | inj₂ refl` = sym ren[p₂,x]≡s[u] 151 | 152 | 153 | prune-sups : ∀ {Sg G D1 D2 T} (i : Inj D1 D2) (t : Tms Sg G D2 T) → 154 | ∀ {G1} (s : Sub< false > Sg G G1) z → rens i z ≡T subs s t → s ≤ toSub (pruner (prune's {i = i} t)) 155 | prune-sups i [] s [] eq = s , (λ S u → sym (ren-id _)) 156 | prune-sups {Sg} {G} i (t ∷ ts) s (z ∷ zs) (eqt ∷ eqts) = down uni , (λ S u → 157 | begin 158 | s S u ≡⟨ proj₂ s≤pr-t S u ⟩ 159 | (proj₁ s≤pr-t ∘s toSub (pruner pr-t)) S u ≡⟨ sub-ext (λ S₁ u₁ → sym (uni∘p₁≡q₁ S₁ u₁)) (toSub (pruner pr-t) S u) ⟩ 160 | ((down uni ∘s toSub p₁) ∘s toSub (pruner pr-t)) S u ≡⟨ sym (Sub∘.subT-∘ {f = down uni} {g = toSub p₁} 161 | (toSub (pruner pr-t) S u)) ⟩ 162 | (down uni ∘s (toSub p₁ ∘s toSub (pruner pr-t))) S u ∎) 163 | where 164 | pr-t = prune' {i = i} t 165 | pr-ts = prune's {i = i} ts 166 | push = pushout (pruner pr-t) (pruner pr-ts) 167 | open Mixed.Pushout {Sg} push 168 | s≤pr-t : s ≤ toSub (pruner (prune' {i = i} t)) 169 | s≤pr-t = prune-sup i t s z eqt 170 | s≤pr-ts : s ≤ toSub (pruner (prune's {i = i} ts)) 171 | s≤pr-ts = prune-sups i ts s zs eqts 172 | eq = ES (λ S u → cong ↓↓ (trans (sym (proj₂ s≤pr-t S u)) (proj₂ s≤pr-ts S u))) 173 | uni = (ESub.⟦ universal (ι (proj₁ s≤pr-t)) (ι (proj₁ s≤pr-ts)) eq ⟧) 174 | uni∘p₁≡q₁ : (down uni ∘s toSub p₁) ≡s proj₁ s≤pr-t 175 | uni∘p₁≡q₁ = \ S u -> trans (≅-to-≡ (↓↓-comm uni (toSub p₁ S u))) 176 | ([]eq (p₁∘universal≡q₁ {q₁ = ι (proj₁ s≤pr-t)} {q₂ = ι (proj₁ s≤pr-ts)} {eq}) S u) 177 | 178 | prune : ∀ {Sg G D1 D2 T} (i : Inj D1 D2) (t : Tm Sg G D2 T) → 179 | Σ (PrunerSub i t) λ pr → ∀ {G1} (s : Sub< false > Sg G G1) z → ren i z ≡T sub s t → s ≤ PrunerSub.pruner pr 180 | prune i t = (Pr toSub (pruner pr) , prunes pr) , prune-sup i t 181 | where 182 | pr = prune' {i = i} t 183 | 184 | -------------------------------------------------------------------------------- /NOTERM/Unification/Specification.agda: -------------------------------------------------------------------------------- 1 | module NOTERM.Unification.Specification where 2 | 3 | open import Data.Sum renaming (map to map⊎) 4 | open import Data.Sum renaming (inj₁ to yes; inj₂ to no) 5 | 6 | open import Support.Equality 7 | open ≡-Reasoning 8 | 9 | open import Injections 10 | 11 | open import Syntax 12 | 13 | open import Unification.Specification.Decr-Sub public 14 | open import Unification.Injections 15 | 16 | Property<_> : ∀ b Sg G -> Set₁ 17 | Property< b > Sg G = (∀ {G2} -> Sub< b > Sg G G2 -> Set) 18 | 19 | Property : ∀ Sg G -> Set₁ 20 | Property Sg G = ∀ {b} -> Property< b > Sg G 21 | 22 | Unifies : ∀ {Sg G1 D S} (x y : Term Sg G1 D S) -> Property Sg G1 23 | Unifies x y σ = subT σ x ≡T subT σ y 24 | 25 | ∃σ_ : ∀ {Sg G1} -> Property< false > Sg G1 -> Set 26 | ∃σ_ P = ∃ \ G2 -> ∃ \ σ -> P {G2} σ 27 | 28 | ∃σ-pat_ : ∀ {Sg G1} -> Property< true > Sg G1 -> Set 29 | ∃σ-pat_ P = ∃ \ G2 -> ∃ \ σ -> P {G2} σ 30 | 31 | Sup : ∀ {Sg G1} -> Property< false > Sg G1 -> Property< true > Sg G1 32 | Sup P σ = {G' : _} (ρ : _) → P {G'} ρ → ρ ≤ σ 33 | 34 | Max : ∀ {Sg G1} -> Property Sg G1 -> Property< true > Sg G1 35 | Max P σ = P σ × Sup P σ 36 | 37 | Extensional : ∀ {b Sg G} -> Property< b > Sg G -> Set 38 | Extensional P = ∀ {G f g} -> f ≡s g -> P {G} f -> P g 39 | 40 | Spec : ∀ {Sg G1 D S} (x y : Term Sg G1 D S) -> Set 41 | Spec x y = ∃σ-pat Max (Unifies x y) ⊎ ¬ ∃σ Unifies x y 42 | 43 | Unifies-ext : ∀ {Sg G1 D S} (x y : Term Sg G1 D S) -> ∀ {b} -> Extensional {b} (Unifies x y) 44 | Unifies-ext x y f≡g subfx≡subfy rewrite subT-ext f≡g x | subT-ext f≡g y = subfx≡subfy 45 | 46 | ≤-∘ : ∀ {Sg g g1 g2 g3}(ρ : Sub< false > Sg g g1)(p : Sub Sg g g2)(q : Sub Sg g2 g3) -> (ρ≤p : ρ ≤ p) -> proj₁ ρ≤p ≤ q -> ρ ≤ (q ∘s p) 47 | ≤-∘ ρ p q (δ , ρ≡δ∘p) (δ' , δ'≡δ∘q) 48 | = δ' , λ S u → begin ρ S u ≡⟨ ρ≡δ∘p S u ⟩ 49 | sub δ (p S u) ≡⟨ sub-ext δ'≡δ∘q (p S u) ⟩ 50 | subT (δ' ∘s q) (p S u) ≡⟨ sym (Sub∘.subT-∘ (p S u)) ⟩ 51 | subT δ' (subT q (p S u)) ∎ 52 | 53 | 54 | map-Unifies : ∀ {Sg g h h' d t} {σ : Sub Sg g h}{σ' : Sub Sg g h'}-> σ ≤ σ' -> {x y : Term Sg g d t} -> Unifies x y σ' -> Unifies x y σ 55 | map-Unifies {σ = σ} {σ'} (δ , σ≡δ∘σ') {x} {y} σ'Unifies[x,y] = ≡-T 56 | (begin subT σ x ≡⟨ subT-ext σ≡δ∘σ' x ⟩ 57 | subT (δ ∘s σ') x ≡⟨ T-≡ (sandwich subT-∘ (T.cong (subT δ) σ'Unifies[x,y])) ⟩ 58 | subT (δ ∘s σ') y ≡⟨ sym (subT-ext σ≡δ∘σ' y) ⟩ 59 | subT σ y ∎) 60 | 61 | 62 | shift_under_by_ : ∀ {Sg G h1 h2 D T} {xs ys : Term Sg G D T} {σ1 : Sub< false > Sg G h1} 63 | -> Unifies xs ys σ1 -> (σ : Sub Sg G h2) -> σ1 ≤ σ -> ∃σ Unifies (subT σ xs) (subT σ ys) 64 | shift_under_by_ eq σ (δ , σ1≡δ∘σ) = _ , δ , sandwich (λ xs₁ → sym (trans (Sub∘.subT-∘ xs₁) 65 | (sym (subT-ext σ1≡δ∘σ xs₁)))) 66 | eq 67 | 68 | 69 | unify-comm : ∀ {Sg G D T} → (x y : Term Sg G D T) → ∃σ Unifies x y → ∃σ Unifies y x 70 | unify-comm _ _ (G , σ , eq) = (G , σ , T.sym eq) 71 | 72 | spec-comm : ∀ {Sg G D T} → (x y : Term Sg G D T) → Spec x y → Spec y x 73 | spec-comm _ _ = map⊎ (λ {(G , σ , eq , max) → G , σ , T.sym eq , (λ {_} ρ x → max ρ (T.sym x))}) (λ x x₁ → x (unify-comm _ _ x₁)) 74 | 75 | 76 | cong-spec : ∀ {Sg G D D' T T'} → (d : DTm Sg G (D' , T') (D , T)) -> {x y : Term Sg G D T} → Spec x y → Spec (wrapD d x) (wrapD d y) 77 | cong-spec d (inj₁ (_ , σ , unifies , sup)) = inj₁ (_ , (σ , (cong-wrapD d unifies , (λ ρ ρ-unifies → sup ρ (inv-wrapD d ρ-unifies))))) 78 | cong-spec d (inj₂ no-unifier) = inj₂ (λ {(_ , σ , σ-unifies) → no-unifier (_ , (σ , inv-wrapD d σ-unifies)) }) 79 | 80 | optimist : ∀ {Sg m l o D T Ts}(x y : Tm Sg m D T)(xs ys : Tms Sg m D Ts) -> 81 | (p : Sub Sg m o) (q : Sub Sg o l) -> 82 | Max (Unifies x y) p 83 | -> Max (Unifies (subT p xs) (subT p ys)) q 84 | -> Max (Unifies (Tms._∷_ x xs) (y ∷ ys)) (q ∘s p) 85 | 86 | optimist x y xs ys p q ([p]Unifies[x,y] , sup-p) ([q]Unifies[px,py] , sup-q) = 87 | (map-Unifies (q , λ S u → refl) {x} {y} [p]Unifies[x,y] ∷ sandwich subT-∘ [q]Unifies[px,py]) , 88 | (λ {ρ ([ρ]Unifies[x,y] ∷ [ρ]Unifies[xs,ys]) 89 | → let ρ≤p : _ 90 | ρ≤p = sup-p ρ [ρ]Unifies[x,y] 91 | δ : _ 92 | δ = proj₁ ρ≤p 93 | ρ≡δ∘p : _ 94 | ρ≡δ∘p = proj₂ ρ≤p 95 | δ≤q : _ 96 | δ≤q 97 | = sup-q δ 98 | (sandwich (λ x₁ → sym (Sub∘.subT-∘ x₁)) 99 | (Unifies-ext xs ys ρ≡δ∘p [ρ]Unifies[xs,ys])) 100 | in ≤-∘ ρ p q ρ≤p δ≤q}) 101 | 102 | 103 | 104 | ∃σMax[Unifies[x,x]] : ∀ {Sg G D T} (x : Term Sg G D T) -> ∃σ-pat Max (Unifies x x) 105 | ∃σMax[Unifies[x,x]] x = _ , 106 | id-s , 107 | refl-T _ , 108 | (λ ρ x₁ → ρ , (λ S u → sym (ren-id _))) 109 | {- 110 | Spec[xs,ys]⇒Spec[σxs,σys] : ∀ {Sg G G1 D T} {xs ys : Term Sg G D T} (σ : Sub Sg G G1) -> Ctx-length G ≡ Ctx-length G1 -> 111 | IsIso σ -> Spec xs ys -> Spec (subT σ xs) (subT σ ys) 112 | Spec[xs,ys]⇒Spec[σxs,σys] {xs = xs} {ys = ys} σ G~G1 ((δ , id≡δ∘σ) , id≡σ∘δ) (yes (_ , σ₁ , [σ₁]Unifies[xs,ys] , sup-σ₁)) 113 | = yes (_ , σ₁ ∘ds (DS δ , inj₁ (sym G~G1 , (σ , id≡σ∘δ) , id≡δ∘σ)) , [σ₁∘δ]Unifies[σxs,σys] , sup-[σ₁∘δ]) 114 | where 115 | [σ₁∘δ]Unifies[σxs,σys] = sandwich (λ ys → 116 | begin subT ⟦ σ₁ ⟧ ys ≡⟨ cong (subT ⟦ σ₁ ⟧) (trans (sym (subT-id ys)) (subT-ext id≡δ∘σ ys)) ⟩ 117 | subT ⟦ σ₁ ⟧ (subT (δ ∘s σ) ys) ≡⟨ subT-∘ ys ⟩ 118 | subT (⟦ σ₁ ⟧ ∘s (δ ∘s σ)) ys ≡⟨ subT-ext (λ S x → subT-∘ (σ S x)) ys ⟩ 119 | subT ((⟦ σ₁ ⟧ ∘s δ) ∘s σ) ys ≡⟨ sym (subT-∘ ys) ⟩ 120 | subT (⟦ σ₁ ⟧ ∘s δ) (subT σ ys) ∎) 121 | [σ₁]Unifies[xs,ys] 122 | 123 | sup-[σ₁∘δ] : Sup (Unifies (subT σ xs) (subT σ ys)) (⟦ σ₁ ⟧ ∘s δ) 124 | sup-[σ₁∘δ] ρ [ρ]Unifies[σxs,σys] = δ' , λ S u → 125 | begin 126 | ρ S u ≡⟨ sym (left-ids ρ S u) ⟩ 127 | sub ρ (id-s S u) ≡⟨ cong (sub ρ) (id≡σ∘δ S u) ⟩ 128 | sub ρ (sub σ (δ S u)) ≡⟨ Sub∘.subT-∘ {f = ρ} {σ} (δ S u) ⟩ 129 | subT (ρ ∘s σ) (δ S u) ≡⟨ subT-ext (proj₂ ρ∘σ≤σ₁) (δ S u) ⟩ 130 | subT (δ' ∘s ⟦ σ₁ ⟧) (δ S u) ≡⟨ sym (Sub∘.subT-∘ (δ S u)) ⟩ 131 | subT δ' (subT ⟦ σ₁ ⟧ (δ S u)) ∎ 132 | where 133 | ρ∘σ≤σ₁ = sup-σ₁ (ρ ∘s σ) (sandwich Sub∘.subT-∘ [ρ]Unifies[σxs,σys]) 134 | δ' = proj₁ ρ∘σ≤σ₁ 135 | 136 | Spec[xs,ys]⇒Spec[σxs,σys] σ G~G1 _ (no ¬p) = no (λ {(_ , σ₁ , eq) → ¬p (_ , σ₁ ∘s σ , sandwich Sub∘.subT-∘ eq)}) 137 | -} 138 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | Higher-Order Pattern Unification in Agda 2 | ======================================== 3 | 4 | Higher-order unification problems can have multiple incompatible solutions, a challenge for many applications which can't afford to enumerate nor to guess, like type inference. A good strategy is instead to defer any decision except for the cases that fall into the pattern fragment, a subset which guarantees unique solutions. 5 | 6 | This repo provides a formalization in Agda of higher-order pattern unification as defined by Miller in ["A Logic Programming Language with Lambda-Abstraction, Function Variables, and Simple Unification"](http://citeseerx.ist.psu.edu/viewdoc/summary?doi=10.1.1.54.8958). 7 | 8 | The code has been tested with [Agda-2.3.2](http://hackage.haskell.org/package/Agda-2.3.2) using [The Agda standard library](http://wiki.portal.chalmers.se/agda/pmwiki.php?n=Libraries.StandardLibrary), if reloading you get an error related to "refl\`" you need to invalidate the interface file for `Support.Equality' because of [this bug](http://code.google.com/p/agda/issues/detail?id=756) with pattern synonyms. 9 | 10 | The presentation is greatly influenced both by ["Higher-Order Dynamic Pattern Unification for Dependent Types and Records"](http://citeseerx.ist.psu.edu/viewdoc/summary?doi=10.1.1.190.4789) and ["First-Order Unification by Structural Recursion"](http://citeseerx.ist.psu.edu/viewdoc/summary?doi=10.1.1.25.1516) but neither are required to follow the work. 11 | 12 | ["A tutorial implementation of dynamic pattern unification"](https://personal.cis.strath.ac.uk/adam.gundry/pattern-unify/) using Haskell and tackling a much more general version of the problem is a great introduction to the field. 13 | 14 | 15 | The key to read the present work is to identify argument lists satisfying the pattern condition with injective renamings, this point of view and the richly typed structures make it very natural to apply concepts from category theory to design the algorithm and its proof of correctness. 16 | 17 | ### Overview of the modules 18 |
19 | - Definitions and Theorems from Category Theory.
20 |  Category - monos, equalizers, pullbacks
21 | 
22 | - Variable Names as Typed de Bruijn Indexes.
23 |  Vars            - `Γ ∋ τ'
24 |    Vars.MatchTwo - matching against two names in one go
25 |    Vars.SumIso   - `(Γ ++ Δ) ∋ τ' ≈ `Γ ∋ τ ⊎ Δ ∋ τ'
26 | 
27 | - Injective Renamings as a first-order datatype.
28 |  Injections 
29 |    Injections.Type - `Inj Γ Δ' ≈ injective `∀ τ → Γ ∋ τ → Δ ∋ τ'
30 |    Injections.Sum  - `Inj Γ (Γ ++ Δ)' and `Inj Δ (Γ ++ Δ)'  
31 | 
32 | - Meta-Renamings send meta-vars to meta-vars applied to an injective renaming.
33 |  MetaRens - MetaRen Γ Δ 
34 | 
35 | - Lambda Calculus Terms with and without the Pattern condition. 
36 |  Syntax
37 |    Syntax.Type             - `Tm< pat? > Σ Γ Δ τ', `ren' applies renamings `Inj Δ₀ Δ₁`
38 |    Syntax.NbE              - Normalization by Evaluation
39 |      Syntax.NbE.Properties
40 |    Syntax.Sub              - `sub' applies meta-substitutions `Sub< pat? > Σ Γ₀ Γ₁`
41 |    Syntax.Equality         - structural term equality, useful for recursing over its proofs
42 |    Syntax.RenOrn           - `i ⁻¹ t' ≈ `∃ λ s → ren i s ≡ t' as an inductive type
43 |    Syntax.OneHoleContext   - `wrap C t` plugs the term t into the hole of the context C 
44 |    Syntax.No-Cycle         - `¬ ren i t ≡ wrap (d ∷ ds) (ren j t)' i.e. you can't go down 
45 |                               a non-empty context and find the same term you started with
46 | 
47 | - Higher-Order Pattern Unification by Structural Recursion (well, almost).
48 |  Unification                            - `Unify : (s t : Tm Sg G D T) → ∃σ-pat Max (Unifies s t) ⊎ ¬ ∃σ Unifies s t`
49 |    Unification.Specification            - `Max (Unifies s t) σ` ≈ σ is the most general unifier of s and t
50 |      Unification.Specification.Decr-Sub - `DSub Σ Γ Δ' ≈ substitutions decreasing the size of the context, or isomorphisms
51 |    Unification.Injections               - equalizers and pullbacks
52 |    Unification.MetaRens                 - coproducts, coequalizers and pushouts
53 |    Unification.Pruning                  - solving for `ρ` in `ren i z ≡ sub ρ t`
54 |      Unification.Pruning.Epi-Decr       - epic `MetaRen's, like `ρ' above, won't enlarge the context
55 |    Unification.Inversion                - solving for `z` in `ren i z ≡ sub ρ t`
56 |    Unification.OccursCheck              - checking whether a meta-var occurs in a term
57 | 
58 | - Wrappers around the Agda Standard Library.
59 |  Support.Product
60 |  Support.List
61 |  Support.Equality
62 |  Support.EqReasoning
63 | 
64 | -------------------------------------------------------------------------------- /Support/EqReasoning.agda: -------------------------------------------------------------------------------- 1 | module Support.EqReasoning where 2 | 3 | open import Relation.Binary 4 | open import Relation.Binary.PropositionalEquality 5 | import Relation.Binary.HeterogeneousEquality as H 6 | open H using (_≅_) 7 | import Relation.Binary.EqReasoning as EqR 8 | 9 | open EqR public using () renaming (begin_ to begin[_]_) 10 | 11 | module _ {s₁ s₂} {S : Setoid s₁ s₂} where 12 | open EqR S public 13 | hiding (begin_) 14 | 15 | infixr 2 _≅⟨_⟩_ 16 | 17 | _≅⟨_⟩_ : (x : _) {y z : _} → 18 | x ≅ y → y IsRelatedTo z → x IsRelatedTo z 19 | _ ≅⟨ x≅y ⟩ y≡z = _ ≡⟨ H.≅-to-≡ x≅y ⟩ y≡z 20 | 21 | module _ {a} {A : Set a} where 22 | open EqR (setoid A) public 23 | using (begin_) 24 | 25 | -------------------------------------------------------------------------------- /Support/Equality.agda: -------------------------------------------------------------------------------- 1 | module Support.Equality where 2 | 3 | open import Relation.Binary.PropositionalEquality public renaming ([_] to ⌞_⌟) 4 | import Relation.Binary.HeterogeneousEquality 5 | module Het = Relation.Binary.HeterogeneousEquality 6 | open import Relation.Binary.HeterogeneousEquality public using (_≅_ ; _≇_ ; refl; ≅-to-≡; ≡-to-≅) 7 | open import Data.Product 8 | 9 | pattern refl` = (refl , refl) 10 | -------------------------------------------------------------------------------- /Support/List.agda: -------------------------------------------------------------------------------- 1 | module Support.List where 2 | 3 | 4 | open import Data.List public using (List; []; _∷_) 5 | {- 6 | data Fwd (X : Set) : Set where 7 | !> : Fwd X 8 | _:>_ : X → Fwd X → Fwd X 9 | -} 10 | 11 | Fwd : Set → Set 12 | Fwd = List 13 | 14 | !> : ∀ {X : Set} → List X 15 | !> = [] 16 | 17 | _:>_ : ∀ {X : Set} → X → List X → List X 18 | _:>_ x xs = x ∷ xs 19 | 20 | infixr 50 _:>_ 21 | 22 | {- 23 | data Bwd (X : Set) : Set where 24 | ′_) 8 | open import Data.Nat.Properties 9 | 10 | open import Support.Equality 11 | -- _≤′_ _<′_ _≥′_ 12 | infix 4 _>′_ 13 | 14 | data _>′_ : (m : ℕ) → ℕ → Set where 15 | >′-refl : ∀ {m n} (m≡n : m ≡ n) → suc m >′ n 16 | >′-step : ∀ {m n} (m≤′n : m >′ n) → suc m >′ n 17 | {- 18 | _<′_ : Rel ℕ Level.zero 19 | m <′ n = suc m ≤′ n 20 | 21 | _≥′_ : Rel ℕ Level.zero 22 | m ≥′ n = n ≤′ m 23 | 24 | _>′_ : Rel ℕ Level.zero 25 | m >′ n = n <′ m 26 | -} 27 | N>′⇒>′ : ∀ {m n} -> m N.>′ n -> m >′ n 28 | N>′⇒>′ ≤′-refl = >′-refl refl 29 | N>′⇒>′ (≤′-step le) = >′-step (N>′⇒>′ le) 30 | 31 | cast : ∀ {m n o} -> m ≡ n -> n > o -> m >′ o 32 | cast refl le = N>′⇒>′ (≤⇒≤′ le) 33 | 34 | -------------------------------------------------------------------------------- /Support/Product.agda: -------------------------------------------------------------------------------- 1 | module Support.Product where 2 | 3 | open import Data.Product public renaming (map to mapΣ) 4 | 5 | mapΣ₂ : ∀ {a b c p q r} 6 | {A : Set a} {B : Set b} {C : Set c} {P : A → Set p} {Q : B → Set q} {R : C -> Set r} → 7 | (f : A -> B -> C) → (∀ {x y} → P x → Q y -> R (f x y)) → 8 | Σ A P → Σ B Q -> Σ C R 9 | mapΣ₂ f g (x₀ , y₀) (x , y) = (f x₀ x , g y₀ y) 10 | -------------------------------------------------------------------------------- /Syntax.agda: -------------------------------------------------------------------------------- 1 | module Syntax where 2 | 3 | open import Data.Nat hiding (_≤_) 4 | open import Data.Empty 5 | open import Data.Unit hiding (_≤_) 6 | open import Data.Sum 7 | 8 | open import Support.Equality 9 | open import Support.EqReasoning 10 | open import Support.Product public 11 | 12 | open import Injections 13 | 14 | open import Syntax.Type public 15 | open import Syntax.NbE 16 | open import Syntax.NbE.Properties 17 | open import Syntax.Sub public 18 | open import Syntax.Equality public 19 | open import Syntax.RenOrn public 20 | 21 | 22 | congTm : ∀ {b1 b2 Sg G D T G1 D1 T1} -> b1 ≡ b2 -> (f : ∀ {b} -> Term< b > Sg G D T -> Term< b > Sg G1 D1 T1) -> 23 | {x : Term< b1 > Sg G D T}{y : Term< b2 > Sg G D T} -> x ≅ y -> f x ≅ f y 24 | congTm refl f refl = refl 25 | 26 | cong-∷ : ∀ {b1 b2 Sg G D T Ts} -> b1 ≡ b2 -> {x : Tm< b1 > Sg G D T}{y : Tm< b2 > Sg G D T} 27 | {xs : Tms< b1 > Sg G D Ts} {ys : Tms< b2 > Sg G D Ts} -> x ≅ y -> xs ≅ ys -> Tms._∷_ x xs ≅ Tms._∷_ y ys 28 | cong-∷ refl refl refl = refl 29 | 30 | cong-[] : ∀ {b1 b2 Sg G D} -> b1 ≡ b2 -> 31 | let xs : Tms< b1 > Sg G D []; xs = []; ys : Tms< b2 > Sg G D []; ys = [] in xs ≅ ys 32 | cong-[] refl = refl 33 | 34 | 35 | module Subid where 36 | 37 | open import Function using (_∘_) 38 | 39 | x∧true≡x : ∀ {x} -> x ∧ true ≡ x 40 | x∧true≡x {false} = refl 41 | x∧true≡x {true} = refl 42 | 43 | mutual 44 | sub-id : ∀ {b Sg G D T} (t : Tm< b > Sg G D T) → sub id-s t ≅ t 45 | sub-id {b} (con c ts) = congTm (x∧true≡x {b}) (con c) (subs-id ts) 46 | sub-id {b} (var x ts) = congTm (x∧true≡x {b}) (var x) (subs-id ts) 47 | sub-id {b} (lam t) = congTm (x∧true≡x {b}) lam (sub-id t) 48 | sub-id {true} (mvar u j) = ≡-to-≅ (cong (mvar u) (right-id j)) 49 | sub-id {false} (mvar u ts) rewrite ≅-to-≡ (subs-id ts) = ≡-to-≅ (cong (mvar u) (begin 50 | reifys (build (λ z → get (evals ts idEnv) (id-i $ z))) 51 | ≡⟨ cong reifys (begin 52 | build (λ z → get (evals ts idEnv) (id-i $ z)) ≡⟨ build-ext (λ v → cong (get (evals ts idEnv)) (id-i$ v)) ⟩ 53 | build (get (evals ts idEnv)) ≡⟨ build-get (evals ts idEnv) ⟩ 54 | evals ts idEnv ∎) ⟩ 55 | nfs ts idEnv ≡⟨ nfs-id ts ⟩ 56 | ts ∎)) 57 | 58 | subs-id : ∀ {b Sg G D T} (t : Tms< b > Sg G D T) → subs id-s t ≅ t 59 | subs-id {b} [] = cong-[] (x∧true≡x {b}) 60 | subs-id {b} (t ∷ ts) = cong-∷ (x∧true≡x {b}) (sub-id t) (subs-id ts) 61 | 62 | sub-id : ∀ {Sg G D T} (t : Tm Sg G D T) → sub id-s t ≡ t 63 | sub-id t = ≅-to-≡ (Subid.sub-id t) 64 | 65 | subs-id : ∀ {Sg G D T} (t : Tms Sg G D T) → subs id-s t ≡ t 66 | subs-id ts = ≅-to-≡ (Subid.subs-id ts) 67 | 68 | left-ids : ∀ {b Sg G G1} -> (f : Sub< b > Sg G G1) -> (f ∘s id-s) ≡s f 69 | left-ids f S u = ren-id _ 70 | 71 | ∧-assoc : ∀ {a b c} -> (a ∧ b) ∧ c ≡ a ∧ (b ∧ c) 72 | ∧-assoc {true} = refl 73 | ∧-assoc {false} = refl 74 | 75 | module Sub∘ where 76 | mutual 77 | sub-∘ : ∀ {b3 b1 b2 Sg G1 G2 G3 D T} {f : Sub< b1 > Sg G2 G3}{g : Sub< b2 > _ _ _} (t : Tm< b3 > Sg G1 D T) → sub f (sub g t) ≅ sub (f ∘s g) t 78 | sub-∘ {b3} (con c ts) = congTm (∧-assoc {b3}) (con c) (subs-∘ ts) 79 | sub-∘ {b3} (var x ts) = congTm (∧-assoc {b3}) (var x) (subs-∘ ts) 80 | sub-∘ {b3} (lam t) = congTm (∧-assoc {b3}) lam (sub-∘ t) 81 | sub-∘ {true} {g = g} (mvar u j) = ≡-to-≅ (sub-nat (g _ u)) 82 | sub-∘ {false} {f = f} {g = g} (mvar u ts) = 83 | Het.trans (≡-to-≅ (nf-nats {f = f} (evals-nats idEnv-nats (idEnv-∘ idEnv) (subs g ts)) (evals-∘ _ (idEnv-∘ idEnv) _) (g _ u))) 84 | (Het.cong (replace ((f ∘s g) _ u)) (subs-∘ {f = f} {g = g} ts)) 85 | 86 | subs-∘ : ∀ {b3 b1 b2 Sg G1 G2 G3 D T} {f : Sub< b1 > Sg G2 G3}{g : Sub< b2 > _ _ _} (t : Tms< b3 > Sg G1 D T) → subs f (subs g t) ≅ subs (f ∘s g) t 87 | subs-∘ {b3} [] = cong-[] (∧-assoc {b3}) 88 | subs-∘ {b3} (t ∷ t₁) = cong-∷ (∧-assoc {b3}) (sub-∘ t) (subs-∘ t₁) 89 | 90 | subT-∘ : ∀ {Sg G1 G2 G3 D T} {f : Sub< false > Sg G2 G3} {g : Sub< true > Sg _ _} (t : Term< true > Sg G1 D T) 91 | → subT f (subT g t) ≡ subT (f ∘s g) t 92 | subT-∘ {Sg} {G1} {G2} {G3} {D} {inj₁ x} t = ≅-to-≡ (sub-∘ {true} {false} {true} t) 93 | subT-∘ {Sg} {G1} {G2} {G3} {D} {inj₂ y} t = ≅-to-≡ (subs-∘ {true} {false} {true} t) 94 | 95 | sub-∘ : ∀ {Sg G1 G2 G3 D T} {f : Sub Sg G2 G3}{g : Sub _ _ _} (t : Tm Sg G1 D T) → sub f (sub g t) ≡ sub (f ∘s g) t 96 | sub-∘ t = ≅-to-≡ (Sub∘.sub-∘ {true} {true} {true} t) 97 | 98 | subs-∘ : ∀ {Sg G1 G2 G3 D T} {f : Sub Sg G2 G3}{g : Sub _ _ _} (t : Tms Sg G1 D T) → subs f (subs g t) ≡ subs (f ∘s g) t 99 | subs-∘ t = ≅-to-≡ (Sub∘.subs-∘ {true} {true} {true} t) 100 | 101 | 102 | -- Substitution respects pointwise equality. 103 | mutual 104 | sub-ext : ∀ {b2 b1 Sg G1 G2 D T} {f g : Sub< b1 > Sg G1 G2} → f ≡s g → (t : Tm< b2 > Sg G1 D T) → sub f t ≡ sub g t 105 | sub-ext q (con c ts) = cong (con c) (subs-ext q ts) 106 | sub-ext {true} q (mvar u j) = cong (ren j) (q _ u) 107 | sub-ext {false} q (mvar u j) = cong₂ replace (q _ u) (subs-ext q j) 108 | sub-ext q (var x ts) = cong (var x) (subs-ext q ts) 109 | sub-ext q (lam t) = cong lam (sub-ext q t) 110 | 111 | subs-ext : ∀ {b1 b2 Sg G1 G2 D T} {f g : Sub< b1 > Sg G1 G2} → f ≡s g → (t : Tms< b2 > Sg G1 D T) → subs f t ≡ subs g t 112 | subs-ext q [] = refl 113 | subs-ext q (t ∷ ts) = cong₂ _∷_ (sub-ext q t) (subs-ext q ts) 114 | 115 | 116 | 117 | subT-id : ∀ {Sg G D T} (t : Term Sg G D T) → subT id-s t ≡ t 118 | subT-id {Sg} {G} {D} {inj₁ x} t = ≅-to-≡ (Subid.sub-id t) 119 | subT-id {Sg} {G} {D} {inj₂ y} t = ≅-to-≡ (Subid.subs-id t) 120 | 121 | subT-ext : ∀ {b1 b2 Sg G1 G2 D T} {f g : Sub< b1 > Sg G1 G2} → f ≡s g → (t : Term< b2 > Sg G1 D T) → subT f t ≡ subT g t 122 | subT-ext {T = inj₁ x} eq t = sub-ext eq t 123 | subT-ext {T = inj₂ y} eq t = subs-ext eq t 124 | 125 | subT-∘ : ∀ {Sg G1 G2 G3 D T} {f : Sub Sg G2 G3}{g : Sub Sg _ _} (t : Term Sg G1 D T) → subT f (subT g t) ≡ subT (f ∘s g) t 126 | subT-∘ {T = inj₁ _} t = sub-∘ t 127 | subT-∘ {T = inj₂ _} t = subs-∘ t 128 | 129 | 130 | 131 | left-idf : ∀ {Sg G G1} -> (f : Sub< false > Sg G G1) -> (f ∘s id-f) ≡s f 132 | left-idf f S u = begin 133 | eval (f S u) (evals (subs f (reifys idEnv)) idEnv) ≡⟨ eval-ext (f S u) (λ {S} x → begin[ dom ] 134 | get (evals (subs f (reifys idEnv)) idEnv) x ≡⟨ cong (λ ts → get (evals ts idEnv) x) (reifys-nats idEnv-nats) ⟩ 135 | get (evals (reifys idEnv) idEnv) x ≈⟨ RelA-unfold idEnv (idEnv-∘ idEnv) x ⟩ 136 | get idEnv x ∎) ⟩ 137 | nf (f S u) idEnv ≡⟨ nf-id (f S u) ⟩ 138 | f S u ∎ 139 | 140 | mutual 141 | right-idf : ∀ {Sg G D T} -> (t : Tm< false > Sg G D T) -> sub id-f t ≡ t 142 | right-idf (con c ts) = cong (con c) (rights-idf ts) 143 | right-idf (var x ts) = cong (var x) (rights-idf ts) 144 | right-idf (lam t) = cong lam (right-idf t) 145 | right-idf (mvar u ts) rewrite rights-idf ts = cong (mvar u) 146 | (begin reifys (evals (reifys idEnv) (evals ts idEnv)) ≡⟨ reifys-ext (RelA-unfold _ (idEnv-∘ (evals ts idEnv))) ⟩ 147 | reifys (evals ts idEnv) ≡⟨ refl ⟩ 148 | nfs ts idEnv ≡⟨ nfs-id ts ⟩ 149 | ts ∎) 150 | 151 | rights-idf : ∀ {Sg G D T} -> (t : Tms< false > Sg G D T) -> subs id-f t ≡ t 152 | rights-idf [] = refl 153 | rights-idf (x ∷ xs) = cong₂ _∷_ (right-idf x) (rights-idf xs) 154 | 155 | ↓↓ : ∀ {b Sg G D T } -> Tm< b > Sg G D T -> Tm< false > Sg G D T 156 | ↓↓ {true} t = sub id-f t 157 | ↓↓ {false} t = t 158 | 159 | _≡d_ : ∀ {b1 b2 Sg G D T } -> (t1 : Tm< b1 > Sg G D T)(t2 : Tm< b2 > Sg G D T) -> Set 160 | t1 ≡d t2 = ↓↓ t1 ≡ ↓↓ t2 161 | 162 | down : ∀ {b Sg G G1} -> Sub< b > Sg G G1 -> Sub< false > Sg G G1 163 | down {b} s = \ S u -> ↓↓ (s S u) 164 | 165 | ↓↓-nat : ∀ {b Sg G D D1 T } -> (i : Inj D D1) (t : Tm< b > Sg G D T) -> ↓↓ (ren i t) ≡ ren i (↓↓ t) 166 | ↓↓-nat {true} i t = sub-nat t 167 | ↓↓-nat {false} i t = refl 168 | 169 | ↓↓-comm : ∀ {b b1 Sg G G1 D T} -> (s : Sub< b > Sg G G1) -> (t : Tm< b1 > Sg G D T) -> sub (down s) t ≅ ↓↓ (sub s t) 170 | ↓↓-comm {true} {true} s t = Het.sym (Sub∘.sub-∘ {f = id-f} {g = s} t) 171 | ↓↓-comm {true} {false} s t = Het.trans (Het.sym (Sub∘.sub-∘ {f = id-f} {g = s} t)) (≡-to-≅ (right-idf (sub s t))) 172 | ↓↓-comm {false} {true} s t = refl 173 | ↓↓-comm {false} {false} s t = refl 174 | 175 | ↓↓-nat₂ : ∀ {b b1 Sg G G1 D T } -> (s : Sub< b1 > Sg G G1) (t : Tm< b > Sg G D T) -> ↓↓ (sub s t) ≡ sub s (↓↓ t) 176 | ↓↓-nat₂ {true} {true} s t = begin 177 | sub id-f (sub s t) ≡⟨ ≅-to-≡ (Sub∘.sub-∘ {f = id-f} {g = s} t) ⟩ 178 | sub (id-f ∘s s) t ≡⟨ sym (sub-ext (left-idf (id-f ∘s s)) t) ⟩ 179 | sub ((id-f ∘s s) ∘s id-f) t ≡⟨ sym (≅-to-≡ (Sub∘.sub-∘ {f = id-f ∘s s} {g = id-f} t)) ⟩ 180 | sub (id-f ∘s s) (↓↓ t) ≡⟨ ≅-to-≡ (↓↓-comm s (↓↓ t)) ⟩ 181 | ↓↓ (sub s (↓↓ t)) ∎ 182 | ↓↓-nat₂ {true} {false} s t = begin 183 | sub s t ≡⟨ sym (sub-ext (left-idf s) t) ⟩ 184 | sub (s ∘s id-f) t ≡⟨ ≅-to-≡ (Het.sym (Sub∘.sub-∘ {f = s} {g = id-f} t)) ⟩ 185 | sub s (sub id-f t) ∎ 186 | ↓↓-nat₂ {false} {true} s t = refl 187 | ↓↓-nat₂ {false} {false} s t = refl 188 | 189 | sub-extd : ∀ {b1 b2 b3 Sg G D G1 T } -> (s1 : Sub< b1 > Sg G G1)(s2 : Sub< b2 > Sg G G1) -> (t : Tm< b3 > Sg G D T) -> 190 | (∀ S x -> s1 S x ≡d s2 S x) -> sub s1 t ≡d sub s2 t 191 | sub-extd s1 s2 t eq = ≅-to-≡ (Het.trans (Het.sym (↓↓-comm s1 t)) (Het.trans (≡-to-≅ (sub-ext eq t)) (↓↓-comm s2 t))) 192 | 193 | cong-ren : ∀ {b1 b2 Sg G D D1 T } -> (i : Inj D D1) -> {t1 : Tm< b1 > Sg G D T}{t2 : Tm< b2 > Sg G D T} -> 194 | t1 ≡d t2 -> ren i t1 ≡d ren i t2 195 | cong-ren i {t1} {t2} eq = begin 196 | ↓↓ (ren i t1) ≡⟨ ↓↓-nat i t1 ⟩ 197 | ren i (↓↓ t1) ≡⟨ cong (ren i) eq ⟩ 198 | ren i (↓↓ t2) ≡⟨ sym (↓↓-nat i t2) ⟩ 199 | ↓↓ (ren i t2) ∎ 200 | 201 | ren-injd : ∀ {b1 b2 Sg G D D0}{T : Ty} → (i : Inj D D0) → (s : Tm< b1 > Sg G D T) (t : Tm< b2 > Sg G D T) 202 | -> ren i s ≡d ren i t -> s ≡d t 203 | ren-injd i s t eq = ren-inj i (↓↓ s) (↓↓ t) (begin 204 | ren i (↓↓ s) ≡⟨ sym (↓↓-nat i s) ⟩ 205 | ↓↓ (ren i s) ≡⟨ eq ⟩ 206 | ↓↓ (ren i t) ≡⟨ ↓↓-nat i t ⟩ 207 | ren i (↓↓ t) ∎) 208 | 209 | cong-↓↓ : ∀ {b1 b2 Sg G D T} -> {t1 : Tm< b1 > Sg G D T}{t2 : Tm< b2 > Sg G D T} -> 210 | b1 ≡ b2 -> t1 ≅ t2 -> t1 ≡d t2 211 | cong-↓↓ {b} refl eq = cong ↓↓ (≅-to-≡ eq) 212 | 213 | cong-sub : ∀ {b1 b2 b3 Sg G D G1 T } -> (s : Sub< b3 > Sg G G1) -> {t1 : Tm< b1 > Sg G D T}{t2 : Tm< b2 > Sg G D T} -> 214 | t1 ≡d t2 -> sub s t1 ≡d sub s t2 215 | cong-sub s {t1} {t2} eq = begin 216 | ↓↓ (sub s t1) ≡⟨ ↓↓-nat₂ s t1 ⟩ 217 | sub s (↓↓ t1) ≡⟨ cong (sub s) eq ⟩ 218 | sub s (↓↓ t2) ≡⟨ sym (↓↓-nat₂ s t2) ⟩ 219 | ↓↓ (sub s t2) ∎ 220 | 221 | mutual 222 | sub-idf-inj : ∀ {Sg G D T} -> (s t : Tm< true > Sg G D T) -> ↓↓ s ≡T ↓↓ t -> s ≡ t 223 | sub-idf-inj (con c ts) (con c₁ ts₁) (con ceq eq) = cong₂ con ceq (subs-idf-inj ts ts₁ eq) 224 | sub-idf-inj (con c ts) (mvar u j) () 225 | sub-idf-inj (con c ts) (var x ts₁) () 226 | sub-idf-inj (mvar u j) (con c ts) () 227 | sub-idf-inj {Sg} {G} (mvar u j) (mvar u₁ j₁) (mvar ueq x₁) = cong₂ mvar ueq 228 | (ext-$ j j₁ 229 | λ {(Ss ->> B) v → 230 | injective (right# Ss) _ _ 231 | (var-inj₀ (≡-T (begin 232 | var (right# Ss $ (j $ v)) (reifys (mapEnv (left# Ss) idEnv)) ≡⟨ sym (apply-injv Ss B v j) ⟩ 233 | eval (reify (Ss ->> B) (injv (right# Ss $ (j $ v)))) idEnv $$ 234 | mapEnv (left# Ss) idEnv ≡⟨ (cong (λ t → eval t idEnv $$ mapEnv (left# Ss) idEnv) (hip v)) ⟩ 235 | eval (reify (Ss ->> B) (injv (right# Ss $ (j₁ $ v)))) idEnv $$ 236 | mapEnv (left# Ss) idEnv ≡⟨ apply-injv Ss B v j₁ ⟩ 237 | var (right# Ss $ (j₁ $ v)) (reifys (mapEnv (left# Ss) idEnv)) ∎)))}) 238 | where 239 | open import Injections.Sum 240 | pointwise : ∀ {Sg G D T} -> {f g : ∀ {S} (x : T ∋ S) -> Dom Sg G D S} 241 | -> ∀ {S} x -> reifys (build f) ≡T reifys (build g) -> reify S (f x) ≡ reify S (g x) 242 | pointwise zero (t ∷ ts) = T-≡ t 243 | pointwise (suc x) (t ∷ ts) = pointwise x ts 244 | apply-injv : ∀ Ss B v j -> 245 | eval (reify (Ss ->> B) (injv (right# Ss $ (j $ v)))) idEnv $$ mapEnv (left# Ss) idEnv 246 | ≡ var (right# Ss $ (j $ v)) (reifys (mapEnv (left# Ss) idEnv)) 247 | apply-injv Ss B v j = begin 248 | eval (reify (Ss ->> B) (injv (right# Ss $ (j $ v)))) idEnv $$ 249 | mapEnv (left# Ss) idEnv ≡⟨ $$-ext 250 | (Rel-unfold (Ss ->> B) idEnv 251 | (injv-∘ (Ss ->> B) idEnv _ _ 252 | (≡-d (get-build (injv {Sg} {G}) (right# Ss $ (j $ v)))))) 253 | {xs = mapEnv (left# Ss) idEnv} {ys = mapEnv (left# Ss) idEnv} reflA ⟩ 254 | injv (right# Ss $ (j $ v)) $$ mapEnv (left# Ss) idEnv ≡⟨ injv-id (right# Ss $ (j $ v)) (mapEnv (left# Ss) idEnv) ⟩ 255 | var (right# Ss $ (j $ v)) (reifys (mapEnv (left# Ss) idEnv)) ∎ 256 | 257 | hip : ∀ {Ss B} (v : _ ∋ (Ss ->> B)) -> reify _ (injv (right# Ss $ (j $ v))) ≡ reify _ (injv (right# Ss $ (j₁ $ v))) 258 | hip {Ss} v = pointwise {f = \ x -> injv (right# Ss $ (j $ x))} {g = \ x -> injv (right# Ss $ (j₁ $ x))} v 259 | (≡-T (begin 260 | reifys (build (λ x → injv (right# Ss $ (j $ x)))) ≡⟨ reifys∘build∘injv-nat j ⟩ 261 | rens (right# Ss) (rens j (reifys idEnv)) ≡⟨ cong (rens (right# Ss)) x₁ ⟩ 262 | rens (right# Ss) (rens j₁ (reifys idEnv)) ≡⟨ sym (reifys∘build∘injv-nat j₁) ⟩ 263 | reifys (build (λ x₂ → injv (right# Ss $ (j₁ $ x₂)))) ∎)) 264 | where 265 | build∘injv-nat : ∀ j -> build (λ x → injv (right# Ss $ (j $ x))) ≡A mapEnv (right# Ss ∘i j) (build injv) 266 | build∘injv-nat j {S} x = begin[ dom ] 267 | get (build (λ v₁ → injv (right# Ss $ (j $ v₁)))) x ≡⟨ get-build (λ x₂ → injv (right# Ss $ (j $ x₂))) x ⟩ 268 | injv (right# Ss $ (j $ x)) ≡⟨ cong injv (sym (apply-∘ (right# Ss) j)) ⟩ 269 | injv ((right# Ss ∘i j) $ x) ≈⟨ symd S (injv-nat (right# Ss ∘i j) x) ⟩ 270 | mapDom (right# Ss ∘i j) (injv x) ≡⟨ sym (cong (mapDom _) (get-build injv x)) ⟩ 271 | mapDom (right# Ss ∘i j) (get idEnv x) ≡⟨ sym (get-nat-≡ {f = mapDom (right# Ss ∘i j)} idEnv x) ⟩ 272 | get (mapEnv (right# Ss ∘i j) idEnv) x ∎ 273 | reifys∘build∘injv-nat : ∀ j -> reifys (build (λ x → injv (right# Ss $ (j $ x)))) ≡ rens (right# Ss) (rens j (reifys (build injv))) 274 | reifys∘build∘injv-nat j = begin 275 | reifys (build (λ x → injv (right# Ss $ (j $ x)))) ≡⟨ reifys-ext (build∘injv-nat j) ⟩ 276 | reifys (mapEnv (right# Ss ∘i j) idEnv) ≡⟨ reifys-nat reflA (right# Ss ∘i j) ⟩ 277 | rens (right# Ss ∘i j) (reifys idEnv) ≡⟨ rens-∘ (reifys idEnv) ⟩ 278 | rens (right# Ss) (rens j (reifys idEnv)) ∎ 279 | sub-idf-inj (mvar u j) (var x ts) () 280 | sub-idf-inj (var x ts) (con c ts₁) () 281 | sub-idf-inj (var x ts) (mvar u j) () 282 | sub-idf-inj (var x ts) (var x₁ ts₁) (var xeq eq) = cong₂ var xeq (subs-idf-inj ts ts₁ eq) 283 | sub-idf-inj (lam s) (lam t) (lam eq) = cong lam (sub-idf-inj s t eq) 284 | 285 | subs-idf-inj : ∀ {Sg G D T} -> (s t : Tms< true > Sg G D T) -> subs id-f s ≡T subs id-f t -> s ≡ t 286 | subs-idf-inj [] [] _ = refl 287 | subs-idf-inj (s ∷ ss) (t ∷ ts) (teq ∷ tseq) = cong₂ _∷_ (sub-idf-inj s t teq) (subs-idf-inj ss ts tseq) 288 | 289 | ↓↓-inj : ∀ {b Sg G D T } -> {s t : Tm< b > Sg G D T} -> s ≡d t -> s ≡ t 290 | ↓↓-inj {true} eq = sub-idf-inj _ _ (≡-T eq) 291 | ↓↓-inj {false} eq = eq 292 | 293 | open import Syntax.No-Cycle public 294 | open import Syntax.OneHoleContext public 295 | -------------------------------------------------------------------------------- /Syntax/Equality.agda: -------------------------------------------------------------------------------- 1 | module Syntax.Equality where 2 | 3 | open import Data.Sum 4 | 5 | open import Support.Equality renaming (sym to ≡-sym; cong to ≡-cong; trans to ≡-trans) 6 | open import Support.Product 7 | open import Support.List 8 | 9 | open import Injections 10 | open import Syntax.Type 11 | 12 | data _≡T_ {b} {Sg} {G} {D} : {T : Ty ⊎ List Ty} -> (x y : Term< b > Sg G D T) -> Set where 13 | con : ∀ {Ss B}{cx : _ ∋ (Ss ->> B)}{xs}{cy ys} -> cx ≡ cy -> xs ≡T ys -> con cx xs ≡T con cy ys 14 | var : ∀ {Ss B}{x : _ ∋ (Ss ->> B)}{xs}{y ys} -> x ≡ y -> xs ≡T ys -> var x xs ≡T var y ys 15 | lam : ∀ {S Ss B}{tx ty : Tm< b > Sg G (S ∷ D) (Ss ->> B)} -> tx ≡T ty -> lam tx ≡T lam ty 16 | mvar : ∀ {Ss B} {ux uy : G ∋ (B <<- Ss)}{i j} -> ux ≡ uy -> i ≡ j -> mvar ux i ≡T mvar uy j 17 | [] : Tms.[] ≡T [] 18 | _∷_ : ∀ {S Ss}{tx ty : Tm< b > Sg G D S}{tsx tsy : Tms< b > Sg G D Ss} -> tx ≡T ty -> tsx ≡T tsy -> (Tms._∷_ tx tsx) ≡T (ty ∷ tsy) 19 | 20 | 21 | con-inj₁ : ∀ {Sg G D B Ss1 Ss2 b} {x : _ ∋ (Ss1 ->> B)}{y : _ ∋ (Ss2 ->> B)} {xs : Tms< b > Sg G D _}{ys} -> con x xs ≡T con y ys 22 | -> x ≅∋ y 23 | con-inj₁ (con refl eq) = refl` 24 | 25 | var-inj₀ : ∀ {Sg G D Ss B b} {x : _ ∋ (Ss ->> B)}{y : _ ∋ (Ss ->> B)} {xs : Tms< b > Sg G D _}{ys} -> var x xs ≡T var y ys 26 | -> x ≡ y 27 | var-inj₀ (var eq _) = eq 28 | 29 | var-inj₁ : ∀ {Sg G D B Ss1 Ss2 b} {x : _ ∋ (Ss1 ->> B)}{y : _ ∋ (Ss2 ->> B)} {xs : Tms< b > Sg G D _}{ys} -> var x xs ≡T var y ys 30 | -> x ≅∋ y 31 | var-inj₁ (var refl eq) = refl` 32 | 33 | mutual 34 | refl-Tm : ∀ {Sg G D T b} -> (x : Tm< b > Sg G D T) -> x ≡T x 35 | refl-Tm (con c ts) = con refl (refl-Tms ts) 36 | refl-Tm (mvar u j) = mvar refl refl 37 | refl-Tm (var x ts) = var refl (refl-Tms ts) 38 | refl-Tm (lam x) = lam (refl-Tm x) 39 | 40 | refl-Tms : ∀ {Sg G D T b} -> (x : Tms< b > Sg G D T) -> x ≡T x 41 | refl-Tms [] = [] 42 | refl-Tms (t ∷ ts) = refl-Tm t ∷ refl-Tms ts 43 | 44 | refl-T : ∀ {Sg G D T b} -> (x : Term< b > Sg G D T) -> x ≡T x 45 | refl-T {T = inj₁ _} = refl-Tm 46 | refl-T {T = inj₂ _} = refl-Tms 47 | 48 | ≡-T : ∀ {Sg G D T b} -> {x y : Term< b > Sg G D T} -> x ≡ y -> x ≡T y 49 | ≡-T {x = x} eq = subst (λ y → _ ≡T y) eq (refl-T x) 50 | 51 | T-≡ : ∀ {Sg G D T b} -> {x y : Term< b > Sg G D T} -> x ≡T y -> x ≡ y 52 | T-≡ (con refl eq) = ≡-cong (con _) (T-≡ eq) 53 | T-≡ (var refl eq) = ≡-cong (var _) (T-≡ eq) 54 | T-≡ (lam eq) = ≡-cong lam (T-≡ eq) 55 | T-≡ (mvar refl refl) = refl 56 | T-≡ [] = refl 57 | T-≡ (eq ∷ eq₁) = cong₂ _∷_ (T-≡ eq) (T-≡ eq₁) 58 | 59 | module T where 60 | sym : ∀ {Sg G D T b}{x y : Term< b > Sg G D T} -> x ≡T y -> y ≡T x 61 | sym {x = x} {y} eq = ≡-T (≡-sym (T-≡ eq)) 62 | 63 | cong : ∀ {Sg Sg1 G G1 D D1 T T1 b}(f : Term< b > Sg G D T -> Term< b > Sg1 G1 D1 T1){x y : Term< b > Sg G D T} -> x ≡T y -> f x ≡T f y 64 | cong f eq = ≡-T (≡-cong f (T-≡ eq)) 65 | 66 | trans : ∀ {Sg G D T b}{x y z : Term< b > Sg G D T} -> x ≡T y -> y ≡T z -> x ≡T z 67 | trans x≡y y≡z = ≡-T (≡-trans (T-≡ x≡y) (T-≡ y≡z)) 68 | 69 | sandwich : ∀ {a b Sg G1 G2 D T} {f g : Term< a > Sg G1 D T -> Term< b > Sg G2 D T} -> (∀ x -> f x ≡ g x) -> ∀ {x y} -> f x ≡T f y -> g x ≡T g y 70 | sandwich eq {x}{y} p rewrite eq x | eq y = p 71 | 72 | -------------------------------------------------------------------------------- /Syntax/NbE.agda: -------------------------------------------------------------------------------- 1 | module Syntax.NbE where 2 | 3 | open import Relation.Binary.PropositionalEquality 4 | open import Relation.Binary 5 | open import Data.Sum 6 | open import Data.Product 7 | open import Function using (_∘_) 8 | 9 | open import Data.List.All public renaming (map to mapAll) 10 | 11 | open import Support.EqReasoning 12 | 13 | open import Injections 14 | 15 | open import Syntax.Type 16 | 17 | mutual 18 | Dom : (Sg : Ctx)(G : MCtx)(D : Ctx) → Ty → Set 19 | Dom Sg G D ([] ->> B) = Tm< false > Sg G D (! B) 20 | Dom Sg G D ((S ∷ Ss) ->> B) = Σ ((D1 : Ctx) → Inj D D1 → Dom Sg G D1 S → Dom Sg G D1 (Ss ->> B)) 21 | \ x → (D1 : Ctx) → (i : Inj D D1) → {s1 s2 : Dom Sg G D1 S} (s : S ∋ s1 ≡d s2) 22 | → (∀ K k → (Ss ->> B) ∋ x K (k ∘i i) (mapDom k s1) ≡d mapDom k (x D1 i s2)) 23 | 24 | _∋_≡d_ : ∀ {Sg G D} T → (x y : Dom Sg G D T) → Set 25 | ([] ->> _) ∋ x ≡d y = x ≡ y 26 | ((T ∷ Ts) ->> B) ∋ x ≡d y = ∀ D i {s1 s2} (s : T ∋ s1 ≡d s2) → (Ts ->> B) ∋ proj₁ x D i s1 ≡d proj₁ y D i s2 27 | 28 | 29 | symd : ∀ {Sg G D} T → {x y : Dom Sg G D T} → T ∋ x ≡d y → T ∋ y ≡d x 30 | symd ([] ->> B) eq = sym eq 31 | symd ((T ∷ Ts) ->> B) eq = λ D i s → symd (Ts ->> B) (eq D i (symd T s)) 32 | 33 | transd : ∀ {Sg G D} T → {x y z : Dom Sg G D T} → T ∋ x ≡d y → T ∋ y ≡d z → T ∋ x ≡d z 34 | transd ([] ->> B) eq1 eq2 = trans eq1 eq2 35 | transd ((T ∷ Ts) ->> B) eq1 eq2 = λ D i s → transd (Ts ->> B) (eq1 D i s) (eq2 D i (transd _ (symd _ s) s)) 36 | 37 | mapDom : ∀ {Sg G D D1 T} → Inj D D1 → Dom Sg G D T → Dom Sg G D1 T 38 | mapDom {T = [] ->> _} i t = ren i t 39 | mapDom {T = (_ ∷ Ts) ->> _} i (t , t-nat) = (λ D2 j s → t D2 (j ∘i i) s) , 40 | (λ D1 i₁ {s1} {s2} s K k → 41 | subst (λ r → Ts ->> _ ∋ t K r (mapDom k s1) ≡d mapDom k (t D1 (i₁ ∘i i) s2)) 42 | assoc-∘i (t-nat _ (i₁ ∘i i) s K k)) 43 | 44 | mapDom-id : ∀ {Sg G D T} → (d : Dom Sg G D T) → T ∋ mapDom id-i d ≡d d 45 | mapDom-id {T = [] ->> _} d = ren-id d 46 | mapDom-id {T = (T ∷ Ts) ->> _} (d , d-nat) = 47 | λ D i s → 48 | transd _ (≡-d (cong (λ r → d D r _) (right-id i))) 49 | (transd _ (symd _ (mapDom-id (d D i _))) 50 | (transd _ (symd _ (d-nat D i (refld T) _ id-i)) 51 | (transd _ (d-nat D i s _ id-i) 52 | (mapDom-id (d D i _))))) 53 | 54 | refld : ∀ {Sg G D} T → {x : Dom Sg G D T} → T ∋ x ≡d x 55 | refld T {x} = transd T (symd T (mapDom-id x)) (mapDom-id x) 56 | 57 | ≡-d : ∀ {Sg G D T} → {x y : Dom Sg G D T} → x ≡ y → T ∋ x ≡d y 58 | ≡-d refl = refld _ 59 | 60 | mapDom-∘ : ∀ {Sg G D D1 D2 T} → (j : Inj D1 D2)(i : Inj D D1) → (d : Dom Sg G D T) → T ∋ mapDom j (mapDom i d) ≡d mapDom (j ∘i i) d 61 | mapDom-∘ {T = [] ->> _} j i d = sym (ren-∘ d) 62 | mapDom-∘ {T = (S ∷ Ss) ->> _} j i d = λ D i₁ {s1} s → transd (Ss ->> _) (≡-d (cong (λ r → proj₁ d D r s1) (sym assoc-∘i))) 63 | (refld ((S ∷ Ss) ->> _) {d} D (i₁ ∘i (j ∘i i)) s) 64 | 65 | dom : ∀ {Sg G D T} -> Setoid _ _ 66 | dom {Sg} {G} {D} {T} = 67 | record { Carrier = Dom Sg G D T; _≈_ = _∋_≡d_ T; isEquivalence = record { refl = refld _; sym = symd _; trans = transd _ } } 68 | 69 | get : ∀ {A : Set} {P : A → Set} {xs : List A} → All P xs → (∀ {x : A} → xs ∋ x → P x) 70 | get (t ∷ ts) zero = t 71 | get (_ ∷ ts) (suc u) = get ts u 72 | 73 | get-nat-≡ : ∀ {A : Set}{P Q : A → Set}{xs : List A} → {f : ∀ {x} → P x → Q x} → (as : All P xs) → 74 | ∀ {T} (v : _ ∋ T) → get (mapAll f as) v ≡ f (get as v) 75 | get-nat-≡ (a ∷ as) zero = refl 76 | get-nat-≡ {f = f} (a ∷ as) (suc v) = get-nat-≡ {f = f} as v 77 | 78 | _≡A_ : ∀ {Sg G D T} → (x y : All (Dom Sg G D) T) → Set 79 | xs ≡A ys = ∀ {S} (x : _ ∋ S) → S ∋ get xs x ≡d get ys x 80 | 81 | reflA : ∀ {Sg G D T} → {x : All (Dom Sg G D) T} → x ≡A x 82 | reflA {x = x} = \ {S} x -> refld S 83 | 84 | mapDom-ext : ∀ {Sg G D D1 T} → (i : Inj D D1) {x y : Dom Sg G D T} → T ∋ x ≡d y → T ∋ mapDom i x ≡d mapDom i y 85 | mapDom-ext {T = [] ->> _} i eq = cong (ren i) eq 86 | mapDom-ext {T = (_ ∷ _) ->> _} i eq = λ D₁ i₁ s → eq D₁ (i₁ ∘i i) s 87 | 88 | mapEnv : ∀ {Sg G D D1 T} → Inj D D1 → All (Dom Sg G D) T → All (Dom Sg G D1) T 89 | mapEnv i = mapAll (mapDom i) 90 | 91 | mapEnv-id : ∀ {Sg G D T} → {xs : All (Dom Sg G D) T} → mapEnv id-i xs ≡A xs 92 | mapEnv-id {xs = x ∷ xs} zero = mapDom-id x 93 | mapEnv-id {xs = x ∷ xs} (suc v) = mapEnv-id v 94 | 95 | mapEnv-ext : ∀ {Sg G D D1 T} → (i : Inj D D1) → {xs ys : All (Dom Sg G D) T} → xs ≡A ys → mapEnv i xs ≡A mapEnv i ys 96 | mapEnv-ext i {_ ∷ _} {_ ∷ _} eq zero = mapDom-ext i (eq zero) 97 | mapEnv-ext i {_ ∷ _} {_ ∷ _} eq (suc x) = mapEnv-ext i (eq ∘ suc) x 98 | 99 | mapEnv-∘ : ∀ {Sg G D D1 D2 T} → (j : Inj D1 D2)(i : Inj D D1) → (xs : All (Dom Sg G D) T) 100 | → mapEnv j (mapEnv i xs) ≡A mapEnv (j ∘i i) xs 101 | mapEnv-∘ j i (x ∷ xs) zero = mapDom-∘ j i x 102 | mapEnv-∘ j i (x ∷ xs) (suc v) = mapEnv-∘ j i xs v 103 | 104 | 105 | cons-ext : ∀ {Sg G D T Ts} → {x y : Dom Sg G D T} → T ∋ x ≡d y 106 | → {xs ys : All (Dom Sg G D) Ts} → xs ≡A ys → (x ∷ xs) ≡A (y ∷ ys) 107 | cons-ext x xs zero = x 108 | cons-ext x xs (suc v) = xs v 109 | 110 | DomN : (Sg : Ctx)(G : MCtx)(D : Ctx) → List Ty → Base -> Set 111 | DomN Sg G D Ss B = Σ (∀ D1 → (i : Inj D D1) → (xs : All (Dom Sg G D1) Ss) → Tm< false > Sg G D1 (! B)) 112 | \ f -> ∀ D i {xs ys} (eq : xs ≡A ys) K k → f K (k ∘i i) (mapEnv k xs) ≡ ren k (f D i ys) 113 | _∋_≡N_ : ∀ {Sg G D T} Ts (f g : DomN Sg G D Ts T) → Set 114 | _∋_≡N_ {Sg} {G} {D} {T} Ts f g = (∀ D1 (i : Inj D D1) {xs ys} -> xs ≡A ys -> proj₁ f D1 i xs ≡ proj₁ g D1 i ys) 115 | 116 | reflN : ∀ {Sg G D T Ts} (f : DomN Sg G D Ts T) -> Ts ∋ f ≡N f 117 | reflN (f , f-nat) D i {xs} {ys} eq = begin 118 | f D i xs ≡⟨ sym (ren-id _) ⟩ 119 | ren id-i (f D i xs) ≡⟨ sym (f-nat D i reflA D id-i) ⟩ 120 | f D (id-i ∘i i) (mapEnv id-i xs) ≡⟨ f-nat D i eq D id-i ⟩ 121 | ren id-i (f D i ys) ≡⟨ ren-id _ ⟩ 122 | f D i ys ∎ 123 | 124 | mapDomN : ∀ {Sg G D Ts T K} (i : Inj D K) -> DomN Sg G D Ts T -> DomN Sg G K Ts T 125 | mapDomN i (f , f-nat) = (λ D1 j xs → f D1 (j ∘i i) xs) 126 | , (λ D i₁ {xs} {ys} eq K k → begin[ dom ] 127 | f K ((k ∘i i₁) ∘i i) (mapEnv k xs) ≈⟨ cong₂ (f K) (sym assoc-∘i) refl ⟩ 128 | f K (k ∘i (i₁ ∘i i)) (mapEnv k xs) ≈⟨ f-nat _ (i₁ ∘i i) eq K k ⟩ 129 | ren k (f D (i₁ ∘i i) ys) ∎) 130 | 131 | applyN : ∀ {Sg G D S Ss T K} -> DomN Sg G D (S ∷ Ss) T -> Inj D K -> Dom Sg G K S -> DomN Sg G K Ss T 132 | applyN (f , f-nat) i s = (λ D2 j xs → f D2 (j ∘i i) (mapDom j s ∷ xs)) , 133 | (λ D₁ i₁ {xs} {ys} eq K k → begin[ dom ] 134 | f K ((k ∘i i₁) ∘i i) (mapDom (k ∘i i₁) s ∷ mapEnv k xs) ≈⟨ reflN (f , f-nat) K ((k ∘i i₁) ∘i i) 135 | (cons-ext (symd _ (mapDom-∘ k i₁ s)) (mapEnv-ext k reflA)) ⟩ 136 | f K ((k ∘i i₁) ∘i i) (mapEnv k (mapDom i₁ s ∷ xs)) ≈⟨ cong₂ (f K) (sym assoc-∘i) refl ⟩ 137 | f K (k ∘i (i₁ ∘i i)) (mapEnv k (mapDom i₁ s ∷ xs)) ≈⟨ f-nat _ (i₁ ∘i i) {mapDom i₁ s ∷ xs} {mapDom i₁ s ∷ ys} 138 | (cons-ext (mapDom-ext i₁ (refld _)) eq) _ k ⟩ 139 | ren k (f D₁ (i₁ ∘i i) (mapDom i₁ s ∷ ys)) ∎) 140 | 141 | applyN-nat : ∀ {Sg G D S Ss T D1} (f : DomN Sg G D (S ∷ Ss) T) (i : Inj D D1) -> 142 | ∀ {s1 s2 : Dom Sg G D1 S} (_ : S ∋ s1 ≡d s2) -> 143 | ∀ {K} (k : Inj D1 K) -> Ss ∋ applyN f (k ∘i i) (mapDom k s1) ≡N mapDomN k (applyN f i s2) 144 | applyN-nat (f , f-nat) i {s1} {s2} s {K} k D i₁ {xs} {ys} xs≡ys = begin 145 | f D (i₁ ∘i (k ∘i i)) (mapDom i₁ (mapDom k s1) ∷ xs) ≈⟨ cong₂ (f D) assoc-∘i refl ⟩ 146 | f D ((i₁ ∘i k) ∘i i) (mapDom i₁ (mapDom k s1) ∷ xs) ≈⟨ (reflN (f , f-nat) D ((i₁ ∘i k) ∘i i) 147 | (cons-ext (transd _ (mapDom-∘ i₁ k s1) (mapDom-ext (i₁ ∘i k) s)) xs≡ys)) ⟩ 148 | f D ((i₁ ∘i k) ∘i i) (mapDom (i₁ ∘i k) s2 ∷ ys) ∎ 149 | 150 | 151 | mutual 152 | 153 | expand : ∀ {Sg G D T} Ts → DomN Sg G D Ts T → Dom Sg G D (Ts ->> T) 154 | expand [] (f , f-nat) = f _ id-i [] 155 | expand (S ∷ Ts) (f , f-nat) = (λ D1 i s → expand Ts (applyN (f , f-nat) i s)) , 156 | (λ D1 i {s1} {s2} s K k → begin[ dom ] 157 | expand Ts (applyN (f , f-nat) (k ∘i i) (mapDom k s1)) ≈⟨ expand-ext Ts (applyN-nat (f , f-nat) i s k) ⟩ 158 | expand Ts (mapDomN k (applyN (f , f-nat) i s2)) ≈⟨ expand-nat Ts (applyN (f , f-nat) i s2) (applyN (f , f-nat) i s2) 159 | (reflN (applyN (f , f-nat) i s2)) K k ⟩ 160 | mapDom k (expand Ts (applyN (f , f-nat) i s2)) ∎) 161 | 162 | injv : ∀ {Sg G D T} → D ∋ T → Dom Sg G D T 163 | injv {T = Ss ->> B} v = expand Ss ((λ D1 i xs → var (i $ v) (reifys xs)) , 164 | (λ D₁ i eq K k → cong₂ var (apply-∘ k i) 165 | (reifys-nat eq k))) 166 | 167 | reify : ∀ {Sg G D}(T : Ty) → Dom Sg G D T → Tm< false > Sg G D T 168 | reify ([] ->> B) x = x 169 | reify ((S ∷ Ss) ->> B) f = lam (reify (Ss ->> B) (proj₁ f _ (weak id-i) (injv zero))) 170 | 171 | reifys : ∀ {Sg G D Ts} → All (Dom Sg G D) Ts → Tms< false > Sg G D Ts 172 | reifys [] = [] 173 | reifys (t ∷ ts) = reify _ t ∷ reifys ts 174 | 175 | expand-nat : ∀ {Sg G D T} Ts (f g : DomN Sg G D Ts T) → (∀ D1 (i : Inj D D1) {xs ys} -> xs ≡A ys -> proj₁ f D1 i xs ≡ proj₁ g D1 i ys) -> 176 | ∀ (K : List Ty) (k : Inj D K) → (Ts ->> T) ∋ expand Ts (mapDomN k f) ≡d mapDom k (expand Ts g) 177 | expand-nat [] (f , f-nat) (g , g-nat) eq K k = begin 178 | f K (id-i ∘i k) [] ≡⟨ eq K (id-i ∘i k) reflA ⟩ 179 | g K (id-i ∘i k) [] ≡⟨ cong₂ (g K) (trans (left-id k) (sym (right-id k))) refl ⟩ 180 | g K (k ∘i id-i) [] ≡⟨ g-nat _ id-i reflA _ k ⟩ 181 | ren k (g _ id-i []) ∎ 182 | expand-nat (T ∷ Ts) (f , _) (g , _) eq K k = λ D i {s1} {s2} s → expand-ext Ts 183 | (λ D₁ i₁ {xs} {ys} xs≡ys → begin[ dom ] 184 | f D₁ ((i₁ ∘i i) ∘i k) (mapDom i₁ s1 ∷ xs) ≈⟨ eq D₁ ((i₁ ∘i i) ∘i k) (cons-ext (mapDom-ext i₁ s) xs≡ys) ⟩ 185 | g D₁ ((i₁ ∘i i) ∘i k) (mapDom i₁ s2 ∷ ys) ≈⟨ cong₂ (g D₁) (sym assoc-∘i) refl ⟩ 186 | g D₁ (i₁ ∘i (i ∘i k)) (mapDom i₁ s2 ∷ ys) ∎) 187 | 188 | expand-ext : ∀ {Sg G D T} Ts {f g : DomN Sg G D Ts T} 189 | → (Ts ∋ f ≡N g) → (Ts ->> T) ∋ expand Ts f ≡d expand Ts g 190 | expand-ext [] eq = eq _ id-i {[]} {[]} (λ ()) 191 | expand-ext (S ∷ Ts) eq = λ D1 i s → expand-ext Ts (λ D i₁ x → eq D (i₁ ∘i i) (cons-ext (mapDom-ext i₁ s) x)) 192 | 193 | 194 | injv-nat : ∀ {Sg G D D1}(i : Inj D D1){T}(v : D ∋ T) → T ∋ mapDom i (injv v) ≡d (injv {Sg} {G} (i $ v)) 195 | injv-nat i {[] ->> B} v = cong (\ v -> var v []) (begin 196 | i $ (id-i $ v) ≡⟨ cong (_$_ i) (id-i$ v) ⟩ 197 | i $ v ≡⟨ sym (id-i$ (i $ v)) ⟩ 198 | id-i $ (i $ v) ∎) 199 | injv-nat i {(S ∷ Ss) ->> B} v = λ D i₁ s → expand-ext Ss (λ D₁ i₂ x → 200 | cong₂ var (begin (i₂ ∘i (i₁ ∘i i)) $ v ≡⟨ cong (λ i₃ → i₃ $ v) assoc-∘i ⟩ 201 | ((i₂ ∘i i₁) ∘i i) $ v ≡⟨ apply-∘ (i₂ ∘i i₁) i ⟩ 202 | (i₂ ∘i i₁) $ (i $ v) ∎) 203 | (reifys-ext (cons-ext (mapDom-ext i₂ s) x))) 204 | 205 | injv-ext : ∀ {Sg G D T} → (v : D ∋ T) → T ∋ injv {Sg} {G} v ≡d injv v 206 | injv-ext {Sg} {G} {D} {Ss ->> B} v = expand-ext Ss (λ D₁ i x → cong (var _) (reifys-ext x)) 207 | 208 | reify-nat : ∀ {Sg G D} T {x y : Dom Sg G D T} → T ∋ x ≡d y → ∀ {K} (k : Inj _ K) → reify T (mapDom k x) ≡ ren k (reify T y) 209 | reify-nat ([] ->> _) eq k = cong (ren k) eq 210 | reify-nat ((S ∷ Ss) ->> _) {x , x-nat} {y , y-nat} eq k 211 | = cong lam (begin 212 | reify (Ss ->> _) (x _ (weak id-i ∘i k) (injv (cons k $ zero))) ≡⟨ reify-ext (Ss ->> _) lemma ⟩ 213 | reify (Ss ->> _) (mapDom (cons k) (y (S ∷ _) (weak id-i) (injv zero))) ≡⟨ reify-nat (Ss ->> _) (refld (Ss ->> _)) (cons k) ⟩ 214 | ren (cons k) (reify (Ss ->> _) (y (S ∷ _) (weak id-i) (injv zero))) ∎) 215 | where 216 | lemma = begin[ dom ] 217 | x _ (weak id-i ∘i k) (injv (cons k $ zero)) 218 | ≡⟨ cong (λ r → x _ r (injv (cons k $ zero))) (ext-∘ (λ _ v → begin 219 | weak id-i $ (k $ v) ≡⟨ iso2 _ _ (k $ v) ⟩ 220 | suc (id-i $ (k $ v)) ≡⟨ cong suc (id-i$ (k $ v)) ⟩ 221 | thin zero _ (k $ v) ≡⟨ sym (iso2 _ _ v) ⟩ 222 | weak k $ v ≡⟨ cong (_$_ (weak k)) (sym (id-i$ v)) ⟩ 223 | cons k $ thin zero _ (id-i $ v) ≡⟨ cong (_$_ (cons k)) (sym (iso2 _ _ v)) ⟩ 224 | cons k $ (weak id-i $ v) ∎)) ⟩ 225 | x _ (cons k ∘i weak id-i) (injv (cons k $ zero)) ≈⟨ eq _ (cons k ∘i weak id-i) (symd _ (injv-nat (cons k) zero)) ⟩ 226 | y _ (cons k ∘i weak id-i) (mapDom (cons k) (injv zero)) ≈⟨ y-nat (S ∷ _) (weak id-i) (injv-ext zero) (_ <: S) (cons k) ⟩ 227 | mapDom (cons k) (y _ (weak id-i) (injv zero)) ∎ 228 | 229 | reify-ext : ∀ {Sg G D} T → {xs ys : Dom Sg G D T} → T ∋ xs ≡d ys → reify T xs ≡ reify T ys 230 | reify-ext ([] ->> _) eq = eq 231 | reify-ext ((S ∷ Ss) ->> B) eq = cong lam (reify-ext (Ss ->> B) (eq _ (weak id-i) (injv-ext zero))) 232 | 233 | reifys-nat : ∀ {Sg G D T} {x y : All (Dom Sg G D) T} → x ≡A y → ∀ {K} (k : Inj _ K) → reifys (mapEnv k x) ≡ rens k (reifys y) 234 | reifys-nat {x = []} {[]} eq k = refl 235 | reifys-nat {x = x ∷ xs} {y ∷ ys} eq k = cong₂ _∷_ (reify-nat _ (eq zero) k) (reifys-nat (eq ∘ suc) k) 236 | 237 | reifys-ext : ∀ {Sg G D Ts} → {xs ys : All (Dom Sg G D) Ts} → xs ≡A ys → reifys xs ≡ reifys ys 238 | reifys-ext {xs = []} {[]} eq = refl 239 | reifys-ext {xs = x ∷ xs} {y ∷ ys} eq = cong₂ _∷_ (reify-ext _ (eq zero)) (reifys-ext (eq ∘ suc)) 240 | 241 | 242 | build : ∀ {A : Set} {P : A → Set} {xs : List A} → (∀ {x : A} → xs ∋ x → P x) → All P xs 243 | build {A} {P} {[]} f = [] 244 | build {A} {P} {x ∷ xs} f = f zero ∷ build (λ x₁ → f (suc x₁)) 245 | 246 | build-ext : ∀ {A : Set} {P : A → Set} {xs : List A} 247 | → {f g : ∀ {x : A} → xs ∋ x → P x} → (∀ {x : A} (v : xs ∋ x) → f v ≡ g v) → build f ≡ build g 248 | build-ext {xs = []} eq = refl 249 | build-ext {xs = x ∷ xs} eq = cong₂ _∷_ (eq zero) (build-ext (\ v → eq (suc v))) 250 | 251 | get-build : ∀ {A : Set} {P : A → Set} {xs : List A} → (f : ∀ {x : A} → xs ∋ x → P x) → 252 | ∀ {T} (x : _ ∋ T) → get (build f) x ≡ f x 253 | get-build f zero = refl 254 | get-build f (suc u) = get-build (λ u₁ → f (suc u₁)) u 255 | 256 | build-get : ∀ {A : Set} {P : A → Set} {xs : List A} → (f : All P xs) 257 | → (build (\ x → get f x)) ≡ f 258 | build-get [] = refl 259 | build-get (x ∷ xs) = cong₂ _∷_ refl (build-get xs) 260 | 261 | mapAll-build : ∀ {A} xs {P Q : A → Set} (f : (∀ {x : A} → xs ∋ x → P x)) (g : ∀ {x} → P x → Q x) → mapAll g (build f) ≡ build (g ∘ f) 262 | mapAll-build [] f g = refl 263 | mapAll-build (x ∷ xs) f g = cong₂ _∷_ refl (mapAll-build xs (f ∘ suc) g) 264 | 265 | _$$_ : ∀ {Sg G D Ts B} → Dom Sg G D (Ts ->> B) → All (Dom Sg G D) Ts → Dom Sg G D (! B) 266 | f $$ [] = f 267 | f $$ (px ∷ xs₁) = proj₁ f _ id-i px $$ xs₁ 268 | 269 | $$-ext : ∀ {Sg G D Ts B} → {x y : Dom Sg G D (Ts ->> B)} → _ ∋ x ≡d y → 270 | ∀ {xs ys : All (Dom Sg G D) Ts} → xs ≡A ys → (! B) ∋ (x $$ xs) ≡d (y $$ ys) 271 | $$-ext f {[]} {[]} eq = f 272 | $$-ext f {x ∷ xs} {y ∷ ys} eq = $$-ext (f _ id-i (eq zero)) (eq ∘ suc) 273 | 274 | $$-nat : ∀ {Sg G D Ts B} → {x y : Dom Sg G D (Ts ->> B)} → _ ∋ x ≡d y → 275 | ∀ {xs ys : All (Dom Sg G D) Ts} → xs ≡A ys → ∀ {K} (k : Inj _ K) → (! B) ∋ (mapDom k x $$ mapEnv k xs) ≡d ren k (y $$ ys) 276 | $$-nat f≡g {[]} {[]} eq k = cong (ren k) f≡g 277 | $$-nat {Ts = T ∷ Ts} {B} {x = f , _} {g , g-nat} f≡g {x ∷ xs} {y ∷ ys} eq k = begin 278 | f _ (id-i ∘i k) (mapDom k x) $$ mapEnv k xs ≡⟨ $$-ext {Ts = Ts} {B = B} {x = f _ (id-i ∘i k) (mapDom k x)} 279 | {mapDom k (g _ id-i y)} lemma 280 | {xs = mapEnv k xs} {ys = mapEnv k ys} 281 | (λ x₁ → mapEnv-ext k (eq ∘ suc) x₁) ⟩ 282 | mapDom k (g _ id-i y) $$ mapEnv k ys ≡⟨ $$-nat {Ts = Ts} {B = B} (refld _) {ys} {ys} reflA k ⟩ 283 | ren k (g _ id-i y $$ ys) ∎ 284 | where 285 | lemma = begin[ dom ] 286 | f _ (id-i ∘i k) (mapDom k x) ≡⟨ cong₂ (f _) (trans (left-id k) (sym (right-id k))) refl ⟩ 287 | f _ (k ∘i id-i) (mapDom k x) ≈⟨ f≡g _ (k ∘i id-i) (mapDom-ext k (eq zero)) ⟩ 288 | g _ (k ∘i id-i) (mapDom k y) ≈⟨ g-nat _ id-i (refld T) _ k ⟩ 289 | mapDom k (g _ id-i y) ∎ 290 | 291 | idEnv : ∀ {Sg G D} → All (Dom Sg G D) D 292 | idEnv = build injv 293 | 294 | get-nat : ∀ {Sg G D D1 T} {xs ys : All (Dom Sg G D1) D} → xs ≡A ys → (v : _ ∋ T) → 295 | ∀ {K} (k : Inj D1 K) → T ∋ get (mapEnv k xs) v ≡d mapDom k (get ys v) 296 | get-nat {xs = x ∷ xs} {y ∷ ys} eq zero k = mapDom-ext k (eq zero) 297 | get-nat {xs = x ∷ xs} {y ∷ ys} eq (suc v) k = get-nat (eq ∘ suc) v k 298 | 299 | mutual 300 | nf : ∀ {b Sg G D D1 T} → Tm< b > Sg G D T → All (Dom Sg G D1) D → Tm< false > Sg G D1 T 301 | nf t g = reify _ (eval t g) 302 | 303 | nfs : ∀ {b Sg G D D1 T} → Tms< b > Sg G D T → All (Dom Sg G D1) D → Tms< false > Sg G D1 T 304 | nfs ts g = reifys (evals ts g) 305 | 306 | nfAs : ∀ {b Sg G D D1 T} → Args b Sg G D T → All (Dom Sg G D1) D → Tms< false > Sg G D1 T 307 | nfAs as g = reifys (evalAs as g) 308 | 309 | eval : ∀ {b Sg G D D1 T} → Tm< b > Sg G D T → All (Dom Sg G D1) D → Dom Sg G D1 T 310 | eval (con c ts) g = con c (nfs ts g) 311 | eval (mvar u j) g = mvar u (nfAs j g) 312 | eval (var x ts) g = get g x $$ evals ts g 313 | eval (lam t) g = (λ D2 i s → eval t (s ∷ mapEnv i g)) , 314 | (λ D1 i {s1} {s2} s K k → begin[ dom ] 315 | eval t (mapDom k s1 ∷ mapEnv (k ∘i i) g) ≈⟨ eval-ext t (cons-ext (mapDom-ext k s) (λ x → symd _ (mapEnv-∘ k i _ x))) ⟩ 316 | eval t (mapDom k s2 ∷ mapEnv k (mapEnv i g)) ≈⟨ eval-nat t reflA k ⟩ 317 | mapDom k (eval t (s2 ∷ mapEnv i g)) ∎) 318 | 319 | evals : ∀ {b Sg G D D1 Ts} → Tms< b > Sg G D Ts → All (Dom Sg G D1) D → All (Dom Sg G D1) Ts 320 | evals [] g = [] 321 | evals (t ∷ ts) g = eval t g ∷ evals ts g 322 | 323 | evalAs : ∀ {b Sg G D D1 Ts} → Args b Sg G D Ts → All (Dom Sg G D1) D → All (Dom Sg G D1) Ts 324 | evalAs {true} i g = build (λ x₁ → get g (i $ x₁)) 325 | evalAs {false} xs g = evals xs g 326 | 327 | eval-nat : ∀ {b Sg G D D1 T} (t : Tm< b > Sg G D T) {xs ys : All (Dom Sg G D1) D} → xs ≡A ys -> 328 | {K : List Ty} (k : Inj D1 K) → T ∋ eval t (mapEnv k xs) ≡d mapDom k (eval t ys) 329 | eval-nat (con c ts) {xs} {ys} eq k = cong (con c) (nfs-nat ts eq k) 330 | eval-nat (mvar u j) {xs} {ys} eq k = cong (mvar u) (nfAs-nat j eq k) 331 | eval-nat (var x ts) {xs} {ys} eq k = begin[ dom ] 332 | get (mapEnv k xs) x $$ evals ts (mapEnv k xs) ≈⟨ $$-ext {x = get (mapEnv k xs) x} {mapDom k (get ys x)} (get-nat eq x k) 333 | {evals ts (mapEnv k xs)} {mapEnv k (evals ts ys)} (evals-nat ts eq k) ⟩ 334 | mapDom k (get ys x) $$ mapEnv k (evals ts ys) ≈⟨ $$-nat {x = get ys x} (refld _) {evals ts ys} reflA k ⟩ 335 | ren k (get ys x $$ evals ts ys) ∎ 336 | eval-nat (lam t) eq k = λ D i₁ s₁ → eval-ext t (cons-ext s₁ (λ x → transd _ (mapEnv-∘ i₁ k _ x) (mapEnv-ext (i₁ ∘i k) eq x))) 337 | 338 | evals-nat : ∀ {b Sg G D D1 T} (t : Tms< b > Sg G D T) {xs ys : All (Dom Sg G D1) D} → xs ≡A ys -> 339 | {K : List Ty} (k : Inj D1 K) → evals t (mapEnv k xs) ≡A mapEnv k (evals t ys) 340 | evals-nat (t ∷ ts) eq k zero = eval-nat t eq k 341 | evals-nat (t ∷ ts) eq k (suc v) = evals-nat ts eq k v 342 | 343 | evalAs-nat : ∀ {b Sg G D D1 T} (t : Args b Sg G D T) {xs ys : All (Dom Sg G D1) D} → xs ≡A ys -> 344 | {K : List Ty} (k : Inj D1 K) → evalAs t (mapEnv k xs) ≡A mapEnv k (evalAs t ys) 345 | evalAs-nat {true} (u ∷ i [ _ ]) eq k zero = get-nat eq u k 346 | evalAs-nat {true} (u ∷ i [ _ ]) eq k (suc v) = evalAs-nat i eq k v 347 | evalAs-nat {false} ts eq k v = evals-nat ts eq k v 348 | 349 | nfs-nat : ∀ {b Sg G D D1 T} → (ts : Tms< b > Sg G D T) → {g1 g2 : All (Dom Sg G D1) D} → g1 ≡A g2 → 350 | ∀ {K} (k : Inj _ K) → nfs ts (mapEnv k g1) ≡ rens k (nfs ts g2) 351 | nfs-nat ts {xs} {ys} eq k = begin 352 | reifys (evals ts (mapEnv k xs)) ≡⟨ reifys-ext (evals-nat ts eq k) ⟩ 353 | reifys (mapEnv k (evals ts ys)) ≡⟨ reifys-nat reflA k ⟩ 354 | rens k (reifys (evals ts ys)) ∎ 355 | 356 | nfAs-nat : ∀ {b Sg G D D1 T} → (ts : Args b Sg G D T) → {g1 g2 : All (Dom Sg G D1) D} → g1 ≡A g2 → 357 | ∀ {K} (k : Inj _ K) → nfAs ts (mapEnv k g1) ≡ rens k (nfAs ts g2) 358 | nfAs-nat j {xs} {ys} eq k = begin 359 | reifys (evalAs j (mapEnv k xs)) ≡⟨ reifys-ext (evalAs-nat j eq k) ⟩ 360 | reifys (mapEnv k (evalAs j ys)) ≡⟨ reifys-nat reflA k ⟩ 361 | rens k (reifys (evalAs j ys)) ∎ 362 | 363 | eval-ext : ∀ {b Sg G D D1 T} → (t : Tm< b > Sg G D T) → {g1 g2 : All (Dom Sg G D1) D} → 364 | g1 ≡A g2 → T ∋ eval t g1 ≡d eval t g2 365 | eval-ext (con c ts) eq = cong (con c) (nfs-ext ts eq) 366 | eval-ext (mvar u j) eq = cong (mvar u) (nfAs-ext j eq) 367 | eval-ext (var x ts) eq = $$-ext (eq x) (evals-ext ts eq) 368 | eval-ext (lam t) eq = λ D i {s1} {s2} s → eval-ext t (cons-ext s (mapEnv-ext i eq)) 369 | 370 | evals-ext : ∀ {b Sg G D D1 T} → (t : Tms< b > Sg G D T) → {g1 g2 : All (Dom Sg G D1) D} → 371 | g1 ≡A g2 → evals t g1 ≡A evals t g2 372 | evals-ext (t ∷ ts) eq zero = eval-ext t eq 373 | evals-ext (t ∷ ts) eq (suc x) = evals-ext ts eq x 374 | 375 | evalAs-ext : ∀ {b Sg G D D1 T} → (t : Args b Sg G D T) → {g1 g2 : All (Dom Sg G D1) D} → g1 ≡A g2 → evalAs t g1 ≡A evalAs t g2 376 | evalAs-ext {true} (i ∷ t [ pf ]) eq zero = eq i 377 | evalAs-ext {true} (i ∷ t [ pf ]) eq (suc x) = evalAs-ext t eq x 378 | evalAs-ext {false} t eq x = evals-ext t eq x 379 | 380 | nf-ext : ∀ {b Sg G D D1 T} → (t : Tm< b > Sg G D T) → {g1 g2 : All (Dom Sg G D1) D} → g1 ≡A g2 → nf t g1 ≡ nf t g2 381 | nf-ext t g = reify-ext _ (eval-ext t g) 382 | 383 | nfs-ext : ∀ {b Sg G D D1 T} → (t : Tms< b > Sg G D T) → {g1 g2 : All (Dom Sg G D1) D} → g1 ≡A g2 → nfs t g1 ≡ nfs t g2 384 | nfs-ext t g = reifys-ext (evals-ext t g) 385 | 386 | nfAs-ext : ∀ {b Sg G D D1 T} → (t : Args b Sg G D T) → {g1 g2 : All (Dom Sg G D1) D} → g1 ≡A g2 → nfAs t g1 ≡ nfAs t g2 387 | nfAs-ext t g = reifys-ext (evalAs-ext t g) 388 | 389 | replace : ∀ {b Sg G D D1 T} → Tm< b > Sg G D T → Tms< false > Sg G D1 D → Tm< false > Sg G D1 T 390 | replace t ts = nf t (evals ts idEnv) 391 | 392 | 393 | 394 | _≤[_]_ : ∀ {Sg G D D1 D2} (g : All (Dom Sg G D1) D) (j : Inj D D2) (g1 : All (Dom Sg G D1) D2) → Set 395 | g ≤[ j ] g2 = ∀ {S} (x : _ ∋ S) → S ∋ get g x ≡d get g2 (j $ x) 396 | 397 | ≤[]-∘ : ∀ {Sg G B C} → (f : All (Dom Sg G C) B) -> 398 | ∀ {D} (i : Inj B D) {K} (j : Inj C K) → ∀ (f' : All (Dom Sg G K) D) 399 | → mapEnv j f ≤[ i ] f' → ∀ {D'} (i' : Inj D D') {K'} (j' : Inj K K') → ∀ (f'' : All (Dom Sg G K') D') 400 | → mapEnv j' f' ≤[ i' ] f'' → mapEnv (j' ∘i j) f ≤[ (i' ∘i i) ] f'' 401 | ≤[]-∘ f i j f' ex i' j' f'' ex' {S} x = begin[ dom ] 402 | get (mapEnv (j' ∘i j) f) x ≈⟨ symd S (mapEnv-∘ j' j f x) ⟩ 403 | get (mapEnv j' (mapEnv j f)) x ≈⟨ get-nat reflA x j' ⟩ 404 | mapDom j' (get (mapEnv j f) x) ≈⟨ mapDom-ext j' (ex x) ⟩ 405 | mapDom j' (get f' (i $ x)) ≈⟨ symd _ (get-nat reflA (i $ x) j') ⟩ 406 | get (mapEnv j' f') (i $ x) ≈⟨ ex' (i $ x) ⟩ 407 | get f'' (i' $ (i $ x)) ≡⟨ cong (get f'') (sym (apply-∘ i' i)) ⟩ 408 | get f'' ((i' ∘i i) $ x) ∎ 409 | 410 | idEnv-≤[] : ∀ {Sg G D} {D1} (i : Inj D D1) → mapEnv i (idEnv {Sg} {G}) ≤[ i ] idEnv 411 | idEnv-≤[] i {S} x = begin[ dom ] 412 | get (mapEnv i idEnv) x ≡⟨ cong (λ g → get g x) (mapAll-build _ injv (mapDom i)) ⟩ 413 | get (build (λ v → mapDom i (injv v))) x ≡⟨ get-build (λ x₁ → mapDom i (injv x₁)) x ⟩ 414 | mapDom i (injv x) ≈⟨ injv-nat i x ⟩ 415 | injv (i $ x) ≡⟨ sym (get-build injv (i $ x)) ⟩ 416 | get idEnv (i $ x) ∎ 417 | 418 | -------------------------------------------------------------------------------- /Syntax/NbE/Properties.agda: -------------------------------------------------------------------------------- 1 | module Syntax.NbE.Properties where 2 | 3 | open import Data.Empty 4 | open import Data.Sum 5 | open import Data.Product 6 | open import Function using (_∘_) 7 | 8 | open import Support.Equality 9 | open import Support.EqReasoning 10 | 11 | open import Injections 12 | 13 | open import Syntax.Type 14 | open import Syntax.NbE 15 | 16 | mutual 17 | eval-tri : ∀ {b Sg G D K D1 T} → (t : Tm< b > Sg G D T) → (i : Inj D K) {g1 : All (Dom Sg G D1) D} {g2 : All (Dom Sg G D1) K} → 18 | g1 ≤[ i ] g2 → T ∋ eval (ren i t) g2 ≡d eval t g1 19 | eval-tri (con c ts) i g1≤g2 = cong (con c) (reifys-ext (evals-tri ts i g1≤g2)) 20 | eval-tri (mvar u j) i g1≤g2 = cong (mvar u) (reifys-ext (evalAs-tri j i g1≤g2)) 21 | eval-tri (var x ts) i {g1} {g2} g1≤g2 = $$-ext {x = get g2 (i $ x)} {get g1 x} (symd _ (g1≤g2 x)) 22 | {evals (rens i ts) g2} {evals ts g1} (evals-tri ts i g1≤g2) 23 | eval-tri (lam t) i {g1} {g2} g1≤g2 = λ D i₁ {s1} {s2} s → eval-tri t (cons i) 24 | (λ { {._} zero → symd _ s; 25 | (suc v) → begin[ dom ] 26 | get (mapEnv i₁ g1) v ≈⟨ get-nat reflA v i₁ ⟩ 27 | mapDom i₁ (get g1 v) ≈⟨ mapDom-ext i₁ (g1≤g2 v) ⟩ 28 | mapDom i₁ (get g2 (i $ v)) ≈⟨ symd _ (get-nat reflA (i $ v) i₁) ⟩ 29 | get (s1 ∷ mapEnv i₁ g2) (thin zero _ (i $ v)) ≡⟨ cong (get (s1 ∷ mapEnv i₁ g2)) (sym (iso2 _ _ v)) ⟩ 30 | get (s1 ∷ mapEnv i₁ g2) (weak i $ v) ∎}) 31 | 32 | evals-tri : ∀ {b Sg G D K D1 T} → (t : Tms< b > Sg G D T) → (i : Inj D K) {g1 : All (Dom Sg G D1) D} {g2 : All (Dom Sg G D1) K} → 33 | g1 ≤[ i ] g2 → evals (rens i t) g2 ≡A evals t g1 34 | evals-tri (t ∷ ts) i le zero = eval-tri t i le 35 | evals-tri (t ∷ ts) i le (suc v) = evals-tri ts i le v 36 | 37 | evalAs-tri : ∀ {b Sg G D K D1 T} → (t : Args b Sg G D T) → (i : Inj D K) {g1 : All (Dom Sg G D1) D} {g2 : All (Dom Sg G D1) K} → 38 | g1 ≤[ i ] g2 → evalAs (renArgs i t) g2 ≡A evalAs t g1 39 | evalAs-tri {true} j i {g1} {g2} le v 40 | rewrite get-build (λ x₁ → get g2 ((i ∘i j) $ x₁)) v 41 | | get-build (λ x₁ → get g1 (j $ x₁)) v = transd _ (≡-d (cong (get g2) (apply-∘ i j))) (symd _ (le (j $ v))) 42 | evalAs-tri {false} ts i le v = evals-tri ts i le v 43 | 44 | -- flip eval maps the square Dom(j) . g1 = g2 . i to Dom(j) . flip eval g1 = flip eval g2 . Tm(i) 45 | -- which makes it a functor from the comma category _∋_ / Dom Sg G of environments to Tm Sg G / Dom Sg G . 46 | eval-square : ∀ {b Sg G D K D1 H T} (t : Tm< b > Sg G D T) (i : Inj D K) (j : Inj D1 H) {g1 : All (Dom Sg G D1) D} {g2 : All (Dom Sg G H) K} → 47 | mapEnv j g1 ≤[ i ] g2 → T ∋ eval (ren i t) g2 ≡d mapDom j (eval t g1) 48 | eval-square t i j {g1} {g2} sq = begin[ dom ] 49 | eval (ren i t) g2 ≈⟨ eval-tri t i sq ⟩ 50 | eval t (mapEnv j g1) ≈⟨ eval-nat t reflA j ⟩ 51 | mapDom j (eval t g1) ∎ 52 | 53 | evals-square : ∀ {b Sg G D K D1 H T} (ts : Tms< b > Sg G D T) (i : Inj D K) (j : Inj D1 H) {g1 : All (Dom Sg G D1) D} {g2 : All (Dom Sg G H) K} → 54 | mapEnv j g1 ≤[ i ] g2 → evals (rens i ts) g2 ≡A mapEnv j (evals ts g1) 55 | evals-square t i j {g1} {g2} sq v = begin[ dom ] 56 | get (evals (rens i t) g2) v ≈⟨ evals-tri t i sq v ⟩ 57 | get (evals t (mapEnv j g1)) v ≈⟨ evals-nat t reflA j v ⟩ 58 | get (mapEnv j (evals t g1)) v ∎ 59 | 60 | 61 | mutual 62 | Rel : ∀ {Sg G B C} T → (f : All (Dom Sg G C) B) → (g : Dom Sg G B T) → (h : Dom Sg G C T) → Set 63 | Rel ([] ->> B) f g h = (! B) ∋ eval g f ≡d h 64 | Rel ((T ∷ Ts) ->> B) f g h = 65 | (∀ D i K j {s1 s2} f' (ex : mapEnv j f ≤[ i ] f') → Rel T f' s1 s2 → Rel (Ts ->> B) f' (proj₁ g D i s1) (proj₁ h K j s2)) 66 | 67 | RelA : ∀ {Sg G A B C} → (f : All (Dom Sg G C) B) 68 | (g : All (Dom Sg G B) A)(h : All (Dom Sg G C) A) → Set 69 | RelA f g h = (∀ {S} (x : _ ∋ S) → Rel S f (get g x) (get h x)) 70 | 71 | transR : ∀ {Sg G B C} T {f : All (Dom Sg G C) B} {g : Dom Sg G B T} {h : Dom Sg G C T} → 72 | Rel T f g h → ∀ {d} → T ∋ h ≡d d → Rel T f g d 73 | transR ([] ->> B) r eq = trans r eq 74 | transR ((T ∷ Ts) ->> B) r eq = (λ D i K j f' ex x → transR (Ts ->> B) (r D i K j f' ex x) (eq K j (refld _))) 75 | 76 | mapRel : ∀ {Sg G B C} T → (f : All (Dom Sg G C) B) {g : Dom Sg G B T} {h : Dom Sg G C T} → Rel T f g h → 77 | ∀ {D} (i : Inj B D) {K} (j : Inj C K) (f' : All (Dom Sg G K) D) 78 | → mapEnv j f ≤[ i ] f' → Rel T f' (mapDom i g) (mapDom j h) 79 | mapRel ([] ->> _) f {g} r i j f' ex = trans (eval-square g i j ex) (cong (ren j) r) 80 | mapRel ((_ ∷ _) ->> _) f r i j f' ex = (λ D i₁ K j₁ f'' ex₁ x₃ → 81 | r D (i₁ ∘i i) K (j₁ ∘i j) f'' (≤[]-∘ f i j f' ex i₁ j₁ f'' ex₁) x₃) 82 | 83 | mapRelA : ∀ {Sg G B C T} (f : All (Dom Sg G C) B) {g : All (Dom Sg G B) T} {h : All (Dom Sg G C) T} → RelA f g h → 84 | ∀ {D} (i : Inj B D) {K} (j : Inj C K) (f' : All (Dom Sg G K) D) 85 | → mapEnv j f ≤[ i ] f' → RelA f' (mapEnv i g) (mapEnv j h) 86 | mapRelA f {g} {h} r i j f' ex {S} x = subst₂ (Rel S f') (sym (get-nat-≡ {f = mapDom i} g x)) (sym (get-nat-≡ {f = mapDom j} h x)) 87 | (mapRel S f (r x) i j f' ex) 88 | 89 | mutual 90 | 91 | expand-∘ : ∀ {Sg G D0 D} Ss {B} (f : All (Dom Sg G D) D0) (d : Dom Sg G D (Ss ->> B)) {e} → 92 | (∀ D i K (j : Inj _ K) {xs1 xs2} f' (ex : mapEnv j f ≤[ i ] f') → RelA f' xs1 xs2 → Rel (! B) f' (proj₁ e D i xs1) (mapDom j d $$ xs2)) 93 | → Rel (Ss ->> B) f (expand Ss e) d 94 | expand-∘ [] f d eq = trans (eq _ id-i _ id-i {[]} {[]} f (λ x → 95 | transd _ (mapEnv-id x) (≡-d (cong (get f) (sym (id-i$ x))))) (λ ())) (ren-id _) 96 | expand-∘ (S ∷ Ss) f (d , d-nat) {e , _} eq = (λ D i K j {x1} {x2} f' ex x₁ → 97 | expand-∘ Ss f' (d K j x2) (λ D₁ i₁ K₁ j₁ {xs1} {xs2} f'' ex₁ x₂ → begin[ dom ] 98 | eval (e D₁ (i₁ ∘i i) (mapDom i₁ x1 ∷ xs1)) f'' ≈⟨ eq _ (i₁ ∘i i) _ (j₁ ∘i j) {mapDom i₁ x1 ∷ xs1} {mapDom j₁ x2 ∷ xs2} 99 | f'' (≤[]-∘ f i j f' ex i₁ j₁ f'' ex₁) 100 | (λ { {._} zero → mapRel S f' x₁ i₁ j₁ f'' ex₁; (suc v) → x₂ v}) ⟩ 101 | d K₁ (id-i ∘i (j₁ ∘i j)) (mapDom j₁ x2) $$ xs2 ≈⟨ $$-ext (≡-d (cong₂ (d K₁) (left-id (j₁ ∘i j)) refl)) 102 | {xs2} {xs2} reflA ⟩ 103 | d K₁ (j₁ ∘i j) (mapDom j₁ x2) $$ xs2 ≈⟨ $$-ext (d-nat _ j (refld _) _ j₁) 104 | {xs2} {xs2} reflA ⟩ 105 | mapDom j₁ (d K j x2) $$ xs2 ∎)) 106 | 107 | injv-∘ : ∀ {Sg G D0 D} T (f : All (Dom Sg G D) D0) (v : _ ∋ T) (d : Dom Sg G D T) 108 | → T ∋ get f v ≡d d → Rel T f (injv v) d 109 | injv-∘ (Ss ->> B) f v d eq = expand-∘ Ss f d (λ D i K j {xs} {ys} f' ex x → 110 | $$-ext {x = get f' (i $ v)} {mapDom j d} (transd _ (transd _ (symd _ (ex v)) (get-nat reflA v j)) (mapDom-ext j eq)) 111 | {evals (reifys xs) f'} {ys} (RelA-unfold f' x)) 112 | 113 | precompose-∘ : ∀ {Sg G Z A B C} → (f : All (Dom Sg G C) B) {g : All (Dom Sg G B) A}{h : All (Dom Sg G C) A} 114 | → RelA f g h → (j : Inj Z A) → RelA f (build (\ z -> get g (j $ z))) (build (\ z -> get h (j $ z))) 115 | precompose-∘ f {g} {h} ar j {S} x = subst₂ (Rel S f) (sym (get-build (λ z → get g (j $ z)) x)) 116 | (sym (get-build (λ z → get h (j $ z)) x)) 117 | (ar (j $ x)) 118 | 119 | idEnv-∘ : ∀ {Sg G D T} (g : All (Dom Sg G D) T) -> RelA g idEnv g 120 | idEnv-∘ {Sg} {G} {D} g x rewrite get-build (injv {Sg} {G}) x = injv-∘ _ g x (get g x) (refld _) 121 | 122 | nf-∘ : ∀ {b Sg G A B C T} → (t : Tm< b > Sg G A T) → (f : All (Dom Sg G C) B){g : All (Dom Sg G B) A}{h : All (Dom Sg G C) A} 123 | → RelA f g h → nf (nf t g) f ≡ nf t h 124 | nf-∘ t f ar = reify-ext _ (Rel-unfold _ f (eval-∘ f ar t)) 125 | 126 | nfs-∘ : ∀ {b Sg G A B C T} → (t : Tms< b > Sg G A T) → (f : All (Dom Sg G C) B){g : All (Dom Sg G B) A}{h : All (Dom Sg G C) A} → 127 | RelA f g h → nfs (nfs t g) f ≡ nfs t h 128 | nfs-∘ [] f ar = refl 129 | nfs-∘ (x ∷ xs) f ar = cong₂ _∷_ (nf-∘ x f ar) (nfs-∘ xs f ar) 130 | 131 | nfAs-∘ : ∀ {b Sg G A B C T} → (t : Args b Sg G A T) → (f : All (Dom Sg G C) B){g : All (Dom Sg G B) A}{h : All (Dom Sg G C) A} → 132 | RelA f g h → nfs (nfAs t g) f ≡ nfAs t h 133 | nfAs-∘ {true} [] f ar = refl 134 | nfAs-∘ {true} (i ∷ t [ pf ]) f ar = cong₂ _∷_ (reify-ext _ (Rel-unfold _ f (ar i))) 135 | (nfAs-∘ t f ar) 136 | nfAs-∘ {false} t f ar = nfs-∘ t f ar 137 | 138 | eval-∘ : ∀ {b Sg G A B C T} → (f : All (Dom Sg G C) B) 139 | {g : All (Dom Sg G B) A}{h : All (Dom Sg G C) A} → RelA f g h → (t : Tm< b > Sg G A T) → Rel T f (eval t g) (eval t h) 140 | eval-∘ f ar (con c ts) = cong (con c) (nfs-∘ ts f ar) 141 | eval-∘ f ar (mvar u j) = cong (mvar u) (nfAs-∘ j f ar) 142 | eval-∘ f ar (var x ts) = $$-∘ f (evals-∘ f ar ts) (ar x) 143 | eval-∘ f {g} {h} ar (lam t) = 144 | (λ D i K j f' ex x → 145 | eval-∘ f' 146 | (λ { {._} zero → x; 147 | {S} (suc v) → subst₂ (Rel S f') (sym (get-nat-≡ {f = mapDom i} g v)) (sym (get-nat-≡ {f = mapDom j} h v)) 148 | (mapRel S f (ar v) i j f' ex)}) 149 | t) 150 | 151 | evals-∘ : ∀ {b Sg G A B C T} → (f : All (Dom Sg G C) B){g : All (Dom Sg G B) A}{h : All (Dom Sg G C) A} → RelA f g h → 152 | (t : Tms< b > Sg G A T) → RelA f (evals t g) (evals t h) 153 | evals-∘ f ar (t ∷ t₁) zero = eval-∘ f ar t 154 | evals-∘ f ar (t ∷ t₁) (suc x) = evals-∘ f ar t₁ x 155 | 156 | $$-∘ : ∀ {Sg G A B C} → (f : All (Dom Sg G C) B) {g : All (Dom Sg G B) A}{h : All (Dom Sg G C) A} → RelA f g h → 157 | ∀ {X} {gx : Dom _ _ _ (_ ->> X)} {hx} → Rel _ f gx hx → _ ∋ eval (gx $$ g) f ≡d (hx $$ h) 158 | $$-∘ f {[]} {[]} ar r = r 159 | $$-∘ f {px ∷ g} {px₁ ∷ h} ar r = $$-∘ f (ar ∘ suc) 160 | (r _ id-i _ id-i f (\ x → transd _ (mapEnv-id x) (≡-d (cong (get f) (sym (id-i$ x))))) (ar zero)) 161 | 162 | Rel-unfold : ∀ {Sg G} A {B C} → (f : All (Dom Sg G C) B) 163 | {g : (Dom Sg G B) A}{h : (Dom Sg G C) A} → Rel A f g h → _ ∋ eval (reify A g) f ≡d h 164 | Rel-unfold ([] ->> _) f r = r 165 | Rel-unfold ((T ∷ Ts) ->> _) f r = λ D i s → 166 | Rel-unfold (Ts ->> _) (_ ∷ mapEnv i f) 167 | (r (T ∷ _) (weak id-i) D i (_ ∷ mapEnv i f) 168 | (λ x → ≡-d (sym (cong (get (_ ∷ mapEnv i f)) (apply-weakid x)))) 169 | (injv-∘ T (_ ∷ mapEnv i f) zero _ s)) 170 | 171 | RelA-unfold : ∀ {Sg G A B C} → (f : All (Dom Sg G C) B) 172 | {g : All (Dom Sg G B) A}{h : All (Dom Sg G C) A} → RelA f g h → evals (reifys g) f ≡A h 173 | RelA-unfold f {gx ∷ g} {hx ∷ h} r zero = Rel-unfold _ f (r zero) 174 | RelA-unfold f {gx ∷ g} {hx ∷ h} r (suc v) = RelA-unfold f (r ∘ suc) v 175 | 176 | 177 | mutual 178 | 179 | expand-id : ∀ {Sg G D T B} f {f-nat} (ts : All (Dom Sg G D) T) -> expand {T = B} T (f , f-nat) $$ ts ≡ f _ id-i ts 180 | expand-id f [] = refl 181 | expand-id f {f-nat} (t ∷ ts) = 182 | begin expand _ (applyN (f , f-nat) id-i t) $$ ts ≡⟨ expand-id _ ts ⟩ 183 | f _ (id-i ∘i id-i) (mapDom id-i t ∷ ts) ≡⟨ cong₂ (f _) (left-id id-i) refl ⟩ 184 | f _ id-i (mapDom id-i t ∷ ts) ≡⟨ reflN (f , f-nat) _ id-i (cons-ext (mapDom-id t) reflA) ⟩ 185 | f _ id-i (t ∷ ts) ∎ 186 | 187 | injv-id : ∀ {Sg G D T B} (v : _ ∋ (_ ->> B)) (ts : All (Dom Sg G D) T) -> injv v $$ ts ≡ var v (reifys ts) 188 | injv-id v ts = begin 189 | expand _ ((λ D1 i xs → var (i $ v) (reifys xs)) , _) $$ ts ≡⟨ expand-id (λ D1 i xs → var (i $ v) (reifys xs)) ts ⟩ 190 | var (id-i $ v) (reifys ts) ≡⟨ cong₂ var (id-i$ v) refl ⟩ 191 | var v (reifys ts) ∎ 192 | 193 | nf-id : ∀ {Sg G D T} (t : Tm< false > Sg G D T) -> nf t idEnv ≡ t 194 | nf-id (con c ts) = cong (con c) (nfs-id ts) 195 | nf-id (mvar u j) = cong (mvar u) (nfs-id j) 196 | nf-id {Sg} {G} (var x ts) = begin 197 | get idEnv x $$ evals ts idEnv ≡⟨ cong (λ f → f $$ evals ts idEnv) (get-build (injv {Sg} {G}) x) ⟩ 198 | injv x $$ evals ts idEnv ≡⟨ expand-id (λ D1 i xs → var (i $ x) (reifys xs)) (evals ts idEnv) ⟩ 199 | var (id-i $ x) (nfs ts idEnv) ≡⟨ cong₂ var (id-i$ x) (nfs-id ts) ⟩ 200 | var x ts ∎ 201 | nf-id (lam t) = 202 | cong lam (begin 203 | nf t (injv zero ∷ mapEnv (weak id-i) (build injv)) 204 | ≡⟨ nf-ext t (cons-ext (injv-ext zero) λ {S} x → begin[ dom ] 205 | get (mapEnv (weak id-i) idEnv) x ≡⟨ cong (λ g → get g x) (mapAll-build _ injv (mapDom (weak id-i))) ⟩ 206 | get (build (λ x₁ → mapDom (weak id-i) (injv x₁))) x ≡⟨ get-build (λ x₁ → mapDom (weak id-i) (injv x₁)) x ⟩ 207 | mapDom (weak id-i) (injv x) ≈⟨ injv-nat (weak id-i) x ⟩ 208 | injv (weak id-i $ x) ≡⟨ cong injv (apply-weakid x) ⟩ 209 | injv (suc x) ≡⟨ sym (get-build (λ x₁ → injv (suc x₁)) x) ⟩ 210 | get (build (λ x₁ → injv (suc x₁))) x ∎) ⟩ 211 | nf t idEnv ≡⟨ nf-id t ⟩ 212 | t ∎) 213 | 214 | nfs-id : ∀ {Sg G D T} (ts : Tms< false > Sg G D T) -> nfs ts idEnv ≡ ts 215 | nfs-id [] = refl 216 | nfs-id (t ∷ ts) = cong₂ _∷_ (nf-id t) (nfs-id ts) 217 | -------------------------------------------------------------------------------- /Syntax/No-Cycle.agda: -------------------------------------------------------------------------------- 1 | module Syntax.No-Cycle where 2 | 3 | open import Data.Empty 4 | open import Data.Nat 5 | open import Data.Nat.Properties 6 | open ≤-Reasoning 7 | open import Relation.Binary.PropositionalEquality 8 | 9 | open import Support.Product 10 | 11 | open import Injections 12 | 13 | open import Syntax.Type 14 | open import Syntax.OneHoleContext 15 | 16 | Height = ℕ 17 | 18 | mutual 19 | height : ∀ {Sg G D T b} -> Tm< b > Sg G D T -> Height 20 | height (con c ts) = suc (heights ts) 21 | height (var x ts) = suc (heights ts) 22 | height (lam t) = suc (height t) 23 | height (mvar u j) = 0 24 | 25 | heights : ∀ {Sg G D T b} -> Tms< b > Sg G D T -> Height 26 | heights [] = 0 27 | heights (t ∷ ts) = suc (height t ⊔ heights ts) 28 | 29 | heightT : ∀ {Sg G D T b} -> Term< b > Sg G D T -> Height 30 | heightT {T = inj₁ _} = height 31 | heightT {T = inj₂ _} = heights 32 | 33 | renT-height : ∀ {T Sg G D D1 b} -> (i : Inj D D1) -> (t : Term< b > Sg G D T) -> heightT t ≡ heightT (renT i t) 34 | renT-height {inj₁ ._} i (con c ts) = cong suc (renT-height i ts) 35 | renT-height {inj₁ ._} i (mvar u j) = refl 36 | renT-height {inj₁ ._} i (var x ts) = cong suc (renT-height i ts) 37 | renT-height {inj₁ ._} i (lam t) = cong suc (renT-height _ t) 38 | renT-height {inj₂ .[]} i [] = refl 39 | renT-height {inj₂ ._} i (t ∷ t₁) = cong₂ (\ x y -> suc (x ⊔ y)) (renT-height i t) (renT-height i t₁) 40 | 41 | open import Data.Nat 42 | 43 | private 44 | n≤m⊔n : ∀ m n → (m ⊔ n) ≥ n 45 | n≤m⊔n zero _ = begin _ ∎ 46 | n≤m⊔n (suc m) zero = z≤n 47 | n≤m⊔n (suc m) (suc n) = s≤s (n≤m⊔n m n) 48 | 49 | wrapD-height : ∀ {Sg G DI TI DO TO b} → (d : DTm< b > Sg G (DI , TI) (DO , TO)) → (t : Term< b > Sg G DO TO) → heightT (wrapD d t) > heightT t 50 | wrapD-height lam t = s≤s (begin heightT t ∎) 51 | wrapD-height (head ts) t = s≤s (m≤m⊔n (height t) (heights ts)) 52 | wrapD-height (tail t) ts = s≤s (n≤m⊔n (height t) (heights ts)) 53 | wrapD-height (con c) ts = s≤s (begin heightT ts ∎) 54 | wrapD-height (var x) ts = s≤s (begin heightT ts ∎) 55 | 56 | wrap-height : ∀ {Sg G I O b} → (C : Context< b > Sg G I O) → (t : Term< b > Sg G (proj₁ O) (proj₂ O)) → heightT (wrap C t) ≥ heightT t 57 | wrap-height [] t = begin heightT t ∎ 58 | wrap-height (d ∷ c) t = begin heightT t ≤⟨ ≤-step (wrap-height c t) ⟩ 59 | suc (heightT (wrap c t)) ≤⟨ wrapD-height d (wrap c t) ⟩ 60 | heightT (wrapD d (wrap c t)) ∎ 61 | 62 | No-Cycle : ∀ {b TI Sg G D1 DI DO X} -> let TO = TI in 63 | (d : DTm< b > Sg G (DI , TI) X) (c : Context< b > Sg G X (DO , TO)) 64 | (t : Term< b > Sg G D1 TO) (i : Inj D1 DI)(j : Inj D1 DO) -> 65 | ¬ renT i t ≡ wrap (d ∷ c) (renT j t) 66 | No-Cycle d c t i j eq = ≡-or-> (cong heightT eq) 67 | (begin 68 | suc (heightT (renT i t)) ≡⟨ cong suc (sym (renT-height i t)) ⟩ 69 | suc (heightT t) ≡⟨ cong suc (renT-height j t) ⟩ 70 | suc (heightT (renT j t)) ≤⟨ s≤s (wrap-height c (renT j t)) ⟩ 71 | suc (heightT (wrap c (renT j t))) ≤⟨ wrapD-height d (wrap c (renT j t)) ⟩ 72 | heightT (wrapD d (wrap c (renT j t))) ∎) 73 | where 74 | ≡-or-> : ∀ {m n} -> m ≡ n -> n > m -> ⊥ 75 | ≡-or-> refl (s≤s ge) = ≡-or-> refl ge 76 | -------------------------------------------------------------------------------- /Syntax/OneHoleContext.agda: -------------------------------------------------------------------------------- 1 | module Syntax.OneHoleContext where 2 | 3 | open import Relation.Binary.PropositionalEquality 4 | open ≡-Reasoning 5 | open import Data.Sum 6 | 7 | open import Support.Product 8 | 9 | open import Injections 10 | 11 | open import Syntax.Type 12 | open import Syntax.Sub 13 | open import Syntax.Equality 14 | 15 | 16 | -- Given an inductive type T = F T we can build the type of its one 17 | -- hole contexts as List (F' T) where F' is the derivative of F, 18 | -- i.e. the elements of F' x are the same as the ones of F x but with 19 | -- one occurrence of x removed. For indexed types like Term we only 20 | -- need to lift this construction pointwise. 21 | -- 22 | -- In our case T = Term Sg G and we define: 23 | -- DTm Sg G as F' T, 24 | -- Context Sg G as List (F' T). 25 | 26 | Index = Ctx × (Ty ⊎ List Ty) 27 | 28 | data DTm<_> (b : Bool) (Sg : Ctx)(G : MCtx) : Index -> Index → Set where 29 | lam : ∀ {D S Ss B} → DTm< b > Sg G (D , inj₁ (S :> Ss ->> B)) (D <: S , inj₁ (Ss ->> B)) 30 | head : ∀ {D S Ss} → (ts : Tms< b > Sg G D Ss) → DTm< b > Sg G (D , inj₂ (S :> Ss)) (D , inj₁ S) 31 | tail : ∀ {D S Ss} → (t : Tm< b > Sg G D S) → DTm< b > Sg G (D , inj₂ (S :> Ss)) (D , inj₂ Ss) 32 | con : ∀ {D Ss B} → (c : Sg ∋ (Ss ->> B)) → DTm< b > Sg G (D , inj₁ (! B)) (D , inj₂ Ss) 33 | var : ∀ {D Ss B} → (x : D ∋ (Ss ->> B)) → DTm< b > Sg G (D , inj₁ (! B)) (D , inj₂ Ss) 34 | 35 | DTm : (Sg : Ctx)(G : MCtx) -> Index -> Index → Set 36 | DTm = DTm< true > 37 | 38 | data IList {I : Set}(T : I → I → Set) (i : I) : (j : I) → Set where 39 | [] : IList T i i 40 | _∷_ : ∀ {k j} → T i k → (xs : IList T k j) → IList T i j 41 | 42 | Context<_> : (b : Bool) (Sg : Ctx)(G : MCtx) → Index -> Index → Set 43 | Context< b > Sg G = IList (DTm< b > Sg G) 44 | 45 | Context : (Sg : Ctx)(G : MCtx) → Index -> Index → Set 46 | Context = Context< true > 47 | 48 | -- Given a Context and an index-compatible filling we can rebuild a Term 49 | wrapD : ∀ {Sg G DI TI DO TO b} → DTm< b > Sg G (DI , TI) (DO , TO) → Term< b > Sg G DO TO → Term< b > Sg G DI TI 50 | wrapD lam t = lam t 51 | wrapD (head ts) t = t ∷ ts 52 | wrapD (tail t) ts = t ∷ ts 53 | wrapD (con c) ts = con c ts 54 | wrapD (var x) ts = var x ts 55 | 56 | wrap : ∀ {Sg G I O b} → Context< b > Sg G I O → Term< b > Sg G (proj₁ O) (proj₂ O) → Term< b > Sg G (proj₁ I) (proj₂ I) 57 | wrap [] t = t 58 | wrap (d ∷ c) t = wrapD d (wrap c t) 59 | 60 | 61 | -- To move a renaming past a λ we need to handle the extra variable, 62 | -- wrapoInj takes care of the induced action of a DTm on Inj. 63 | wrapD-Ctx : ∀ {Sg G I1 I2 b} -> DTm< b > Sg G I1 I2 -> Ctx -> Ctx 64 | wrapD-Ctx (lam {S = S}) D = S ∷ D 65 | wrapD-Ctx _ D = D 66 | 67 | wrapD-Inj : ∀ {Sg G DI I1 I2 b} -> (d : DTm< b > Sg G I1 I2) -> Inj DI (proj₁ I1) -> Inj (wrapD-Ctx d DI) (proj₁ I2) 68 | wrapD-Inj lam i = cons i 69 | wrapD-Inj (head ts) i = i 70 | wrapD-Inj (tail t) i = i 71 | wrapD-Inj (con c) i = i 72 | wrapD-Inj (var x) i = i 73 | 74 | 75 | subD : ∀ {Sg G1 G2 TI TO b1 b2} -> Sub< b1 > Sg G1 G2 -> DTm< b2 > Sg G1 TI TO -> DTm< b2 ∧ b1 > Sg G2 TI TO 76 | subD s lam = lam 77 | subD s (head ts) = head (subs s ts) 78 | subD s (tail t) = tail (sub s t) 79 | subD s (con c) = con c 80 | subD s (var x) = var x 81 | 82 | 83 | subC : ∀ {Sg G1 G2 TI TO b1 b2} -> (Sub< b1 > Sg G1 G2) -> Context< b2 > Sg G1 TI TO -> Context< b2 ∧ b1 > Sg G2 TI TO 84 | subC s [] = [] 85 | subC s (x ∷ c) = subD s x ∷ (subC s c) 86 | 87 | wrapD-sub : ∀ {Sg G1 G2 TI TO b1 b2} -> (s : Sub< b1 > Sg G1 G2) -> (d : DTm< b2 > Sg G1 TI TO) -> 88 | ∀ t -> subT s (wrapD d t) ≡ wrapD (subD s d) (subT s t) 89 | wrapD-sub s lam t = refl 90 | wrapD-sub s (head ts) t = refl 91 | wrapD-sub s (tail t) ts = refl 92 | wrapD-sub s (con c) ts = refl 93 | wrapD-sub s (var x) ts = refl 94 | 95 | wrap-sub : ∀ {Sg G1 G2 TI TO b1 b2} -> (s : Sub< b1 > Sg G1 G2) -> (c : Context< b2 > Sg G1 TI TO) -> 96 | ∀ t -> subT s (wrap c t) ≡ wrap (subC s c) (subT s t) 97 | wrap-sub s [] t = refl 98 | wrap-sub s (d ∷ c) t = trans (wrapD-sub s d _) (cong (wrapD (subD s d)) (wrap-sub s c t)) 99 | 100 | 101 | cong-wrapD : ∀ {Sg G1 G2 TI TO b1 b2} -> {s : Sub< b1 > Sg G1 G2} -> (d : DTm< b2 > Sg G1 TI TO) -> 102 | ∀ {x y} -> subT s x ≡T subT s y -> subT s (wrapD d x) ≡T subT s (wrapD d y) 103 | cong-wrapD {s = s} d {x} {y} eq = ≡-T 104 | (begin 105 | subT s (wrapD d x) ≡⟨ wrapD-sub s d x ⟩ 106 | wrapD (subD s d) (subT s x) ≡⟨ cong (wrapD (subD s d)) (T-≡ eq) ⟩ 107 | wrapD (subD s d) (subT s y) ≡⟨ sym (wrapD-sub s d y) ⟩ 108 | subT s (wrapD d y) ∎) 109 | 110 | wrapD-inj : ∀ {Sg G1 TI TO b} -> (d : DTm< b > Sg G1 TI TO) -> ∀ {x y} -> wrapD d x ≡ wrapD d y -> x ≡ y 111 | wrapD-inj lam refl = refl 112 | wrapD-inj (head ts) refl = refl 113 | wrapD-inj (tail t) refl = refl 114 | wrapD-inj (con c) refl = refl 115 | wrapD-inj (var x) refl = refl 116 | 117 | inv-wrapD : ∀ {Sg G1 G2 TI TO b1 b2} -> {s : Sub< b1 > Sg G1 G2} -> (d : DTm< b2 > Sg G1 TI TO) -> 118 | ∀ {x y} -> subT s (wrapD d x) ≡T subT s (wrapD d y) -> subT s x ≡T subT s y 119 | inv-wrapD {s = s} d {x} {y} eq = ≡-T (wrapD-inj (subD s d) 120 | (begin 121 | wrapD (subD s d) (subT s x) ≡⟨ sym (wrapD-sub s d x) ⟩ 122 | subT s (wrapD d x) ≡⟨ T-≡ eq ⟩ 123 | subT s (wrapD d y) ≡⟨ wrapD-sub s d y ⟩ 124 | wrapD (subD s d) (subT s y) ∎)) 125 | -------------------------------------------------------------------------------- /Syntax/RenOrn.agda: -------------------------------------------------------------------------------- 1 | module Syntax.RenOrn where 2 | 3 | open import Relation.Binary.PropositionalEquality 4 | open import Data.Sum 5 | open import Data.Bool 6 | 7 | open import Support.Product 8 | 9 | open import Injections 10 | 11 | open import Syntax.Type 12 | 13 | -- RTm i t is the inductive version of Σ (Tm Sg G D T) λ s → ren i s ≡ t 14 | -- In the language of "Ornamental Algebras, Algebraic Ornaments" RTm i t 15 | -- would be the latter and we'd get both RTm and the remember/forget 16 | -- conversions for free by expressing ren i as a fold. 17 | mutual 18 | data RTm {b : Bool} {Sg : Ctx}{G : MCtx}{D : Ctx}{K : Ctx} (i : Inj D K) : {T : Ty} → Tm< b > Sg G K T → Set where 19 | con : {Ss : Fwd Ty}{B : Base} → 20 | (c : Sg ∋ (Ss ->> B)) → ∀ {tms} → (RTms i {Ss} tms) → RTm i {(! B)} (con c tms) 21 | mvar : {Ss : Bwd Ty}{B : Base} → 22 | (u : G ∋ (B <<- Ss)) → ∀ {k} → RArgs b i k → RTm i {(! B)} (mvar u k) 23 | var : forall {Ss B} → (v : D ∋ (Ss ->> B)) → ∀ {x} → (i $ v) ≡ x → ∀ {tms} → RTms i {Ss} tms → RTm i {(! B)} (var x tms) 24 | lam : {S : Ty}{Ss : Fwd Ty}{B : Base} → ∀ {b} → 25 | RTm (cons i) {(Ss ->> B)} b → RTm i {(S :> Ss ->> B)} (lam b) 26 | 27 | data RTms {b : Bool}{Sg : Ctx}{G : MCtx}{D K : Ctx}(i : Inj D K) : {Ss : Fwd Ty} → Tms< b > Sg G K Ss → Set where 28 | [] : RTms i { !> } [] 29 | _∷_ : {S : Ty}{Ss : Fwd Ty} → ∀ {x xs} → 30 | RTm i {S} x → RTms i {Ss} xs → RTms i {(S :> Ss)} (x ∷ xs) 31 | 32 | RArgs : (b : Bool){Sg : Ctx}{G : MCtx}{D K : Ctx}(i : Inj D K) -> {Ss : Fwd Ty} → Args b Sg G K Ss → Set 33 | RArgs true i k = ∃ (λ j → i ∘i j ≡ k) 34 | RArgs false i xs = RTms i xs 35 | 36 | _⁻¹_ : ∀ {b}{Sg : Ctx}{G : MCtx}{D : Ctx}{K : Ctx} (i : Inj D K) {T} → Term< b > Sg G K T → Set 37 | _⁻¹_ i {inj₁ _} t = RTm i t 38 | _⁻¹_ i {inj₂ _} ts = RTms i ts 39 | 40 | mutual 41 | forget : ∀ {b Sg G D D0}{T : Ty} → {i : Inj D D0} → {t : Tm< b > Sg G D0 T} → (x : RTm i t) → Σ (Tm< b > Sg G D T) \ s → ren i s ≡ t 42 | forget (con c x) = mapΣ (con c) (cong (con c)) (forgets x) 43 | forget {true} (mvar u (j , eq)) = mvar u j , cong (mvar u) eq 44 | forget {false} (mvar u ts) = mapΣ (mvar u) (cong (mvar u)) (forgets ts) 45 | forget {i = i} (var v refl x₂) = mapΣ (var v) (cong (var (i $ v))) (forgets x₂) 46 | forget (lam x) = mapΣ lam (cong lam) (forget x) 47 | 48 | forgets : ∀ {b Sg G D D0}{Ts} → {i : Inj D D0} → {t : Tms< b > Sg G D0 Ts} → (x : RTms i t) → Σ (Tms< b > Sg G D Ts) \ s → rens i s ≡ t 49 | forgets [] = [] , refl 50 | forgets (t ∷ ts) = mapΣ₂ _∷_ (cong₂ _∷_) (forget t) (forgets ts) 51 | 52 | mutual 53 | remember : ∀ {b Sg G D D0}{T : Ty} → (i : Inj D D0) → (s : Tm< b > Sg G D T) → Σ (RTm i (ren i s)) \ rt -> proj₁ (forget rt) ≡ s 54 | remember i (con c ts) = mapΣ (con c) (cong (con c)) (remembers i ts) 55 | remember {true} i (mvar u j) = mvar u (j , refl) , refl 56 | remember {false} i (mvar u ts) = mapΣ (mvar u) (cong (mvar u)) (remembers i ts) 57 | remember i (var x ts) = mapΣ (var x refl) (cong (var x)) (remembers i ts) 58 | remember i (lam s) = mapΣ lam (cong lam) (remember (cons i) s) 59 | 60 | remembers : ∀ {b Sg G D D0}{T} → (i : Inj D D0) → (s : Tms< b > Sg G D T) → Σ (RTms i (rens i s)) \ rt -> proj₁ (forgets rt) ≡ s 61 | remembers i [] = [] , refl 62 | remembers i (t ∷ ts) = mapΣ₂ _∷_ (cong₂ _∷_) (remember i t) (remembers i ts) 63 | 64 | 65 | -- One advantage of this representation is the good interaction with 66 | -- pattern matching as shown here in unique: since x and y share the 67 | -- index t we don't have to consider the patterns where they are 68 | -- matched against different constructors. 69 | mutual 70 | unique : ∀ {b Sg G D D0}{T : Ty} → {i : Inj D D0} → {t : Tm< b > Sg G D0 T} → (x y : RTm i t) → x ≡ y 71 | unique (con c xs) (con .c ys) = cong (con c) (uniques xs ys) 72 | unique (lam x) (lam y) = cong lam (unique x y) 73 | unique {false} (mvar u xs) (mvar .u ys) = cong (mvar u) (uniques xs ys) 74 | unique {i = i} (var v i$v≡x xs) (var w i$w≡x ys) with injective i v w (trans i$v≡x (sym i$w≡x)) 75 | unique (var v i$v≡x xs) (var .v i$w≡x ys) | refl 76 | = cong₂ (λ a b → var v a b) (proof-irrelevance i$v≡x i$w≡x) (uniques xs ys) 77 | unique {true} {i = i} (mvar u (j , i∘j≡k)) (mvar .u (j₁ , i∘j₁≡k)) with ∘i-inj i j j₁ (trans i∘j≡k (sym i∘j₁≡k)) 78 | unique {true} {i = i} (mvar u (j , i∘j≡k)) (mvar .u (.j , i∘j₁≡k)) | refl 79 | = cong (λ e → mvar u (j , e)) (proof-irrelevance i∘j≡k i∘j₁≡k) 80 | 81 | uniques : ∀ {b Sg G D D0}{Ss : Fwd Ty} → {i : Inj D D0} → {t : Tms< b > Sg G D0 Ss} → (x y : RTms i t) → x ≡ y 82 | uniques [] [] = refl 83 | uniques (x ∷ xs) (y ∷ ys) = cong₂ _∷_ (unique x y) (uniques xs ys) 84 | 85 | ren-inj : ∀ {b Sg G D D0}{T : Ty} → (i : Inj D D0) → (s t : Tm< b > Sg G D T) -> ren i s ≡ ren i t -> s ≡ t 86 | ren-inj i s t eq 87 | with remember i s | remember i t 88 | ... | (rs , fogrs≡s) | (rt , fogrt≡t) rewrite eq = trans (sym fogrs≡s) (trans (cong (λ r → proj₁ (forget r)) (unique rs rt)) fogrt≡t) 89 | -------------------------------------------------------------------------------- /Syntax/Sub.agda: -------------------------------------------------------------------------------- 1 | module Syntax.Sub where 2 | 3 | open import Data.Unit using (⊤) 4 | open import Data.Sum 5 | 6 | open import Support.Equality 7 | open ≡-Reasoning 8 | open import Support.Product 9 | 10 | open import Injections 11 | 12 | open import Syntax.Type 13 | open import Syntax.NbE 14 | open import Syntax.NbE.Properties 15 | 16 | -- The type of meta-substitutions 17 | Sub<_> : ∀ b → Ctx → MCtx → MCtx → Set 18 | Sub< b > Sg G1 G2 = ∀ S → G1 ∋ S → Tm< b > Sg G2 (ctx S) (! type S) 19 | 20 | Sub : Ctx → MCtx → MCtx → Set 21 | Sub = Sub< true > 22 | 23 | _≡s_ : ∀ {b Sg G G1}(s s₁ : Sub< b > Sg G G1) → Set 24 | s ≡s s₁ = ∀ S u → s S u ≡ s₁ S u 25 | 26 | thin-s : ∀ {Sg}{G T} → (u : G ∋ T) → Sub Sg (G - u) G 27 | thin-s u = \ S v → mvar (thin u S v) id-i 28 | 29 | mutual 30 | sub : ∀ {b1 b2 Sg G1 G2 D T} → Sub< b1 > Sg G1 G2 → Tm< b2 > Sg G1 D T → Tm< b2 ∧ b1 > Sg G2 D T 31 | sub s (con c ts) = con c (subs s ts) 32 | sub {b1} {true} s (mvar u i) = ren i (s _ u) 33 | sub {b1} {false} s (mvar u ts) = replace (s _ u) (subs s ts) 34 | sub s (var x ts) = var x (subs s ts) 35 | sub s (lam t) = lam (sub s t) 36 | 37 | subs : ∀ {b1 b2 Sg G1 G2 D Ss} → Sub< b1 > Sg G1 G2 → Tms< b2 > Sg G1 D Ss → Tms< b2 ∧ b1 > Sg G2 D Ss 38 | subs s [] = [] 39 | subs s (t ∷ ts) = sub s t ∷ subs s ts 40 | 41 | id-s : ∀ {Sg G} → Sub Sg G G 42 | id-s = \ S x → mvar x id-i 43 | 44 | id-f : ∀ {Sg G} → Sub< false > Sg G G 45 | id-f = \ S u → mvar u (reifys idEnv) 46 | 47 | _∘s_ : ∀ {b1 b2 Sg G1 G2 G3} → Sub< b1 > Sg G2 G3 → Sub< b2 > Sg G1 G2 → Sub< b2 ∧ b1 > Sg G1 G3 48 | r ∘s s = λ S x → sub r (s S x) 49 | 50 | -- f ≤ g iff g can be specialized to f 51 | _≤_ : ∀ {Sg G G1 G2 b} → Sub< b > Sg G G1 → Sub Sg G G2 → Set 52 | f ≤ g = ∃ \ r → ∀ S u → f S u ≡ (r ∘s g) S u 53 | 54 | subT : ∀ {b1 b2 T Sg G1 G2 D} → (Sub< b1 > Sg G1 G2) → Term< b2 > Sg G1 D T → Term< b2 ∧ b1 > Sg G2 D T 55 | subT {T = inj₁ _} i t = sub i t 56 | subT {T = inj₂ _} i ts = subs i ts 57 | 58 | mutual 59 | Dom-nats : ∀ {Sg G1 G2 D b} T → (f : Sub< b > Sg G1 G2) (x : Dom Sg G1 D T)(y : Dom Sg G2 D T) → Set 60 | Dom-nats ([] ->> _) f x y = sub f x ≡ y 61 | Dom-nats ((S ∷ Ss) ->> B) f x y = ∀ D1 r {s1 s2} → Rel S idEnv s2 s2 → 62 | Dom-nats S f s1 s2 → Dom-nats (Ss ->> B) f (proj₁ x D1 r s1) (proj₁ y D1 r s2) 63 | 64 | Env-nats : ∀ {Sg G1 G2 D Ts b} → (f : Sub< b > Sg G1 G2) (x : All (Dom Sg G1 D) Ts)(y : All (Dom Sg G2 D) Ts) → Set 65 | Env-nats f [] [] = ⊤ 66 | Env-nats f (x ∷ xs) (y ∷ ys) = Dom-nats _ f x y × Env-nats f xs ys 67 | 68 | RelId : ∀ {Sg G D Ts} → (y : All (Dom Sg G D) Ts) → Set 69 | RelId f = RelA idEnv f f 70 | 71 | get-nats : ∀ {b1 Sg G1 G2 D1 D2} {f : Sub< b1 > Sg G1 G2} {g1 : All (Dom Sg G1 D1) D2}{g2 : All (Dom Sg G2 D1) D2} 72 | → Env-nats f g1 g2 → ∀ {T} x → Dom-nats T f (get g1 x) (get g2 x) 73 | get-nats {g1 = _ ∷ _} {_ ∷ _} (dnat , _) zero = dnat 74 | get-nats {g1 = _ ∷ _} {_ ∷ _} (_ , enat) (suc x) = get-nats enat x 75 | 76 | mutual 77 | sub-nat : ∀ {b1 b2 Sg G1 G2 D1 D2 T} {f : Sub< b1 > Sg G1 G2} {i : Inj D1 D2} (t : Tm< b2 > Sg G1 D1 T) → sub f (ren i t) ≡ ren i (sub f t) 78 | sub-nat (con c ts) = cong (con c) (subs-nat ts) 79 | sub-nat (var x ts) = cong (var _) (subs-nat ts) 80 | sub-nat (lam t) = cong lam (sub-nat t) 81 | sub-nat {b1} {true} (mvar u j) = ren-∘ _ 82 | sub-nat {b1} {false} {f = f} {i} (mvar u xs) = begin 83 | replace (f _ u) (subs f (rens i xs)) ≡⟨ cong (replace (f _ u)) (subs-nat xs) ⟩ 84 | replace (f _ u) (rens i (subs f xs)) ≡⟨ eval-ext (f _ u) (evals-square (subs f xs) i i (idEnv-≤[] i)) ⟩ 85 | eval (f _ u) (mapEnv i (evals (subs f xs) idEnv)) ≡⟨ eval-nat (f _ u) reflA i ⟩ 86 | ren i (sub f (mvar u xs)) ∎ 87 | 88 | subs-nat : ∀ {b1 b2 Sg G1 G2 D1 D2 T} {f : Sub< b1 > Sg G1 G2} {i : Inj D1 D2} (t : Tms< b2 > Sg G1 D1 T) → subs f (rens i t) ≡ rens i (subs f t) 89 | subs-nat [] = refl 90 | subs-nat (t ∷ ts) = cong₂ _∷_ (sub-nat t) (subs-nat ts) 91 | 92 | mutual 93 | eval-nats : ∀ {b1 b2 Sg G1 G2 D1 D2 T} {f : Sub< b1 > Sg G1 G2} {g1 : All (Dom Sg G1 D1) D2}{g2 : All (Dom Sg G2 D1) D2} 94 | → Env-nats f g1 g2 → RelId g2 → (t : Tm< b2 > Sg G1 D2 T) → Dom-nats T f (eval t g1) (eval (sub f t) g2) 95 | eval-nats g gr (con c ts) = cong (con c) (nfs-nats g gr ts) 96 | eval-nats g gr (var x ts) = $$-nats _ (get-nats g x) (evals-nats g gr ts) (evals-∘ idEnv gr _) 97 | eval-nats g gr (lam t) = λ D1 r sr x → eval-nats (x , mapEnv-nats _ _ r g) 98 | (λ { {._} zero → sr; {S} (suc v) → mapRelA idEnv gr r r idEnv (idEnv-≤[] r) v }) t 99 | eval-nats {_} {true} {f = f} {g1} {g2} g gr (mvar u j) = begin 100 | eval (f _ u) (evals (subs f (reifys (build (λ x₁ → get g1 (j $ x₁))))) idEnv) 101 | ≡⟨ cong (λ ts → eval (f _ u) (evals ts idEnv)) (reifys-nats {f = f} (evalAs-nats g j)) ⟩ 102 | eval (f _ u) (evals (reifys (build (λ z → get g2 (j $ z)))) idEnv) 103 | ≡⟨ eval-ext (f _ u) (RelA-unfold idEnv (precompose-∘ idEnv gr j)) ⟩ 104 | eval (f _ u) (build (λ z → get g2 (j $ z))) 105 | ≡⟨ sym (eval-tri (f _ u) j (λ x → ≡-d (get-build (λ v → get g2 (j $ v)) x))) ⟩ 106 | eval (ren j (f _ u)) g2 ∎ 107 | 108 | eval-nats {_} {false} {f = f} {g1} {g2} g gr (mvar u ts) = begin 109 | replace (f _ u) (subs f (reifys (evals ts g1))) ≡⟨ cong (replace (f _ u)) (reifys-nats {f = f} (evals-nats g gr ts)) ⟩ 110 | nf (f _ u) (evals (reifys (evals (subs f ts) g2)) idEnv) ≡⟨ sym 111 | (nf-∘ (f _ u) g2 112 | (λ {S} x → 113 | transR S (evals-∘ g2 (idEnv-∘ g2) (subs f ts) x) 114 | (symd _ (RelA-unfold idEnv (evals-∘ idEnv gr (subs f ts)) x)))) ⟩ 115 | nf (nf (f _ u) (evals (subs f ts) idEnv)) g2 ∎ 116 | 117 | 118 | evals-nats : ∀ {b1 b2 Sg G1 G2 D1 D2 T} {f : Sub< b1 > Sg G1 G2} {g1 : All (Dom Sg G1 D1) D2}{g2 : All (Dom Sg G2 D1) D2} 119 | → Env-nats f g1 g2 → RelId g2 → (t : Tms< b2 > Sg G1 D2 T) → Env-nats f (evals t g1) (evals (subs f t) g2) 120 | evals-nats enat gr [] = _ 121 | evals-nats enat gr (x ∷ xs) = eval-nats enat gr x , evals-nats enat gr xs 122 | 123 | evalAs-nats : ∀ {b1 Sg G1 G2 D1 D2 T} {f : Sub< b1 > Sg G1 G2} {g1 : All (Dom Sg G1 D1) D2}{g2 : All (Dom Sg G2 D1) D2} 124 | → Env-nats f g1 g2 → (t : Args true Sg G1 D2 T) → Env-nats f (evalAs t g1) (evalAs t g2) 125 | evalAs-nats g [] = _ 126 | evalAs-nats g (i ∷ t [ pf ]) = get-nats g i , evalAs-nats g t 127 | 128 | expand-nats : ∀ {b1 Sg G1 G2 D B} Ss {f : Sub< b1 > Sg G1 G2}{c1 c2} 129 | → (∀ D r {xs ys} → Env-nats f xs ys → sub f (proj₁ c1 D r xs) ≡ proj₁ c2 D r ys) 130 | → Dom-nats {D = D} (Ss ->> B) f (expand Ss c1) (expand Ss c2) 131 | expand-nats [] h = h _ id-i _ 132 | expand-nats (S ∷ Ss) h = λ D1 r sr x₁ → expand-nats Ss (λ D r₁ eq → h D (r₁ ∘i r) (mapDom-nats S r₁ x₁ , eq)) 133 | 134 | injv-nats : ∀ {b1 Sg G1 G2 D T} {f : Sub< b1 > Sg G1 G2} → (v : D ∋ T) → Dom-nats T f (injv v) (injv v) 135 | injv-nats {T = _ ->> _} v = expand-nats _ (λ D r x → cong (var _) (reifys-nats x)) 136 | 137 | reify-nats : ∀ {b1 Sg G1 G2 D1} T {f : Sub< b1 > Sg G1 G2} {t1 : (Dom Sg G1 D1) T}{t2 : (Dom Sg G2 D1) T} 138 | → Dom-nats T f t1 t2 → sub f (reify T t1) ≡ reify T t2 139 | reify-nats ([] ->> _) t = t 140 | reify-nats ((S ∷ Ss) ->> _) t = cong lam (reify-nats (Ss ->> _) (t _ (weak id-i) (injv-∘ S idEnv zero (injv zero) (refld _)) (injv-nats zero))) 141 | 142 | reifys-nats : ∀ {b1 Sg G1 G2 D1 D2} {f : Sub< b1 > Sg G1 G2} {g1 : All (Dom Sg G1 D1) D2}{g2 : All (Dom Sg G2 D1) D2} 143 | → Env-nats f g1 g2 → subs f (reifys g1) ≡ reifys g2 144 | reifys-nats {g1 = []} {[]} eq = refl 145 | reifys-nats {g1 = x ∷ xs} {y ∷ ys} eq = cong₂ _∷_ (reify-nats _ (proj₁ eq)) (reifys-nats (proj₂ eq)) 146 | 147 | nf-nats : ∀ {b1 b2 Sg G1 G2 D1 D2 T} {f : Sub< b1 > Sg G1 G2} {g1 : All (Dom Sg G1 D1) D2}{g2 : All (Dom Sg G2 D1) D2} 148 | → Env-nats f g1 g2 → RelId g2 → (t : Tm< b2 > Sg G1 D2 T) → sub f (nf t g1) ≡ nf (sub f t) g2 149 | nf-nats enat gr t = reify-nats _ (eval-nats enat gr t) 150 | 151 | nfs-nats : ∀ {b1 b2 Sg G1 G2 D1 D2 T} {f : Sub< b1 > Sg G1 G2} {g1 : All (Dom Sg G1 D1) D2}{g2 : All (Dom Sg G2 D1) D2} 152 | → Env-nats f g1 g2 → RelId g2 → (t : Tms< b2 > Sg G1 D2 T) → subs f (nfs t g1) ≡ nfs (subs f t) g2 153 | nfs-nats g gr [] = refl 154 | nfs-nats g gr (x ∷ xs) = cong₂ _∷_ (nf-nats g gr x) (nfs-nats g gr xs) 155 | 156 | mapDom-nats : ∀ {Sg G1 G2 D D1 b} Ts → {f : Sub< b > Sg G1 G2} {x : (Dom Sg G1 D) Ts}{y : (Dom Sg G2 D) Ts} 157 | → (i : Inj D D1) → Dom-nats Ts f x y → Dom-nats Ts f (mapDom i x) (mapDom i y) 158 | mapDom-nats ([] ->> _) {x = x} i d = trans (sub-nat x) (cong (ren i) d) 159 | mapDom-nats ((S ∷ Ss) ->> _) i d = λ D1 r x₁ → d D1 (r ∘i i) x₁ 160 | 161 | mapEnv-nats : ∀ {Sg G1 G2 D D1 Ts b} → {f : Sub< b > Sg G1 G2} (x : All (Dom Sg G1 D) Ts)(y : All (Dom Sg G2 D) Ts) 162 | → (i : Inj D D1) → Env-nats f x y → Env-nats f (mapEnv i x) (mapEnv i y) 163 | mapEnv-nats [] [] i e = _ 164 | mapEnv-nats (x ∷ xs) (y ∷ ys) i (s , e) = mapDom-nats _ i s , mapEnv-nats xs ys i e 165 | 166 | $$-nats : ∀ {Sg G G1 D Ts B b} (s : Sub< b > Sg G G1) 167 | {f1 : Dom Sg G D (Ts ->> B)} {f2 : Dom Sg G1 D (Ts ->> B)} {xs : All (Dom Sg G D) Ts} {xs1 : All (Dom Sg G1 D) Ts} → 168 | Dom-nats (Ts ->> B) s f1 f2 → Env-nats s xs xs1 → RelId xs1 → Dom-nats (! B) s (f1 $$ xs) (f2 $$ xs1) 169 | $$-nats s {xs = []} {[]} dnat enat gr = dnat 170 | $$-nats s {xs = px ∷ xs} {py ∷ ys} dnat enat gr = $$-nats s (dnat _ id-i (gr zero) (proj₁ enat)) (proj₂ enat) (\ x → gr (suc x)) 171 | 172 | build-nats : ∀ {b Sg G G1 D} Ts (s : Sub< b > Sg G G1) (f : ∀ {S} (x : Ts ∋ S) → Dom Sg G D S) -> 173 | (g : ∀ {S} (x : Ts ∋ S) → Dom Sg G1 D S) -> 174 | (∀ {S} (x : Ts ∋ S) → Dom-nats S s (f x) (g x)) → Env-nats s (build f) (build g) 175 | build-nats [] s f g fg-nats = _ 176 | build-nats (x ∷ Ts) s f g fg-nats = fg-nats zero , build-nats Ts s (λ x₁ → f (suc x₁)) (λ x₁ → g (suc x₁)) (λ x₁ → fg-nats (suc x₁)) 177 | 178 | idEnv-nats : ∀ {D b Sg G1 G2} {f : Sub< b > Sg G1 G2} → Env-nats f (idEnv {D = D}) idEnv 179 | idEnv-nats = build-nats _ _ injv injv injv-nats 180 | 181 | -------------------------------------------------------------------------------- /Syntax/Type.agda: -------------------------------------------------------------------------------- 1 | module Syntax.Type where 2 | 3 | open import Data.Sum public using (_⊎_; module _⊎_) 4 | open _⊎_ public 5 | open import Data.Bool public 6 | 7 | open import Support.Equality 8 | open ≡-Reasoning 9 | open import Support.List public 10 | 11 | open import Injections 12 | 13 | postulate 14 | -- it'd more properly be a module parameter 15 | Base : Set 16 | 17 | 18 | -- Ty contains the types of our STLC: 19 | -- t1 → ... → tn → b is represented as [t1, ... , tn] ->> b. 20 | data Ty : Set where 21 | _->>_ : Fwd Ty → Base → Ty 22 | 23 | !_ : Base → Ty 24 | !_ B = !> ->> B 25 | 26 | infixl 40 _->>_ 27 | 28 | -- Typing contexts are snoc-lists of types: 29 | -- lambdas are cuter this way. 30 | Ctx : Set 31 | Ctx = Bwd Ty 32 | 33 | -- Meta-variables have their own context, rather than a function type. 34 | record MTy : Set where 35 | constructor _<<-_ 36 | field 37 | type : Base 38 | ctx : Ctx 39 | 40 | open MTy public 41 | infixr 40 _<<-_ 42 | 43 | -- Meta-Context to type meta-vars. 44 | MCtx : Set 45 | MCtx = Bwd MTy 46 | 47 | -- Tm Sg G D T represents lambda terms in beta-short, eta-long normal form with added meta-vars. 48 | -- The pattern condition is ensured by providing meta-vars arguments through injective renamings. 49 | -- 50 | -- Sg - context for top level constants i.e. signature 51 | -- G - context for meta-vars 52 | -- D - context for object-vars (or just vars) 53 | -- T - term's type 54 | -- 55 | -- We use the "spiny" representation where each var/constant is paired with as many arguments as its type requires. 56 | mutual 57 | data Tm<_> (b : Bool) (Sg : Ctx)(G : MCtx)(D : Ctx) : Ty → Set where 58 | con : ∀ {Ss B} → 59 | (c : Sg ∋ (Ss ->> B)) → (ts : Tms< b > Sg G D Ss) → Tm< b > Sg G D (! B) 60 | mvar : ∀ {Ss B} → 61 | (u : G ∋ (B <<- Ss)) → (j : Args b Sg G D Ss) → Tm< b > Sg G D (! B) 62 | var : ∀ {Ss B} → 63 | (x : D ∋ (Ss ->> B)) → (ts : Tms< b > Sg G D Ss) → Tm< b > Sg G D (! B) 64 | lam : ∀ {S Ss B} → 65 | (t : Tm< b > Sg G (D <: S) (Ss ->> B)) → Tm< b > Sg G D (S :> Ss ->> B) 66 | 67 | data Tms<_> (b : Bool) (Sg : Ctx)(G : MCtx)(D : Ctx) : Fwd Ty → Set where 68 | [] : Tms< b > Sg G D !> 69 | _∷_ : {S : Ty}{Ss : Fwd Ty} → 70 | (t : Tm< b > Sg G D S) → (ts : Tms< b > Sg G D Ss) → Tms< b > Sg G D (S :> Ss) 71 | 72 | Args : Bool -> (Sg : Ctx)(G : MCtx)(D : Ctx) -> Fwd Ty -> Set 73 | Args true Sg G D Ss = Inj Ss D 74 | Args false Sg G D Ss = Tms< false > Sg G D Ss 75 | 76 | module Tm = Tm<_> 77 | module Tms = Tms<_> 78 | 79 | Tm : (Sg : Ctx)(G : MCtx)(D : Ctx) -> Ty → Set 80 | Tm = Tm< true > 81 | 82 | Tms : (Sg : Ctx)(G : MCtx)(D : Ctx) -> Fwd Ty → Set 83 | Tms = Tms< true > 84 | 85 | Term<_> : (b : Bool)(Sg : Ctx)(G : MCtx)(DI : Ctx) (TI : Ty ⊎ Fwd Ty) → Set 86 | Term< b > Sg G D (inj₁ T) = Tm< b > Sg G D T 87 | Term< b > Sg G D (inj₂ Ts) = Tms< b > Sg G D Ts 88 | 89 | Term : (Sg : Ctx)(G : MCtx)(DI : Ctx) (TI : Ty ⊎ Fwd Ty) → Set 90 | Term = Term< true > 91 | 92 | -- ren makes Tm a functor over Inj 93 | mutual 94 | ren : ∀ {b Sg G D D0}{T : Ty} → Inj D D0 → Tm< b > Sg G D T → Tm< b > Sg G D0 T 95 | ren rho (con c ts) = con c (rens rho ts) 96 | ren rho (mvar f xs) = mvar f (renArgs rho xs) 97 | ren rho (var x xs) = var (rho $ x) (rens rho xs) 98 | ren rho (lam t) = lam (ren (cons rho) t) 99 | 100 | rens : ∀ {b Sg G D D0}{Ts : Fwd Ty} → Inj D D0 → Tms< b > Sg G D Ts → Tms< b > Sg G D0 Ts 101 | rens rho [] = [] 102 | rens rho (x ∷ ts) = ren rho x ∷ rens rho ts 103 | 104 | renArgs : ∀ {b Sg G D D0}{Ts : Fwd Ty} → Inj D D0 → Args b Sg G D Ts → Args b Sg G D0 Ts 105 | renArgs {true} rho xs = rho ∘i xs 106 | renArgs {false} rho xs = rens rho xs 107 | 108 | renT : ∀ {T Sg G D1 D2 b} -> (Inj D1 D2) -> Term< b > Sg G D1 T -> Term< b > Sg G D2 T 109 | renT {inj₁ _} i t = ren i t 110 | renT {inj₂ _} i ts = rens i ts 111 | 112 | mutual 113 | ren-∘ : ∀ {b Sg G1 D1 D2 D3 T} {j : Inj D2 D3} {i : Inj D1 D2} (t : Tm< b > Sg G1 D1 T) → ren (j ∘i i) t ≡ ren j (ren i t) 114 | ren-∘ (con c ts) = cong (con c) (rens-∘ ts) 115 | ren-∘ (mvar u j₁) = cong (mvar u) (renAs-∘ j₁) 116 | ren-∘ (var x ts) = cong₂ var (apply-∘ _ _) (rens-∘ ts) 117 | ren-∘ (lam t) = cong lam (trans (cong (λ k → ren k t) (cons-∘i _ _)) (ren-∘ t)) 118 | 119 | rens-∘ : ∀ {b Sg G1 D1 D2 D3 T} {j : Inj D2 D3} {i : Inj D1 D2} (t : Tms< b > Sg G1 D1 T) → rens (j ∘i i) t ≡ rens j (rens i t) 120 | rens-∘ [] = refl 121 | rens-∘ (t ∷ ts) = cong₂ _∷_ (ren-∘ t) (rens-∘ ts) 122 | 123 | renAs-∘ : ∀ {b Sg G1 D1 D2 D3 T} {j : Inj D2 D3} {i : Inj D1 D2} (t : Args b Sg G1 D1 T) → renArgs (j ∘i i) t ≡ renArgs j (renArgs i t) 124 | renAs-∘ {true} t = sym assoc-∘i 125 | renAs-∘ {false} t = rens-∘ t 126 | 127 | mutual 128 | ren-id : ∀ {b Sg G D T} (t : Tm< b > Sg G D T) → ren id-i t ≡ t 129 | ren-id (con c ts) = cong (con c) (rens-id ts) 130 | ren-id (mvar u j) = cong (mvar u) (renAs-id j) 131 | ren-id (var x ts) = cong₂ var (id-i$ x) (rens-id ts) 132 | ren-id (lam t) = cong lam (trans (cong (λ k → ren k t) cons-id) (ren-id t)) 133 | 134 | rens-id : ∀ {b Sg G D T} (t : Tms< b > Sg G D T) → rens id-i t ≡ t 135 | rens-id [] = refl 136 | rens-id (t ∷ ts) = cong₂ _∷_ (ren-id t) (rens-id ts) 137 | 138 | renAs-id : ∀ {b Sg G D T} (t : Args b Sg G D T) → renArgs id-i t ≡ t 139 | renAs-id {true} j = left-id j 140 | renAs-id {false} xs = rens-id xs 141 | -------------------------------------------------------------------------------- /Unification.agda: -------------------------------------------------------------------------------- 1 | module Unification where 2 | 3 | open import Support.Nat 4 | open import Data.Empty 5 | open import Data.Sum renaming (inj₁ to yes; inj₂ to no; map to map⊎) 6 | 7 | open import Support.Equality 8 | open ≡-Reasoning 9 | 10 | open import Injections 11 | open import MetaRens 12 | 13 | open import Syntax 14 | 15 | open import Unification.Specification 16 | open import Unification.MetaRens 17 | open import Unification.Injections 18 | open import Unification.OccursCheck 19 | open import Unification.Pruning 20 | open import Unification.Inversion 21 | 22 | 23 | flexSame : ∀ {Sg G D S} → (u : G ∋ S) → (i j : Inj (ctx S) D) → ∃⟦σ⟧ Max (Unifies {Sg} (Tm.mvar u i) (mvar u j)) 24 | flexSame {Sg} {G} {D} {B <<- Ss} u i j = _ , (DS σ , singleton-Decreasing e u (equalizer-Decr i j)) 25 | , [σ]Unifies[mvar[u,i],mvar[u,j]] 26 | , sup-σ 27 | where 28 | i,j⇒e = equalizer i j 29 | open Equalizer i,j⇒e 30 | 31 | σ = toSub (singleton u e) 32 | 33 | [σ]Unifies[mvar[u,i],mvar[u,j]] : ren i (σ _ u) ≡T ren j (σ _ u) 34 | [σ]Unifies[mvar[u,i],mvar[u,j]] rewrite thick-refl u = ≡-T (cong (mvar zero) commutes) 35 | 36 | sup-σ : Sup (Unifies (mvar u i) (mvar u j)) σ 37 | sup-σ {G'} ρ ρ-unifies = δ , ρ≡δ∘σ where 38 | 39 | ∃s[ren[e,s]≡ρ[u]] = forget (lift-equalizer i,j⇒e (ρ (B <<- Ss) u) ρ-unifies) 40 | 41 | δ : Sub< false > Sg (B <<- E ∷ G - u) G' 42 | δ ._ zero = proj₁ ∃s[ren[e,s]≡ρ[u]] 43 | δ S₁ (suc v) = ρ _ (thin u _ v) 44 | 45 | ρ≡δ∘σ : ρ ≡s (δ ∘s σ) 46 | ρ≡δ∘σ S v with thick u v 47 | ρ≡δ∘σ S .(thin u S w) | inj₁ (w , refl) = sym (ren-id (ρ S (thin u S w))) 48 | ρ≡δ∘σ .(B <<- Ss) .u | inj₂ refl` = sym (proj₂ ∃s[ren[e,s]≡ρ[u]]) 49 | 50 | 51 | flexRigid : ∀ {Sg G D S} (u : G ∋ S) (i : Inj (ctx S) D) (s : Tm Sg (G - u) D (! type S)) → Spec (mvar u i) (sub (thin-s u) s) 52 | flexRigid {Sg} {G} {S = S} u i s with prune i s 53 | ... | ((Pr ρ , decr , m) , ρ-sup) 54 | with invertTm i s ρ m 55 | ... | no NotInv = no λ {(_ , σ , σ-unifies) → 56 | let 57 | eq = begin 58 | ren i (σ S u) ≡⟨ T-≡ σ-unifies ⟩ 59 | subT σ (subT (thin-s u) s) ≡⟨ Sub∘.subT-∘ s ⟩ 60 | subT (σ ∘s thin-s u) s ∎ 61 | σ≤ρ = ρ-sup (σ ∘s thin-s u) (σ S u) (≡-T eq) 62 | 63 | in NotInv (proj₁ σ≤ρ) (σ S u , ≡-T (begin ren i (σ S u) ≡⟨ eq ⟩ 64 | subT (σ ∘s thin-s u) s ≡⟨ subT-ext (proj₂ σ≤ρ) s ⟩ 65 | subT (proj₁ σ≤ρ ∘s ρ) s ≡⟨ sym (Sub∘.subT-∘ s) ⟩ 66 | subT (proj₁ σ≤ρ) (subT ρ s) ∎))} 67 | 68 | ... | yes (t , ren[i,t]≡sub[ρ,s]) = yes 69 | (_ , (DS σ , inj₂ (rigid-decr u decr)) , 70 | ≡-T (begin 71 | ren i (σ _ u) ≡⟨ cong (ren i) σ[u]≡t ⟩ 72 | ren i t ≡⟨ ren[i,t]≡sub[ρ,s] ⟩ 73 | sub ρ s ≡⟨ sub-ext ρ≡σ∘thin[u] s ⟩ 74 | sub (σ ∘s thin-s u) s ≡⟨ sym (sub-∘ s) ⟩ 75 | sub σ (sub (thin-s u) s) ∎) , σ-sup ) 76 | where 77 | σ : Sub Sg G _ 78 | σ S v with thick u v 79 | σ S v | inj₁ (w , eq) = ρ _ w 80 | σ ._ .u | inj₂ refl` = t 81 | 82 | σ[u]≡t : σ _ u ≡ t 83 | σ[u]≡t rewrite thick-refl u = refl 84 | 85 | ρ≡σ∘thin[u] : ρ ≡s (σ ∘s thin-s u) 86 | ρ≡σ∘thin[u] S y rewrite thick-thin u y = sym (ren-id _) 87 | 88 | σ-sup : Sup (Unifies (mvar u i) (sub (thin-s u) s)) σ 89 | σ-sup ρ₁ ρ₁-unifies = δ , ρ₁≡δ∘σ where 90 | ren[i,ρ₁[u]]≡sub[ρ₁∘thin[u],s] = begin 91 | sub ρ₁ (mvar u i) ≡⟨ T-≡ ρ₁-unifies ⟩ 92 | sub ρ₁ (sub (thin-s u) s) ≡⟨ Sub∘.subT-∘ s ⟩ 93 | sub (ρ₁ ∘s thin-s u) s ∎ 94 | 95 | ρ₁∘thin[u]≤ρ = ρ-sup (ρ₁ ∘s thin-s u) (ρ₁ _ u) (≡-T ren[i,ρ₁[u]]≡sub[ρ₁∘thin[u],s]) 96 | δ = proj₁ ρ₁∘thin[u]≤ρ 97 | ρ₁∘thin[u]≡δ∘ρ = proj₂ ρ₁∘thin[u]≤ρ 98 | 99 | ρ₁≡δ∘σ : ρ₁ ≡s (δ ∘s σ) 100 | ρ₁≡δ∘σ S u₁ with thick u u₁ 101 | ρ₁≡δ∘σ S ._ | inj₁ (v , refl) = begin 102 | ρ₁ S (thin u S v) ≡⟨ sym (ren-id (ρ₁ S (thin u S v))) ⟩ 103 | sub ρ₁ (thin-s u S v) ≡⟨ ρ₁∘thin[u]≡δ∘ρ S v ⟩ 104 | sub δ (ρ S v) ∎ 105 | ρ₁≡δ∘σ ._ .u | inj₂ refl` = ren-inj i (ρ₁ _ u) (sub δ t) -- crucial use of injectivity to show 106 | (begin -- that we got the most general solution 107 | ren i (ρ₁ _ u) ≡⟨ ren[i,ρ₁[u]]≡sub[ρ₁∘thin[u],s] ⟩ 108 | sub (ρ₁ ∘s thin-s u) s ≡⟨ sub-ext ρ₁∘thin[u]≡δ∘ρ s ⟩ 109 | sub (δ ∘s ρ) s ≡⟨ sym (Sub∘.subT-∘ s) ⟩ 110 | sub δ (sub ρ s) ≡⟨ cong (sub δ) (sym ren[i,t]≡sub[ρ,s]) ⟩ 111 | sub δ (ren i t) ≡⟨ sub-nat t ⟩ 112 | ren i (sub δ t) ∎) 113 | 114 | 115 | flexAny : ∀ {Sg G D S} → (u : G ∋ S) → (i : Inj (ctx S) D) → (t : Tm Sg G D (! (type S))) → Spec (mvar u i) t 116 | flexAny u i t with check u t 117 | flexAny u i .(sub (thin-s u) s) | inj₁ (s , refl) = flexRigid u i s 118 | flexAny u i .(mvar u j) | inj₂ (G1 , j , [] , refl) = yes (flexSame u i j) 119 | flexAny u i .(wrap (d ∷ c) (mvar u j)) | inj₂ (G1 , j , d ∷ c , refl) = no λ {(D1 , s , eq) → 120 | No-Cycle (subD s d) (subC s c) (s _ u) i j 121 | (trans (T-≡ eq) (wrap-sub s (d ∷ c) (mvar u j)))} 122 | 123 | mutual 124 | unify : ∀ {Sg G D T} → (x y : Tm Sg G D T) → ∀ n -> n ≡ Ctx-length G -> Spec x y 125 | -- congruence and directly failing cases 126 | unify (con c xs) (con c₁ ys) n l with c ≅∋? c₁ 127 | unify (con c xs) (con c₁ ys) n l | no c≢c₁ = no (λ {(_ , _ , eq) → c≢c₁ (con-inj₁ eq)}) 128 | unify (con c xs) (con .c ys) n l | yes refl` = cong-spec (con c) (unifyTms xs ys n l) 129 | unify (var x xs) (var y ys) n l with x ≅∋? y 130 | unify (var x xs) (var y ys) n l | no x≢y = no (λ {(_ , _ , eq) → x≢y (var-inj₁ eq)}) 131 | unify (var x xs) (var .x ys) n l | yes refl` = cong-spec (var x) (unifyTms xs ys n l) 132 | unify (lam x) (lam y) n l = cong-spec lam {x} {y} (unify x y n l) 133 | unify (con _ _) (var _ _) n l = no λ {(_ , _ , ())} 134 | unify (var _ _) (con _ _) n l = no λ {(_ , _ , ())} 135 | 136 | -- flexible cases 137 | unify (mvar u i) t _ l = flexAny u i t 138 | unify t (mvar u i) _ l = spec-comm (mvar u i) t (flexAny u i t) 139 | 140 | unifyTms : ∀ {Sg G D Ts} → (x y : Tms Sg G D Ts) → ∀ n -> n ≡ Ctx-length G -> Spec x y 141 | unifyTms [] [] _ _ = yes (∃σMax[Unifies[x,x]] []) 142 | unifyTms (s ∷ xs) (t ∷ ys) n l 143 | with unify s t n l 144 | ... | no ¬unify[s,t] = no λ {(_ , ρ , eq ∷ _) → ¬unify[s,t] (_ , ρ , eq)} 145 | ... | yes (_ , σ , eq , sup) 146 | with under σ unifyTms xs ys n l 147 | ... | no ¬unify[σxs,σys] = no λ {(_ , σ1 , eqt ∷ eqts) → ¬unify[σxs,σys] (shift eqts under ⟦ σ ⟧ by sup σ1 eqt) } 148 | ... | yes (_ , σ1 , eq1 , sup1) = yes (_ , (σ1 ∘ds σ) , optimist s t xs ys ⟦ σ ⟧ ⟦ σ1 ⟧ (eq , sup) (eq1 , sup1)) 149 | 150 | -- termination overhead 151 | under_unifyTms : ∀ {Sg G D Ts} -> 152 | ∀ {G1} (σ : DSub Sg G G1) -> (xs ys : Tms Sg G D Ts) -> ∀ n -> n ≡ Ctx-length G -> Spec (subs ⟦ σ ⟧ xs) (subs ⟦ σ ⟧ ys) 153 | under (DS σ , inj₁ (G~G1 , σ-is-iso)) unifyTms xs ys n l = Spec[xs,ys]⇒Spec[σxs,σys] σ G~G1 σ-is-iso (unifyTms xs ys n l) 154 | under (DS σ , inj₂ G>G1 ) unifyTms xs ys n n≡G = under-not-iso σ unifyTms xs ys n (cast n≡G G>G1) 155 | 156 | under-not-iso_unifyTms : ∀ {Sg G D Ts} -> ∀ {G1} (σ : Sub Sg G G1) -> (xs ys : Tms Sg G D Ts) -> 157 | ∀ n -> n >′ Ctx-length G1 -> Spec (subs σ xs) (subs σ ys) 158 | under-not-iso_unifyTms σ xs ys .(suc n) (>′-refl {n} n≡G1) = unifyTms (subs σ xs) (subs σ ys) n n≡G1 159 | under-not-iso_unifyTms σ xs ys .(suc n) (>′-step {n} n>G1) = under-not-iso σ unifyTms xs ys n n>G1 160 | 161 | 162 | -- unifyTms (subs σ xs) (subs σ ys) n z 163 | Unify : ∀ {Sg G D T} → (s t : Tm Sg G D T) → ∃σ-pat Max (Unifies s t) ⊎ ¬ ∃σ Unifies s t 164 | Unify x y = map⊎ (λ {(G1 , σ , σ-max) → G1 , ⟦ σ ⟧ , σ-max}) (λ ¬p → ¬p) (unify x y _ refl) 165 | -------------------------------------------------------------------------------- /Unification/Injections.agda: -------------------------------------------------------------------------------- 1 | module Unification.Injections where 2 | 3 | open import Data.Product 4 | open import Data.Empty 5 | 6 | open import Support.Equality 7 | open ≡-Reasoning 8 | 9 | open import Injections 10 | import Category 11 | 12 | module _ {A : Set} where 13 | open Category (List A) Inj _∘i_ id-i _≡_ public 14 | 15 | -- The category of injective maps between finite sets has all the 16 | -- pullbacks and equalizers, "pullback" and "equalizer" compute 17 | -- them. 18 | -- 19 | -- In this category both Equalizers and Pullbacks are ways to select 20 | -- the largest subset of the input where the maps agree, the 21 | -- difference is that Pullbacks allow for mediating permutations while 22 | -- Equalizers don't. 23 | -- 24 | -- The implementation takes a bit of record field's shuffling but it's straightforward. 25 | 26 | abstract 27 | e$u≡m : ∀ {A : Set}{S T : List A}{f g : Inj S T} -> (equ : Equalizer f g) -> let open Equalizer equ in 28 | (a : A) (m : S ∋ a) → f $ m ≡ g $ m → Σ (E ∋ a) (λ u → m ≡ e $ u) 29 | e$u≡m equ a m eq = 30 | let 31 | open Equalizer equ 32 | [m] = m ∷ [] [ _ ] 33 | u = universal [m] (ext-∘ (λ {.a zero → eq; x (suc ())})) 34 | in u $ zero , sym (trans (sym (apply-∘ e u)) (cong (λ f → f $ zero) e∘universal≡m)) 35 | 36 | p$u≡q : ∀ {A : Set} {X Y Z : List A}(f : Inj X Z)(g : Inj Y Z) -> (p : Pullback f g) -> let open Pullback p in 37 | (a : A) (q₂ : Y ∋ a) (q₁ : X ∋ a) → f $ q₁ ≡ g $ q₂ → Σ (P ∋ a) (λ u → p₁ $ u ≡ q₁ × p₂ $ u ≡ q₂) 38 | p$u≡q f g p a q₂ q₁ eq = 39 | let 40 | open Pullback p 41 | [q₁] = q₁ ∷ [] [ _ ] 42 | [q₂] = q₂ ∷ [] [ _ ] 43 | u : Inj (a ∷ []) P 44 | u = universal [q₁] [q₂] (ext-∘ (λ {._ zero → eq; _ (suc ())})) 45 | in u $ zero , 46 | trans (sym (apply-∘ p₁ u)) (cong (λ f₁ → f₁ $ zero) p₁∘universal≡q₁) , 47 | trans (sym (apply-∘ p₂ u)) (cong (λ f₁ → f₁ $ zero) p₂∘universal≡q₂) 48 | 49 | 50 | []-unique : ∀ {A : Set}{Q : List A}{f g : Inj [] Q} -> f ≡ g 51 | []-unique {f = []}{[]} = refl 52 | 53 | 54 | weak-equalizer : ∀ {A : Set}{a : A}{S T : List A}(f g : Inj S T){E e} -> IsEqualizer f g E e -> IsEqualizer (weak {x = a} f) (weak g) E e 55 | weak-equalizer {a = a} f g {E} {e} equ = 56 | record { 57 | commutes = trans (Inj-thin-∘i zero f e) (trans (cong (Inj-thin zero) commutes) 58 | (sym (Inj-thin-∘i zero g e))); 59 | universal = λ m commutes₁ → universal m 60 | (Inj-thin-inj zero (f ∘i m) (g ∘i m) (trans (sym (Inj-thin-∘i zero f m)) (trans commutes₁ (Inj-thin-∘i zero g m)))); 61 | universal-unique = universal-unique; 62 | e∘universal≡m = e∘universal≡m } 63 | where 64 | open IsEqualizer equ 65 | 66 | 67 | _→zero∷e⇒i∷f,i∷g : ∀ {A : Set}{a : A}{S T : List A}{v : T ∋ a}{f g : Inj S T}{pf pg}{E e} -> 68 | IsEqualizer f g E e -> IsEqualizer (v ∷ f [ pf ]) (v ∷ g [ pg ]) (a ∷ E) (zero ∷[] e) 69 | _→zero∷e⇒i∷f,i∷g {a = a} {v = u} {f} {g} {pf} {pg} {E} {e} equ = 70 | record { 71 | commutes = ext-∘ aux; 72 | universal = λ m commutes₁ → proj₁ (uni m commutes₁); 73 | universal-unique = λ {Q} {m} {commutes₁} u₁ e∘u≡m → ∘i-inj (zero ∷[] e) u₁ _ (trans e∘u≡m (sym (proj₂ (uni m commutes₁)))); 74 | e∘universal≡m = \ {Q} {m} {comm} -> proj₂ (uni m comm) } 75 | where 76 | open IsEqualizer equ 77 | aux : (x : _) (v : a ∷ E ∋ x) → (u ∷ f [ pf ]) $ ((zero ∷[] e) $ v) ≡ (u ∷ g [ pg ]) $ ((zero ∷[] e) $ v) 78 | aux .a zero = refl 79 | aux x (suc v) rewrite Inj-thin-$ {x = a} zero e v = ∘-ext commutes _ v 80 | uni : {Q : _} (m : Inj Q _) → (u ∷ f [ pf ]) ∘i m ≡ (u ∷ g [ pg ]) ∘i m → Σ (Inj Q (a ∷ E)) (λ z → (zero ∷[] e) ∘i z ≡ m ) 81 | uni = 82 | Equ-universal-quote (u ∷ f [ pf ]) (u ∷ g [ pg ]) (zero ∷[] e) 83 | (∋-case (λ x → zero , refl) 84 | (λ a₁ y x → 85 | let rec = e$u≡m (Equ _ , _ , equ) a₁ y x 86 | in suc (proj₁ rec) , 87 | trans (cong suc (proj₂ rec)) 88 | (sym (Inj-thin-$ zero e (proj₁ (e$u≡m (Equ _ , _ , equ) a₁ y x)))))) 89 | 90 | 91 | cons-equalizer : ∀ {A : Set}{a : A}{S T : List A}(f g : Inj S T) -> Equalizer f g -> Equalizer (cons {x = a} f) (cons g) 92 | cons-equalizer {A} {a} f g (Equ E , e , equ) = Equ _ , _ , (weak-equalizer f g equ) →zero∷e⇒i∷f,i∷g 93 | 94 | eta-empty : ∀ {A : Set}{Q : List A}(f g : Inj Q []) -> f ≡ g 95 | eta-empty [] [] = refl 96 | eta-empty (() ∷ f [ pf ]) g 97 | 98 | empty-equalizer : ∀ {A : Set}{xs : List A} -> IsEqualizer {A}{[]}{xs} [] [] [] [] 99 | empty-equalizer {A} = record { 100 | commutes = refl; 101 | universal = λ m commutes → m; 102 | universal-unique = λ u e∘u≡m → eta-empty _ _; 103 | e∘universal≡m = eta-empty _ _ } 104 | 105 | _,_→weake⇒i∷f,j∷g : ∀ {A : Set}{a : A}{S T : List A}{u v : T ∋ a} -> u ≢ v -> 106 | ∀ {f g : Inj S T}{pf pg}{E e} -> IsEqualizer f g E e -> IsEqualizer (u ∷ f [ pf ]) (v ∷ g [ pg ]) E (weak e) 107 | _,_→weake⇒i∷f,j∷g {a = a} {u = u}{v} u≢v {f} {g} {pf} {pg} {E} {e} equ = 108 | record { 109 | commutes = ext-∘ aux; 110 | universal = λ m commutes₁ → proj₁ (uni m commutes₁); 111 | universal-unique = λ {Q} {m} {commutes₁} u₁ e∘u≡m → ∘i-inj (weak e) u₁ _ (trans e∘u≡m (sym (proj₂ (uni m commutes₁)))); 112 | e∘universal≡m = \ {Q} {m} {comm} -> proj₂ (uni m comm) } 113 | where 114 | open IsEqualizer equ 115 | aux : (x : _) (v₁ : E ∋ x) → (u ∷ f [ pf ]) $ (weak e $ v₁) ≡ (v ∷ g [ pg ]) $ (weak e $ v₁) 116 | aux x v₁ rewrite Inj-thin-$ {x = a} zero e v₁ = ∘-ext commutes _ v₁ 117 | 118 | uni : {Q : _} (m : Inj Q _) → (u ∷ f [ pf ]) ∘i m ≡ (v ∷ g [ pg ]) ∘i m → Σ (Inj Q E) (λ z → (weak e) ∘i z ≡ m ) 119 | uni = Equ-universal-quote (u ∷ f [ pf ]) (v ∷ g [ pg ]) (weak e) (∋-case (λ u≡v → ⊥-elim (u≢v u≡v)) 120 | (λ a₁ y x → 121 | let rec : _ 122 | rec = e$u≡m (Equ _ , _ , equ) a₁ y x 123 | in proj₁ rec , 124 | trans (cong suc (proj₂ rec)) 125 | (sym (Inj-thin-$ zero e (proj₁ (e$u≡m (Equ _ , _ , equ) a₁ y x)))))) 126 | 127 | equalizer : ∀ {A : Set}{S T : List A}(f g : Inj S T) -> Equalizer f g 128 | equalizer [] [] = Equ [] , [] , empty-equalizer 129 | equalizer (i ∷ f [ pf ]) (j ∷ g [ pf₁ ]) with i ≅∋? j | equalizer f g 130 | equalizer (i ∷ f [ pf ]) (.i ∷ g [ pf₁ ]) | yes refl` | Equ E , e , e⇒f,g = Equ _ ∷ E , zero ∷[] e , e⇒f,g →zero∷e⇒i∷f,i∷g 131 | equalizer (i ∷ f [ pf ]) (j ∷ g [ pf₁ ]) | no ¬p | Equ E , e , e⇒f,g = Equ E , weak e , i≢j , e⇒f,g →weake⇒i∷f,j∷g 132 | where 133 | i≢j : i ≢ j 134 | i≢j = (λ x → ¬p (refl , ≡-to-≅ x)) 135 | 136 | empty-pullback : ∀ {A : Set} {X Z : List A} → {f : Inj X Z} → IsPullback f [] [] [] [] 137 | empty-pullback {A} {X} {Z} {f} = record { 138 | commutes = []-unique; 139 | universal = λ q₁ q₂ commutes → q₂; 140 | universal-unique = λ u p₁∘u≡q₁ p₂∘u≡q₂ → eta-empty u _; 141 | p₁∘universal≡q₁ = \ {Q q₁ q₂ commutes} -> proof {Q} {q₁} {q₂} {commutes}; 142 | p₂∘universal≡q₂ = eta-empty _ _ } 143 | where 144 | proof : {Q : List A} {q₁ : Inj Q X} {q₂ : Inj Q []} 145 | {commutes : f ∘i q₁ ≡ [] ∘i q₂} → 146 | [] ∘i q₂ ≡ q₁ 147 | proof {.[]} {q₁} {[]} = []-unique 148 | proof {._} {q₁} {() ∷ q₂ [ pf ]} 149 | 150 | ∈-pullback : ∀ {A : Set} {X Y Z P : List A} {f : Inj X Z}{g : Inj Y Z}{p₁ p₂} -> IsPullback f g P p₁ p₂ -> 151 | ∀ {a x pf x∉p₁ } → IsPullback f ((f $ x) ∷ g [ pf ]) (a ∷ P) (x ∷ p₁ [ x∉p₁ ]) (zero ∷[] p₂) 152 | ∈-pullback {A}{X}{Y}{Z}{P}{f = f} {g} {p₁}{p₂}pull{a}{x}{pf}{x∉p₁} = record { 153 | commutes = ext-∘ aux; 154 | universal = λ q₁ q₂ commutes₁ → proj₁ (uni q₁ q₂ commutes₁); 155 | universal-unique = λ u p₁∘u≡q₁ p₂∘u≡q₂ → 156 | ∘i-inj (x ∷ p₁ [ x∉p₁ ]) u (proj₁ (uni _ _ _)) (trans p₁∘u≡q₁ (proj₁ (proj₂ (uni _ _ _)))); 157 | p₁∘universal≡q₁ = sym (proj₁ (proj₂ (uni _ _ _))); 158 | p₂∘universal≡q₂ = sym (proj₂ (proj₂ (uni _ _ _))) } 159 | where 160 | open IsPullback pull 161 | abstract 162 | uni : ∀ {Q} -> (q₁ : Inj Q X) (q₂ : Inj Q (a ∷ Y)) -> f ∘i q₁ ≡ ((f $ x) ∷ g [ pf ]) ∘i q₂ -> 163 | ∃ \ u -> q₁ ≡ (x ∷ p₁ [ x∉p₁ ]) ∘i u × q₂ ≡ (zero ∷[] p₂) ∘i u 164 | uni {Q} q₁ q₂ eq = Pull-universal-quote f ((f $ x) ∷ g [ pf ]) (x ∷ p₁ [ x∉p₁ ]) (zero ∷[] p₂) 165 | (∋-case (λ x₁ x₂ → zero , ((injective f x x₁ (sym x₂)) , refl)) (λ a₁ y x₁ x₂ → 166 | let p : Σ _ _; p = p$u≡q f g (Pull P , p₁ , p₂ , pull) a₁ y x₁ x₂ 167 | in (suc (proj₁ p)) , (proj₁ (proj₂ p)) , (trans (iso2 _ _ (proj₁ p)) (cong suc (proj₂ (proj₂ p)))))) 168 | q₁ q₂ eq 169 | 170 | aux : ∀ t (v : _ ∋ t) -> f $ ((x ∷ p₁ [ x∉p₁ ]) $ v) ≡ 171 | ((f $ x) ∷ g [ pf ]) $ 172 | ((zero ∷[] p₂) $ v) 173 | aux .a zero = refl 174 | aux t (suc v) = trans (∘-ext (IsPullback.commutes pull) _ v) (cong (_$_ ((f $ x) ∷ g [ pf ])) (sym (Inj-thin-$ zero p₂ v))) 175 | 176 | 177 | comm-pullback : ∀ {A : Set} {X Y Z P : List A} {f : Inj X Z}{g : Inj Y Z}{p₁ p₂} -> IsPullback f g P p₁ p₂ -> 178 | IsPullback g f P p₂ p₁ 179 | comm-pullback pull = record { 180 | commutes = sym commutes; 181 | universal = λ q₁ q₂ commutes₁ → universal q₂ q₁ (sym commutes₁); 182 | universal-unique = λ u p₁∘u≡q₁ p₂∘u≡q₂ → universal-unique u p₂∘u≡q₂ p₁∘u≡q₁; 183 | p₁∘universal≡q₁ = p₂∘universal≡q₂; 184 | p₂∘universal≡q₂ = p₁∘universal≡q₁ } 185 | where 186 | open IsPullback pull 187 | 188 | 189 | ∉-pullback : ∀ {A : Set} {X Y Z P : List A} {f : Inj X Z}{g : Inj Y Z}{p₁ p₂} -> IsPullback f g P p₁ p₂ -> 190 | ∀ {a} {i : Z ∋ a} {pf} → i ∉Im f -> IsPullback f (i ∷ g [ pf ]) P p₁ (weak p₂) 191 | ∉-pullback {A}{X}{Y}{Z}{f = f} {g = g} {p₁}{p₂}pull {a} {i} {pf} i∉f = record { 192 | commutes = trans commutes (ext-∘ (λ x v → 193 | cong (_$_ (_ ∷ g [ _ ])) (sym (Inj-thin-$ zero p₂ v)))); 194 | universal = λ q₁ q₂ commutes₁ → proj₁ (uni q₁ q₂ commutes₁); 195 | universal-unique = λ u p₁∘u≡q₁ p₂∘u≡q₂ → ∘i-inj p₁ u (proj₁ (uni _ _ _)) (trans p₁∘u≡q₁ (proj₁ (proj₂ (uni _ _ _)))); 196 | p₁∘universal≡q₁ = sym (proj₁ (proj₂ (uni _ _ _))); 197 | p₂∘universal≡q₂ = sym (proj₂ (proj₂ (uni _ _ _))) } 198 | where 199 | open IsPullback pull 200 | abstract 201 | uni : ∀ {Q} -> (q₁ : Inj Q X) (q₂ : Inj Q (a ∷ Y)) -> f ∘i q₁ ≡ (i ∷ g [ pf ]) ∘i q₂ -> ∃ \ u -> q₁ ≡ p₁ ∘i u × q₂ ≡ (weak p₂) ∘i u 202 | uni {Q} q₁ q₂ eq = Pull-universal-quote f (i ∷ g [ pf ]) p₁ (weak p₂) (∋-case (λ x x₁ → ⊥-elim (i∉f x (sym x₁))) 203 | (λ a₁ y x x₁ → let p : Σ _ _ 204 | p = p$u≡q f g (Pull _ , p₁ , p₂ , pull) a₁ y x x₁ 205 | in (proj₁ p) , ((proj₁ (proj₂ p)) , trans (iso2 _ _ (proj₁ p)) (cong suc (proj₂ (proj₂ p)))))) 206 | q₁ q₂ eq 207 | 208 | pullback : ∀ {A : Set} {X Y Z : List A} → (f : Inj X Z)(g : Inj Y Z) → Pullback f g 209 | pullback f [] = Pull [] , [] , [] , empty-pullback 210 | pullback f (i ∷ g [ pf ]) with invert f i | pullback f g 211 | pullback f (.(f $ x) ∷ g [ pf ]) | yes (x , refl) | Pull P , p₁ , p₂ , p₁,p₂⇒f,g = 212 | Pull (_ ∷ P) , (x ∷ p₁ [ x∉p₁ ]) , (zero ∷[] p₂) , ∈-pullback p₁,p₂⇒f,g 213 | where 214 | open IsPullback p₁,p₂⇒f,g 215 | x∉p₁ : x ∉ p₁ 216 | x∉p₁ = ∉Im-∉ p₁ x λ b x≡k$b → ∉-∉Im g (f $ x) pf (p₂ $ b) 217 | (begin f $ x ≡⟨ cong (_$_ f) x≡k$b ⟩ 218 | f $ (p₁ $ b) ≡⟨ ∘-ext commutes _ b ⟩ 219 | g $ (p₂ $ b) ∎) 220 | pullback f (i ∷ g [ pf ]) | no i∉f | Pull P , p₁ , p₂ , p₁,p₂⇒f,g = 221 | Pull P , p₁ , weak p₂ , ∉-pullback p₁,p₂⇒f,g (λ b x → i∉f (b , (sym x))) 222 | 223 | weak-pullback : ∀ {A : Set} {X Y Z P : List A} {f : Inj X Z}{g : Inj Y Z}{p₁ p₂} -> IsPullback f g P p₁ p₂ -> 224 | ∀ {a : A} -> IsPullback (weak {x = a} f) (weak g) P p₁ p₂ 225 | weak-pullback pull = record { 226 | commutes = trans (Inj-thin-∘i zero _ _) 227 | (trans (cong (Inj-thin zero) commutes) 228 | (sym (Inj-thin-∘i zero _ _))); 229 | universal = λ q₁ q₂ commutes₁ → universal q₁ q₂ 230 | (Inj-thin-inj zero (_ ∘i q₁) (_ ∘i q₂) 231 | (trans (sym (Inj-thin-∘i zero _ q₁)) 232 | (trans commutes₁ (Inj-thin-∘i zero _ q₂)))); 233 | universal-unique = universal-unique; 234 | p₁∘universal≡q₁ = p₁∘universal≡q₁; 235 | p₂∘universal≡q₂ = p₂∘universal≡q₂ } 236 | where 237 | open IsPullback pull 238 | 239 | cons-pullback : ∀ {A : Set}{a : A}{X Y Z : List A}(f : Inj X Z)(g : Inj Y Z) -> Pullback f g -> Pullback (cons {x = a} f) (cons g) 240 | cons-pullback {A} {a} f g (Pull P , p₁ , p₂ , p₁,p₂⇒f,g) = Pull a ∷ P , cons p₁ , cons p₂ , 241 | ∈-pullback (comm-pullback (∉-pullback (comm-pullback (weak-pullback p₁,p₂⇒f,g)) (∉-∉Im (weak g) zero (v∉Inj-thinv zero g)))) 242 | 243 | 244 | -------------------------------------------------------------------------------- /Unification/Inversion.agda: -------------------------------------------------------------------------------- 1 | module Unification.Inversion where 2 | 3 | open import Relation.Binary.PropositionalEquality 4 | open import Data.Empty 5 | open import Data.Sum renaming (map to map⊎) 6 | open import Data.Sum renaming (inj₁ to yes; inj₂ to no) 7 | 8 | open import Injections 9 | 10 | open import Syntax 11 | open import Unification.Pruning 12 | 13 | 14 | NotInv : ∀ {Sg G D D' T} (i : Inj D D') (t : Term Sg G D' T) → Set 15 | NotInv {Sg} {G} i t = ∀ {G1} (σ : Sub< false > Sg G G1) -> ¬ ∃ \ s -> renT i s ≡T subT σ t 16 | 17 | map-NI : ∀ {Sg G DI D T D' T'}{i : Inj D' DI}{t : Term Sg G D T} 18 | (d : DTm Sg G (DI , T') (D , T)) → NotInv (wrapD-Inj d i) t → NotInv i (wrapD d t) 19 | map-NI lam notinv σ (lam s , lam eq) = notinv σ (s , eq) 20 | map-NI (head ts) notinv σ (s ∷ ss , eq ∷ eq') = notinv σ (s , eq) 21 | map-NI (tail t) notinv σ (s ∷ ss , eq ∷ eq') = notinv σ (ss , eq') 22 | map-NI (con c) notinv σ (con .c ss , con refl eq) = notinv σ (ss , eq) 23 | map-NI (var ._) notinv σ (var x ss , var refl eq) = notinv σ (ss , eq) 24 | map-NI (con c) notinv σ (var x ss , ()) 25 | map-NI (con c) notinv σ (mvar u j , ()) 26 | map-NI (var x) notinv σ (con c ss , ()) 27 | map-NI (var x) notinv σ (mvar u j , ()) 28 | 29 | NI-var : ∀ {Sg G D D1 Ss B} {i : Inj D D1} {ts : Tms Sg G D1 Ss} {x : D1 ∋ (Ss ->> B)} 30 | → ¬ ∃ (λ y → i $ y ≡ x) → NotInv {T = inj₁ _} i (var x ts) 31 | NI-var ¬∃y[iy≡x] σ (con c ts , ()) 32 | NI-var ¬∃y[iy≡x] σ (mvar u j , ()) 33 | NI-var ¬∃y[iy≡x] σ (var y _ , var iy≡x _) = ¬∃y[iy≡x] (y , iy≡x) 34 | 35 | mutual 36 | invertTm' : ∀ {Sg G Ss D T} (i : Inj Ss D) (t : Tm Sg G D T) → 37 | AllMV∈ i t → i ⁻¹ t ⊎ NotInv i t 38 | invertTm' i (con c ts) (con m) = map⊎ (con c) (map-NI (con c)) (invertTm's i ts m) 39 | invertTm' i (mvar v h) (mvar (k , comm)) = yes (mvar v (k , comm)) 40 | invertTm' i (var x ts) (var m) 41 | with invert i x | invertTm's i ts m 42 | ... | yes (y , iy≡x) | yes i⁻¹ts = yes (var y iy≡x i⁻¹ts) 43 | ... | _ | no NI[ts] = no (map-NI (var x) NI[ts]) 44 | ... | no ¬∃y[iy≡x] | _ = no (NI-var ¬∃y[iy≡x]) 45 | invertTm' i (lam t) (lam m) = map⊎ lam (map-NI {t = t} lam) (invertTm' (cons i) t m) 46 | 47 | invertTm's : ∀ {Sg G Ss D T} (i : Inj Ss D) (t : Tms Sg G D T) → 48 | AllMV∈ i t → i ⁻¹ t ⊎ NotInv i t 49 | invertTm's i [] m = yes [] 50 | invertTm's i (t ∷ ts) (mt ∷ mts) 51 | with invertTm' i t mt | invertTm's i ts mts 52 | ... | yes p | yes ps = yes (p ∷ ps) 53 | ... | yes p | no ¬ps = no (map-NI (tail t) ¬ps) 54 | ... | no ¬p | _ = no (map-NI {t = t} (head ts) ¬p) 55 | 56 | invertTm : ∀ {Sg G G1 Ss D T} (i : Inj Ss D) → (t : Tm Sg G D T) → (ρ : Sub Sg G G1) → ρ / t ∈ i 57 | → (∃ \ s → ren i s ≡ sub ρ t) ⊎ NotInv i (sub ρ t) 58 | invertTm i t ρ m = map⊎ forget (λ x → x) (invertTm' i (sub ρ t) m) 59 | 60 | -------------------------------------------------------------------------------- /Unification/OccursCheck.agda: -------------------------------------------------------------------------------- 1 | module Unification.OccursCheck where 2 | 3 | open import Data.Nat 4 | open import Data.Empty 5 | open import Data.Unit 6 | open import Data.Sum renaming (inj₁ to no; inj₂ to yes) 7 | 8 | open import Support.Equality 9 | 10 | open import Injections hiding (Dec) 11 | 12 | open import Syntax 13 | 14 | _OccursIn_ : ∀ {Sg G D T S} (u : G ∋ S) (t : Term Sg G D T) → Set 15 | _OccursIn_ {Sg} {G} {D} {T} {S} u t 16 | = ∃ \ D' → ∃ \ (j : Inj (ctx S) D') → ∃ \ (C : Context Sg G (D , T) (D' , inj₁ ([] ->> type S))) → wrap C (mvar u j) ≡ t 17 | where open import Data.Sum 18 | 19 | _NotOccursIn_ : ∀ {Sg G D T S} (u : G ∋ S) (t : Term Sg G D T) → Set 20 | u NotOccursIn t = ∃ \ (s : Term _ _ _ _) → subT (thin-s u) s ≡ t 21 | 22 | Dec_OccursIn_ : ∀ {Sg G D T S} (u : G ∋ S) (t : Term Sg G D T) → Set 23 | Dec u OccursIn t = u NotOccursIn t ⊎ u OccursIn t 24 | 25 | map-occ : ∀ {Sg G S D T D' T'}{u : G ∋ S}{t : Term Sg G D T} (d : DTm Sg G (D' , T') (D , T)) → u OccursIn t → u OccursIn wrapD d t 26 | map-occ d (Dj , j , C , eq) = (Dj , j , (d ∷ C) , cong (wrapD d) eq) 27 | 28 | _∙_ : ∀ {Sg G S D T D' T'}{u : G ∋ S}{t : Term Sg G D T} (d : DTm Sg (G - u) (D' , T') (D , T)) 29 | → Dec u OccursIn t → Dec u OccursIn wrapD (subD (thin-s u) d) t 30 | _∙_ {u = u} d (yes occ) = yes (map-occ (subD (thin-s u) d) occ) 31 | _∙_ {u = u} d (no (s , eq)) = no (wrapD d s , trans (wrapD-sub _ d s) (cong (wrapD (subD (thin-s u) d)) eq)) 32 | 33 | -- "check" decides whether u occurs in t. 34 | -- Note: without the pattern condition we'd have to consider 35 | -- the occurrences' "rigid"-ness. 36 | mutual 37 | check : ∀ {Sg G D T S} (u : G ∋ S) (t : Tm Sg G D T) → Dec u OccursIn t 38 | check u (con c ts) = con c ∙ checks u ts 39 | check u (var x ts) = var x ∙ checks u ts 40 | check u (lam t) = lam ∙ check u t 41 | check u (mvar w j) with thick u w 42 | ... | no (z , eq) = no (mvar z j , cong₂ mvar eq (right-id j)) 43 | check u (mvar .u j) | yes refl` = yes (_ , j , [] , refl) 44 | 45 | checks : ∀ {Sg G D Ts S} (u : G ∋ S) (ts : Tms Sg G D Ts) → Dec u OccursIn ts 46 | checks u [] = no ([] , refl) 47 | checks u (t ∷ ts) 48 | with check u t | checks u ts 49 | ... | yes x | _ = yes (map-occ (head ts) x) 50 | ... | _ | yes xs = yes (map-occ (tail t) xs) 51 | ... | no x | no xs = no (mapΣ₂ _∷_ (cong₂ _∷_) x xs) 52 | 53 | -------------------------------------------------------------------------------- /Unification/Pruning.agda: -------------------------------------------------------------------------------- 1 | module Unification.Pruning where 2 | 3 | open import Data.Nat using (_≥_) 4 | open import Data.Empty 5 | open import Data.Sum 6 | 7 | open import Support.Equality 8 | open ≡-Reasoning 9 | 10 | open import Injections 11 | open import MetaRens 12 | 13 | open import Syntax 14 | 15 | open import Unification.Specification 16 | open import Unification.MetaRens 17 | open import Unification.Injections 18 | open import Unification.Pruning.Epi-Decr 19 | 20 | data AllMV∈ {Sg : Ctx} {G : MCtx} {D0 D : Ctx} (i : Inj D0 D) : ∀ {T} → Term Sg G D T → Set where 21 | [] : AllMV∈ i {inj₂ []} [] 22 | _∷_ : ∀ {S Ss t ts} → (m : AllMV∈ i {inj₁ S} t) → (ms : AllMV∈ i {inj₂ Ss} ts) → AllMV∈ i {inj₂ (S ∷ Ss)} (t ∷ ts) 23 | con : ∀ {Ss B c ts} → (ms : AllMV∈ i {inj₂ Ss} ts) → AllMV∈ i {inj₁ (! B)} (con c ts) 24 | var : ∀ {Ss B x ts} → (ms : AllMV∈ i {inj₂ Ss} ts) → AllMV∈ i {inj₁ (! B)} (var x ts) 25 | lam : ∀ {S Ss B t} → (m : AllMV∈ (cons i) {inj₁ (Ss ->> B)} t) → AllMV∈ i {inj₁ ((S ∷ Ss) ->> B)} (lam t) 26 | mvar : ∀ {Ss B} {v : G ∋ (B <<- Ss)}{h} → (∃ \ k → i ∘i k ≡ h) → AllMV∈ i {inj₁ (! B)} (mvar v h) 27 | 28 | -- s / t ∈ i holds iff all the variables appearing as arguments to 29 | -- meta-vars in (sub s t) are present in the image of i 30 | _/_∈_ : ∀ {Sg G1 G2 D1 D2 T} → Sub Sg G1 G2 → Term Sg G1 D2 T → Inj D1 D2 → Set 31 | _/_∈_ s t i = AllMV∈ i (subT s t) 32 | 33 | -- (\ s -> s / t ∈ i) is closed under post-composition, we'll need this in the _∷_ case of prune. 34 | _/_∈_-∘Closed : ∀ {Sg G1 G2 G3 D1 D2 T} (f : Sub Sg G2 G3) {g : Sub Sg G1 G2} {i : Inj D1 D2} → 35 | ∀ {t : Term Sg G1 D2 T} → g / t ∈ i → (f ∘s g) / t ∈ i 36 | _/_∈_-∘Closed f m = subst (AllMV∈ _) (subT-∘ _) (go f m) where 37 | mutual 38 | go2 : ∀ {Sg G X Y Z T} {i : Inj Y Z} {h : Inj X Z} → (∃ \ k → i ∘i k ≡ h) → 39 | (t : Tm Sg G X T) → AllMV∈ i (renT h t) 40 | go2 le (con c ts) = con (go2s le ts) 41 | go2 le (var x ts) = var (go2s le ts) 42 | go2 (k , i∘k≡h) (lam t) = lam (go2 (cons k , trans (sym (cons-∘i _ _)) (cong cons i∘k≡h)) t) 43 | go2 (k , i∘k≡h) (mvar u j) = mvar (k ∘i j , trans assoc-∘i (cong (λ h → h ∘i j) i∘k≡h)) 44 | 45 | go2s : ∀ {Sg G X Y Z T} {i : Inj Y Z}{h : Inj X Z} → (∃ \ k → i ∘i k ≡ h) → 46 | (t : Tms Sg G X T) → AllMV∈ i (renT h t) 47 | go2s le [] = [] 48 | go2s le (t ∷ ts) = go2 le t ∷ go2s le ts 49 | 50 | go : ∀ {Sg G1 G3 D1 D2 T} (f : Sub Sg G1 G3) {i : Inj D1 D2} → 51 | {t : Term Sg G1 D2 T} → AllMV∈ i t → f / t ∈ i 52 | go f [] = [] 53 | go f (m ∷ ms) = go f m ∷ go f ms 54 | go f (con ms) = con (go f ms) 55 | go f (var ms) = var (go f ms) 56 | go f (lam m) = lam (go f m) 57 | go f {i = i} (mvar {v = v} {h = h} (k , i∘k≡h)) = go2 (k , i∘k≡h) (f _ v) 58 | 59 | 60 | -- A few properties of substitutions respected by _/_∈_ 61 | _/_∈_-∘ : ∀ {Sg G1 G2 G3 D1 D2 T} {f : Sub Sg G2 G3} (g : Sub Sg G1 G2) {i : Inj D1 D2} → 62 | ∀ (t : Term Sg G1 D2 T) → f / subT g t ∈ i → (f ∘s g) / t ∈ i 63 | _/_∈_-∘ g t m = subst (AllMV∈ _) (subT-∘ _) m 64 | 65 | _/_∈_-ext : ∀ {Sg G G1 D1 D2 T} {i : Inj D1 D2} {f g : Sub Sg G G1} → 66 | f ≡s g → ∀ {t : Term Sg G D2 T} → f / t ∈ i → g / t ∈ i 67 | _/_∈_-ext f≡g {t} m = subst (AllMV∈ _) (subT-ext f≡g t) m 68 | 69 | -- In the flexible-rigid case we'll need to find z and ρ such that ren i z ≡ sub ρ t, 70 | -- this module is about finding such a ρ, which we call the pruner. 71 | -- Its role is to handle occurrences in t like (mvar u j) where there are variables in j 72 | -- which are not in i: ρ will substitute u with a term that ignores them, 73 | -- since their presence would make finding z impossible. 74 | record Pruner {Sg G D1 D2 T} (i : Inj D1 D2) (t : Term Sg G D2 T) : Set where 75 | constructor Pr_,_,_ 76 | field 77 | {G1} : MCtx 78 | pruner : MetaRen G G1 79 | epic : MRop.Monic pruner 80 | prunes : toSub pruner / t ∈ i 81 | 82 | record PrunerSub {Sg G D1 D2 T} (i : Inj D1 D2) (t : Term Sg G D2 T) : Set where 83 | constructor Pr_,_,_ 84 | field 85 | {G1} : MCtx 86 | pruner : Sub Sg G G1 87 | decr : Ctx-length G ≥ Ctx-length G1 88 | prunes : pruner / t ∈ i 89 | 90 | open Pruner using (pruner; prunes; epic) 91 | 92 | _∙_ : ∀ {Sg G D1 D2 T} → {i : Inj D1 D2} {t : Term Sg G D2 T} → 93 | ∀ {D1 D2 T} → {j : Inj D1 D2} {s : Term Sg G D2 T} → 94 | (∀ {G1}{σ : Sub Sg G G1} → σ / t ∈ i → σ / s ∈ j) → Pruner i t → Pruner j s 95 | f ∙ (Pr ρ , ρ-decr , m) = Pr ρ , ρ-decr , f m 96 | 97 | mutual 98 | prune' : ∀ {Sg G D1 D2 T} {i : Inj D1 D2} (t : Tm Sg G D2 T) 99 | → Pruner i t 100 | -- congruence cases 101 | prune' (con c ts) = con ∙ prune's ts 102 | prune' (var x ts) = var ∙ prune's ts 103 | prune' (lam t) = lam ∙ prune' t 104 | 105 | prune' {i = i} (mvar u j) = Pr singleton u p₂ , singleton-epic u p₂ , mvar aux where 106 | open Pullback (pullback i j) 107 | -- (toSub (singleton u p₂)) (mvar u j) = mvar zero (j ∘i p₂), so we 108 | -- need aux to show that (j ∘i p₂) only contains variables in i 109 | aux : ∃ \ k → i ∘i k ≡ j ∘i ρ-env (singleton u p₂ _ u) 110 | aux rewrite thick-refl u = p₁ , commutes 111 | 112 | prune's : ∀ {Sg G D1 D2 T} {i : Inj D1 D2} (t : Tms Sg G D2 T) 113 | → Pruner i t 114 | prune's [] = Pr idmr , id-epic , [] 115 | prune's {i = i} (t ∷ ts) = Pr p₁ ∘mr pruner pr-t , p₁∘pruner-epic , (prunes[t] ∷ prunes[ts]) 116 | where 117 | pr-t = prune' {i = i} t 118 | pr-ts = prune's {i = i} ts 119 | push = (pushout (pruner pr-t) (pruner pr-ts)) 120 | open Mixed.Pushout push 121 | prunes[t] = _/_∈_-∘Closed (toSub p₁) {toSub (pruner pr-t)} {i} {t} (prunes pr-t) 122 | prunes[ts] = _/_∈_-ext {f = toSub (p₂ ∘mr pruner pr-ts)} (λ S u → ↓↓-inj (sym ([]eq commutes S u))) 123 | (_/_∈_-∘Closed (toSub p₂) {_} {i} {ts} (prunes pr-ts)) 124 | open MRopProps using (_∘mono_) 125 | open ESubopProps using (mono-pullback-stable) 126 | 127 | p₁∘pruner-epic = epic pr-t ∘mono 128 | epic-from-sub p₁ (mono-pullback-stable _ _ (Mixed.Pushout.to-sub push) 129 | (epic-to-sub (pruner pr-ts) (epic pr-ts))) 130 | 131 | 132 | 133 | -- prune-sup makes use of the universal property of pullbacks to prove 134 | -- that the pruner computed above is more general than any possible 135 | -- solution to the equation runT i z ≡ sub s t from which we started. 136 | mutual 137 | prune-sup : ∀ {Sg G D1 D2 T} (i : Inj D1 D2) (t : Tm Sg G D2 T) → 138 | ∀ {G1} (s : Sub< false > Sg G G1) z → ren i z ≡T sub s t → s ≤ toSub (pruner (prune' {i = i} t)) 139 | prune-sup i (con c ts) s (con c₁ ts₁) (con _ eq) = prune-sups i ts s ts₁ eq 140 | prune-sup i (var x ts) s (var x₁ ts₁) (var eqv eq) = prune-sups i ts s ts₁ eq 141 | prune-sup i (lam t) s (lam z) (lam eq) = prune-sup (cons i) t s z eq 142 | 143 | prune-sup i (con c ts) s (mvar u j) () 144 | prune-sup i (con c ts) s (var x ts₁) () 145 | prune-sup i (var x ts) s (con c ts₁) () 146 | prune-sup i (var x ts) s (mvar u j) () 147 | 148 | prune-sup {Sg} {G} i (mvar {Ss = Ss} {B} u j) {G2} s z eq = δ , s≡δ∘pruner 149 | where 150 | pull = pullback i j 151 | open Pullback pull 152 | open Σ (forget (lift-pullback pull z (s (B <<- Ss) u) eq)) renaming 153 | (proj₁ to x; proj₂ to ren[p₂,x]≡s[u]) 154 | δ : (S : MTy) → B <<- P ∷ G - u ∋ S → Tm< false > Sg G2 (ctx S) ([] ->> type S) 155 | δ .(B <<- P) zero = x 156 | δ S (suc v) = s S (thin u S v) 157 | s≡δ∘pruner : (S : MTy) (v : G ∋ S) → s S v ≡ sub δ (toSub (singleton u p₂) S v) 158 | s≡δ∘pruner S v with thick u v 159 | s≡δ∘pruner S .(thin u S v) | inj₁ (v , refl) = sym (ren-id _) 160 | s≡δ∘pruner .(B <<- Ss) .u | inj₂ refl` = sym ren[p₂,x]≡s[u] 161 | 162 | 163 | prune-sups : ∀ {Sg G D1 D2 T} (i : Inj D1 D2) (t : Tms Sg G D2 T) → 164 | ∀ {G1} (s : Sub< false > Sg G G1) z → rens i z ≡T subs s t → s ≤ toSub (pruner (prune's {i = i} t)) 165 | prune-sups i [] s [] eq = s , (λ S u → sym (ren-id _)) 166 | prune-sups {Sg} {G} i (t ∷ ts) s (z ∷ zs) (eqt ∷ eqts) = down uni , (λ S u → 167 | begin 168 | s S u ≡⟨ proj₂ s≤pr-t S u ⟩ 169 | (proj₁ s≤pr-t ∘s toSub (pruner pr-t)) S u ≡⟨ sub-ext (λ S₁ u₁ → sym (uni∘p₁≡q₁ S₁ u₁)) (toSub (pruner pr-t) S u) ⟩ 170 | ((down uni ∘s toSub p₁) ∘s toSub (pruner pr-t)) S u ≡⟨ sym (Sub∘.subT-∘ {f = down uni} {g = toSub p₁} 171 | (toSub (pruner pr-t) S u)) ⟩ 172 | (down uni ∘s (toSub p₁ ∘s toSub (pruner pr-t))) S u ∎) 173 | where 174 | pr-t = prune' {i = i} t 175 | pr-ts = prune's {i = i} ts 176 | push = pushout (pruner pr-t) (pruner pr-ts) 177 | open Mixed.Pushout {Sg} push 178 | s≤pr-t : s ≤ toSub (pruner (prune' {i = i} t)) 179 | s≤pr-t = prune-sup i t s z eqt 180 | s≤pr-ts : s ≤ toSub (pruner (prune's {i = i} ts)) 181 | s≤pr-ts = prune-sups i ts s zs eqts 182 | eq = ES (λ S u → cong ↓↓ (trans (sym (proj₂ s≤pr-t S u)) (proj₂ s≤pr-ts S u))) 183 | uni = (ESub.⟦ universal (ι (proj₁ s≤pr-t)) (ι (proj₁ s≤pr-ts)) eq ⟧) 184 | uni∘p₁≡q₁ : (down uni ∘s toSub p₁) ≡s proj₁ s≤pr-t 185 | uni∘p₁≡q₁ = \ S u -> trans (≅-to-≡ (↓↓-comm uni (toSub p₁ S u))) 186 | ([]eq (p₁∘universal≡q₁ {q₁ = ι (proj₁ s≤pr-t)} {q₂ = ι (proj₁ s≤pr-ts)} {eq}) S u) 187 | 188 | prune : ∀ {Sg G D1 D2 T} (i : Inj D1 D2) (t : Tm Sg G D2 T) → 189 | Σ (PrunerSub i t) λ pr → ∀ {G1} (s : Sub< false > Sg G G1) z → ren i z ≡T sub s t → s ≤ PrunerSub.pruner pr 190 | prune i t = (Pr toSub (pruner pr) , epi-decr (pruner pr , epic pr) , prunes pr) , prune-sup i t 191 | where 192 | pr = prune' {i = i} t 193 | 194 | -------------------------------------------------------------------------------- /Unification/Pruning/Epi-Decr.agda: -------------------------------------------------------------------------------- 1 | module Unification.Pruning.Epi-Decr where 2 | 3 | open import Data.Nat using (_≥_; zero; suc; z≤n; s≤s; module ≤-Reasoning) 4 | open import Data.Nat.Properties 5 | open import Data.Empty 6 | open import Data.Sum 7 | 8 | open import Support.Equality 9 | 10 | open import Injections 11 | open import MetaRens 12 | 13 | open import Syntax 14 | 15 | open import Unification.Specification.Decr-Sub 16 | open import Unification.Injections 17 | 18 | Epi : _ -> _ -> Set 19 | Epi G G1 = ∃ \ (f : MetaRen G G1) -> MRop.Monic f 20 | 21 | Inj-incr : ∀ {A : Set}{X : List A}{Y} -> Inj X Y -> length Y ≥ length X 22 | Inj-incr [] = z≤n 23 | Inj-incr {A} {x ∷ xs} {Y} (u ∷ i [ pf ]) rewrite length-del u = s≤s (Inj-incr j) 24 | where 25 | jf : (x₁ : A) → xs ∋ x₁ → Y - u ∋ x₁ 26 | jf _ v with thick u (i $ v) 27 | ... | inj₁ (w , _) = w 28 | jf ._ v | inj₂ (refl , eq) = ⊥-elim (∉-∉Im i u pf v (≅-to-≡ eq)) 29 | 30 | jf-inj : (x : A) {a b : xs ∋ x} -> jf x a ≡ jf x b -> a ≡ b 31 | jf-inj x {a} {b} eq 32 | with thick u (i $ a) | thick u (i $ b) 33 | jf-inj x₁ refl | inj₁ (._ , eq1) | inj₁ (_ , eq2) = injective i _ _ (trans (sym eq1) eq2) 34 | jf-inj ._ _ | inj₁ x₂ | inj₂ (refl , eq) = ⊥-elim (∉-∉Im i u pf _ (≅-to-≡ eq)) 35 | jf-inj ._ _ | inj₂ (refl , eq) | p = ⊥-elim (∉-∉Im i u pf _ (≅-to-≡ eq)) 36 | 37 | j : Inj xs (Y - u) 38 | j = quo jf {jf-inj} 39 | 40 | 41 | punchout : ∀ {x G G1} -> (f : Epi (x ∷ G) G1) -> Epi G G1 ⊎ Epi G (G1 - body (proj₁ f _ zero)) 42 | punchout {x} {G} {G1} (f , f-epic) 43 | with any? (λ v → body (f _ (suc v)) ≅∋? body (f _ zero)) 44 | ... | yes (_ , v , eqty , eqbody) = inj₁ ((λ S x → f S (suc x)) , fsuc-epic) 45 | where 46 | fsuc-epic : MRop.Monic (λ S x → f S (suc x)) 47 | fsuc-epic {C} {g1} {g2} eq = f-epic {C} {g1} {g2} (∋-case (cong (map-Vc (ρ-env (f _ zero))) aux) eq) 48 | where 49 | aux : g1 _ (body (f _ zero)) ≡ g2 _ (body (f _ zero)) 50 | aux with type x <<- Ψ (f _ zero) | body (f _ zero) | eqty | eqbody 51 | aux | ._ | ._ | refl | refl = map-Vc-inj (ρ-env (f _ (suc v))) (to-vc (eq _ v)) 52 | ... | no ¬p = inj₂ (f' , f'-epic) where 53 | u = body (f _ zero) 54 | 55 | f' : MetaRen G (G1 - u) 56 | f' S x with f S (suc x) | inspect (f S) (suc x) 57 | f' S x | i / v | ⌞ eq ⌟ with thick u v 58 | ... | inj₁ x₂ = i / proj₁ x₂ 59 | f' (._ <<- _) x | i / ._ | ⌞ eq ⌟ | inj₂ refl` = ⊥-elim (¬p (_ , _ , cong (_<<-_ _) (_≈vc_.Ψeq (to-vc eq)) , _≈vc_.beq (to-vc eq))) 60 | 61 | f'-epic : MRop.Monic f' 62 | f'-epic {C} {g1} {g2} eq S y = wk-inj 63 | (begin wk (g1 S y) ≡⟨ sym (shift-thin g1 S y) ⟩ 64 | shift g1 S (thin u S y) ≡⟨ f-epic {_} {shift g1} {shift g2} (∋-case zero-eq suc-eq) S (thin u S y) ⟩ 65 | shift g2 S (thin u S y) ≡⟨ shift-thin g2 S y ⟩ 66 | wk (g2 S y) ∎ ) 67 | where 68 | open ≡-Reasoning 69 | shift : MetaRen (G1 - u) C -> MetaRen G1 (_ ∷ C) 70 | shift g1 S x with thick u x 71 | shift g1 S₁ x₁ | inj₁ x₂ = wk (g1 _ (proj₁ x₂)) 72 | shift g1 ._ ._ | inj₂ refl` = id-i / zero 73 | 74 | shift-thin : ∀ g S y -> shift g _ (thin u S y) ≡ wk (g S y) 75 | shift-thin g S y rewrite thick-thin u y = refl 76 | 77 | shift-refl : ∀ g -> shift g _ u ≡ id-i / zero 78 | shift-refl g rewrite thick-refl u = refl 79 | 80 | shift-refl2 : ∀ {S}{v : _ ∋ S} g g1 -> u ≅∋ v -> shift g _ v ≡ shift g1 _ v 81 | shift-refl2 {.(type x) <<- .(Ψ (f x zero))} g g1 refl` = trans (shift-refl g) (sym (shift-refl g1)) 82 | 83 | zero-eq = cong (map-Vc _) (begin shift g1 _ u ≡⟨ shift-refl g1 ⟩ 84 | id-i / zero ≡⟨ sym (shift-refl g2) ⟩ 85 | shift g2 _ u ∎) 86 | 87 | suc-eq : (shift g1 ∘mr (\ S y -> f S (suc y))) ≡mr (shift g2 ∘mr (\ S y -> f S (suc y))) 88 | suc-eq S x with thick u (body (f S (suc x))) | eq S x | shift-refl2 {v = (body (f S (suc x)))} g1 g2 89 | suc-eq S₁ x₁ | inj₁ x₂ | p | _ = cong wk p 90 | suc-eq S₁ x₁ | inj₂ y₁ | p | q = cong (map-Vc _) (q y₁) 91 | 92 | 93 | epi-decr : ∀ {G G1} -> Epi G G1 -> Ctx-length G ≥ Ctx-length G1 94 | epi-decr {[]} {[]} (f , f-epic) = z≤n 95 | epi-decr {[]} {x ∷ G1} (f , f-epic) 96 | with f-epic {x ∷ x ∷ G1} {∋-case (id-i / zero) (\ _ x -> id-i / suc (suc x))} 97 | {∋-case (id-i / suc zero) (\ _ x -> id-i / suc (suc x))} 98 | (\ _ ()) _ zero 99 | ... | () 100 | epi-decr {x ∷ G} (f , f-epic) with f _ zero | Data.Sum.map epi-decr epi-decr (punchout (f , f-epic)) 101 | epi-decr {x ∷ G} {G1} (f , f-epic) | i / u | inj₁ G≥G1 = begin 102 | Ctx-length G1 ≤⟨ G≥G1 ⟩ 103 | Ctx-length G ≤⟨ n≤m+n (suc (length (ctx x))) (Ctx-length G) ⟩ 104 | Ctx-length (x ∷ G) ∎ where open ≤-Reasoning 105 | epi-decr {x ∷ G} {G1} (f , f-epic) | i / u | inj₂ G≥G1-u rewrite Ctx-length-lemma u = s≤s (Inj-incr i) +-mono G≥G1-u 106 | 107 | -------------------------------------------------------------------------------- /Unification/Specification.agda: -------------------------------------------------------------------------------- 1 | module Unification.Specification where 2 | 3 | open import Data.Sum renaming (map to map⊎) 4 | open import Data.Sum renaming (inj₁ to yes; inj₂ to no) 5 | 6 | open import Support.Equality 7 | open ≡-Reasoning 8 | 9 | open import Injections 10 | 11 | open import Syntax 12 | 13 | open import Unification.Specification.Decr-Sub public 14 | open import Unification.Injections 15 | 16 | Property<_> : ∀ b Sg G -> Set₁ 17 | Property< b > Sg G = (∀ {G2} -> Sub< b > Sg G G2 -> Set) 18 | 19 | Property : ∀ Sg G -> Set₁ 20 | Property Sg G = ∀ {b} -> Property< b > Sg G 21 | 22 | Unifies : ∀ {Sg G1 D S} (x y : Term Sg G1 D S) -> Property Sg G1 23 | Unifies x y σ = subT σ x ≡T subT σ y 24 | 25 | ∃⟦σ⟧_ : ∀ {Sg G1} -> Property< true > Sg G1 -> Set 26 | ∃⟦σ⟧ P = ∃ \ G2 -> ∃ \ σ -> P {G2} ⟦ σ ⟧ 27 | 28 | ∃σ_ : ∀ {Sg G1} -> Property< false > Sg G1 -> Set 29 | ∃σ_ P = ∃ \ G2 -> ∃ \ σ -> P {G2} σ 30 | 31 | ∃σ-pat_ : ∀ {Sg G1} -> Property< true > Sg G1 -> Set 32 | ∃σ-pat_ P = ∃ \ G2 -> ∃ \ σ -> P {G2} σ 33 | 34 | Sup : ∀ {Sg G1} -> Property< false > Sg G1 -> Property< true > Sg G1 35 | Sup P σ = {G' : _} (ρ : _) → P {G'} ρ → ρ ≤ σ 36 | 37 | Max : ∀ {Sg G1} -> Property Sg G1 -> Property< true > Sg G1 38 | Max P σ = P σ × Sup P σ 39 | 40 | Extensional : ∀ {b Sg G} -> Property< b > Sg G -> Set 41 | Extensional P = ∀ {G f g} -> f ≡s g -> P {G} f -> P g 42 | 43 | Spec : ∀ {Sg G1 D S} (x y : Term Sg G1 D S) -> Set 44 | Spec x y = ∃⟦σ⟧ Max (Unifies x y) ⊎ ¬ ∃σ Unifies x y 45 | 46 | Unifies-ext : ∀ {Sg G1 D S} (x y : Term Sg G1 D S) -> ∀ {b} -> Extensional {b} (Unifies x y) 47 | Unifies-ext x y f≡g subfx≡subfy rewrite subT-ext f≡g x | subT-ext f≡g y = subfx≡subfy 48 | 49 | ≤-∘ : ∀ {Sg g g1 g2 g3}(ρ : Sub< false > Sg g g1)(p : Sub Sg g g2)(q : Sub Sg g2 g3) -> (ρ≤p : ρ ≤ p) -> proj₁ ρ≤p ≤ q -> ρ ≤ (q ∘s p) 50 | ≤-∘ ρ p q (δ , ρ≡δ∘p) (δ' , δ'≡δ∘q) 51 | = δ' , λ S u → begin ρ S u ≡⟨ ρ≡δ∘p S u ⟩ 52 | sub δ (p S u) ≡⟨ sub-ext δ'≡δ∘q (p S u) ⟩ 53 | subT (δ' ∘s q) (p S u) ≡⟨ sym (Sub∘.subT-∘ (p S u)) ⟩ 54 | subT δ' (subT q (p S u)) ∎ 55 | 56 | 57 | map-Unifies : ∀ {Sg g h h' d t} {σ : Sub Sg g h}{σ' : Sub Sg g h'}-> σ ≤ σ' -> {x y : Term Sg g d t} -> Unifies x y σ' -> Unifies x y σ 58 | map-Unifies {σ = σ} {σ'} (δ , σ≡δ∘σ') {x} {y} σ'Unifies[x,y] = ≡-T 59 | (begin subT σ x ≡⟨ subT-ext σ≡δ∘σ' x ⟩ 60 | subT (δ ∘s σ') x ≡⟨ T-≡ (sandwich subT-∘ (T.cong (subT δ) σ'Unifies[x,y])) ⟩ 61 | subT (δ ∘s σ') y ≡⟨ sym (subT-ext σ≡δ∘σ' y) ⟩ 62 | subT σ y ∎) 63 | 64 | 65 | shift_under_by_ : ∀ {Sg G h1 h2 D T} {xs ys : Term Sg G D T} {σ1 : Sub< false > Sg G h1} 66 | -> Unifies xs ys σ1 -> (σ : Sub Sg G h2) -> σ1 ≤ σ -> ∃σ Unifies (subT σ xs) (subT σ ys) 67 | shift_under_by_ eq σ (δ , σ1≡δ∘σ) = _ , δ , sandwich (λ xs₁ → sym (trans (Sub∘.subT-∘ xs₁) 68 | (sym (subT-ext σ1≡δ∘σ xs₁)))) 69 | eq 70 | 71 | 72 | unify-comm : ∀ {Sg G D T} → (x y : Term Sg G D T) → ∃σ Unifies x y → ∃σ Unifies y x 73 | unify-comm _ _ (G , σ , eq) = (G , σ , T.sym eq) 74 | 75 | spec-comm : ∀ {Sg G D T} → (x y : Term Sg G D T) → Spec x y → Spec y x 76 | spec-comm _ _ = map⊎ (λ {(G , σ , eq , max) → G , σ , T.sym eq , (λ {_} ρ x → max ρ (T.sym x))}) (λ x x₁ → x (unify-comm _ _ x₁)) 77 | 78 | 79 | cong-spec : ∀ {Sg G D D' T T'} → (d : DTm Sg G (D' , T') (D , T)) -> {x y : Term Sg G D T} → Spec x y → Spec (wrapD d x) (wrapD d y) 80 | cong-spec d (inj₁ (_ , σ , unifies , sup)) = inj₁ (_ , (σ , (cong-wrapD d unifies , (λ ρ ρ-unifies → sup ρ (inv-wrapD d ρ-unifies))))) 81 | cong-spec d (inj₂ no-unifier) = inj₂ (λ {(_ , σ , σ-unifies) → no-unifier (_ , (σ , inv-wrapD d σ-unifies)) }) 82 | 83 | optimist : ∀ {Sg m l o D T Ts}(x y : Tm Sg m D T)(xs ys : Tms Sg m D Ts) -> 84 | (p : Sub Sg m o) (q : Sub Sg o l) -> 85 | Max (Unifies x y) p 86 | -> Max (Unifies (subT p xs) (subT p ys)) q 87 | -> Max (Unifies (Tms._∷_ x xs) (y ∷ ys)) (q ∘s p) 88 | 89 | optimist x y xs ys p q ([p]Unifies[x,y] , sup-p) ([q]Unifies[px,py] , sup-q) = 90 | (map-Unifies (q , λ S u → refl) {x} {y} [p]Unifies[x,y] ∷ sandwich subT-∘ [q]Unifies[px,py]) , 91 | (λ {ρ ([ρ]Unifies[x,y] ∷ [ρ]Unifies[xs,ys]) 92 | → let ρ≤p : _ 93 | ρ≤p = sup-p ρ [ρ]Unifies[x,y] 94 | δ : _ 95 | δ = proj₁ ρ≤p 96 | ρ≡δ∘p : _ 97 | ρ≡δ∘p = proj₂ ρ≤p 98 | δ≤q : _ 99 | δ≤q 100 | = sup-q δ 101 | (sandwich (λ x₁ → sym (Sub∘.subT-∘ x₁)) 102 | (Unifies-ext xs ys ρ≡δ∘p [ρ]Unifies[xs,ys])) 103 | in ≤-∘ ρ p q ρ≤p δ≤q}) 104 | 105 | 106 | 107 | ∃σMax[Unifies[x,x]] : ∀ {Sg G D T} (x : Term Sg G D T) -> ∃⟦σ⟧ Max (Unifies x x) 108 | ∃σMax[Unifies[x,x]] x = _ , 109 | (DS id-s , inj₁ (refl , IsIso-id)) , 110 | refl-T _ , 111 | (λ ρ x₁ → ρ , (λ S u → sym (ren-id _))) 112 | 113 | Spec[xs,ys]⇒Spec[σxs,σys] : ∀ {Sg G G1 D T} {xs ys : Term Sg G D T} (σ : Sub Sg G G1) -> Ctx-length G ≡ Ctx-length G1 -> 114 | IsIso σ -> Spec xs ys -> Spec (subT σ xs) (subT σ ys) 115 | Spec[xs,ys]⇒Spec[σxs,σys] {xs = xs} {ys = ys} σ G~G1 ((δ , id≡δ∘σ) , id≡σ∘δ) (yes (_ , σ₁ , [σ₁]Unifies[xs,ys] , sup-σ₁)) 116 | = yes (_ , σ₁ ∘ds (DS δ , inj₁ (sym G~G1 , (σ , id≡σ∘δ) , id≡δ∘σ)) , [σ₁∘δ]Unifies[σxs,σys] , sup-[σ₁∘δ]) 117 | where 118 | [σ₁∘δ]Unifies[σxs,σys] = sandwich (λ ys → 119 | begin subT ⟦ σ₁ ⟧ ys ≡⟨ cong (subT ⟦ σ₁ ⟧) (trans (sym (subT-id ys)) (subT-ext id≡δ∘σ ys)) ⟩ 120 | subT ⟦ σ₁ ⟧ (subT (δ ∘s σ) ys) ≡⟨ subT-∘ ys ⟩ 121 | subT (⟦ σ₁ ⟧ ∘s (δ ∘s σ)) ys ≡⟨ subT-ext (λ S x → subT-∘ (σ S x)) ys ⟩ 122 | subT ((⟦ σ₁ ⟧ ∘s δ) ∘s σ) ys ≡⟨ sym (subT-∘ ys) ⟩ 123 | subT (⟦ σ₁ ⟧ ∘s δ) (subT σ ys) ∎) 124 | [σ₁]Unifies[xs,ys] 125 | 126 | sup-[σ₁∘δ] : Sup (Unifies (subT σ xs) (subT σ ys)) (⟦ σ₁ ⟧ ∘s δ) 127 | sup-[σ₁∘δ] ρ [ρ]Unifies[σxs,σys] = δ' , λ S u → 128 | begin 129 | ρ S u ≡⟨ sym (left-ids ρ S u) ⟩ 130 | sub ρ (id-s S u) ≡⟨ cong (sub ρ) (id≡σ∘δ S u) ⟩ 131 | sub ρ (sub σ (δ S u)) ≡⟨ Sub∘.subT-∘ {f = ρ} {σ} (δ S u) ⟩ 132 | subT (ρ ∘s σ) (δ S u) ≡⟨ subT-ext (proj₂ ρ∘σ≤σ₁) (δ S u) ⟩ 133 | subT (δ' ∘s ⟦ σ₁ ⟧) (δ S u) ≡⟨ sym (Sub∘.subT-∘ (δ S u)) ⟩ 134 | subT δ' (subT ⟦ σ₁ ⟧ (δ S u)) ∎ 135 | where 136 | ρ∘σ≤σ₁ = sup-σ₁ (ρ ∘s σ) (sandwich Sub∘.subT-∘ [ρ]Unifies[σxs,σys]) 137 | δ' = proj₁ ρ∘σ≤σ₁ 138 | 139 | Spec[xs,ys]⇒Spec[σxs,σys] σ G~G1 _ (no ¬p) = no (λ {(_ , σ₁ , eq) → ¬p (_ , σ₁ ∘s σ , sandwich Sub∘.subT-∘ eq)}) 140 | -------------------------------------------------------------------------------- /Unification/Specification/Decr-Sub.agda: -------------------------------------------------------------------------------- 1 | module Unification.Specification.Decr-Sub where 2 | 3 | open import Data.Nat renaming (_≤_ to _≤ℕ_) 4 | open import Relation.Nullary 5 | open import Data.Nat.Properties 6 | open import Algebra 7 | open CommutativeSemiring commutativeSemiring using (+-comm; +-assoc) 8 | open import Data.Sum 9 | 10 | open import Support.Equality 11 | 12 | open import Injections 13 | open import MetaRens 14 | 15 | open import Syntax 16 | 17 | open import Unification.Injections 18 | 19 | -- We define a measure of meta-contexts to help with proving 20 | -- termination of the main unification algorithm and pruning. 21 | Ctx-length : MCtx -> ℕ 22 | Ctx-length [] = zero 23 | Ctx-length (type <<- ctx ∷ m) = suc (length ctx + Ctx-length m) 24 | 25 | IsIso : ∀ {Sg G1 G2} -> (s : Sub Sg G1 G2) -> Set 26 | IsIso s = Σ (∃ \ r -> id-s ≡s (r ∘s s)) \le -> id-s ≡s (s ∘s proj₁ le) 27 | 28 | -- The substitutions we produce are going to either be isomorphims or 29 | -- produce terms in a smaller context, so it'll be fine to recurse on 30 | -- their results. 31 | Decreasing : ∀ {Sg G1 G2} -> (s : Sub Sg G1 G2) -> Set 32 | Decreasing {Sg} {G1} {G2} s = (Ctx-length G1 ≡ Ctx-length G2 × IsIso s) 33 | ⊎ (Ctx-length G1 > Ctx-length G2) 34 | 35 | record DSub (Sg : Ctx) (G1 : MCtx) (G2 : MCtx) : Set where 36 | constructor DS_,_ 37 | field 38 | ⟦_⟧ : Sub Sg G1 G2 39 | decreasing : Decreasing ⟦_⟧ 40 | 41 | open DSub public 42 | 43 | -- Below we prove that DSub forms a category and that other 44 | -- substitutions of interest are Decreasing. 45 | 46 | Ctx-length-lemma : ∀ {G Ss B} -> (u : G ∋ B <<- Ss) -> Ctx-length G ≡ Ctx-length (G - u <: B <<- Ss) 47 | Ctx-length-lemma zero = refl 48 | Ctx-length-lemma {._ ∷ G} {Ss} (suc {S = _ <<- ctx} u) = 49 | begin 50 | suc (length ctx) + Ctx-length G ≡⟨ cong (_+_ (suc (length ctx))) (Ctx-length-lemma u) ⟩ 51 | suc (length ctx) + (suc (length Ss) + Ctx-length (G - u)) ≡⟨ sym (+-assoc (suc (length ctx)) (suc (length Ss)) _) ⟩ 52 | suc (length ctx) + suc (length Ss) + Ctx-length (G - u) ≡⟨ cong (λ x → x + Ctx-length (G - u)) (+-comm (suc (length ctx)) (suc (length Ss))) ⟩ 53 | suc (length Ss) + suc (length ctx) + Ctx-length (G - u) ≡⟨ +-assoc (suc (length Ss)) (suc (length ctx)) _ ⟩ 54 | suc (length Ss) + (suc (length ctx) + Ctx-length (G - u)) ∎ 55 | where open ≡-Reasoning 56 | 57 | IsIso-id : ∀ {Sg G} -> IsIso {Sg} {G} {G} id-s 58 | IsIso-id = λ {Sg} {G} → (id-s , (λ S u → sym (ren-id _))) , (λ S u → sym (ren-id _)) 59 | 60 | IsIso-∘ : ∀ {Sg G1 G2 G3} -> (s : Sub Sg G2 G3) -> (s' : Sub Sg G1 G2) -> IsIso s -> IsIso s' -> IsIso (s ∘s s') 61 | IsIso-∘ s s' ((δ , p) , p') ((δ' , q) , q') = (δ' ∘s δ , 62 | (λ S u → 63 | begin 64 | mvar u id-i ≡⟨ q S u ⟩ 65 | sub δ' (s' S u) ≡⟨ cong (sub δ') (sym (sub-id (s' S u))) ⟩ 66 | sub δ' (sub id-s (s' S u)) ≡⟨ cong (sub δ') (sub-ext p (s' S u)) ⟩ 67 | sub δ' (sub (δ ∘s s) (s' S u)) ≡⟨ cong (sub δ') (sym (sub-∘ (s' S u))) ⟩ 68 | sub δ' (sub δ (sub s (s' S u))) ≡⟨ sub-∘ {f = δ'} {g = δ} (sub s (s' S u)) ⟩ 69 | sub (δ' ∘s δ) (sub s (s' S u)) ∎)), 70 | 71 | (λ S u → 72 | begin 73 | mvar u id-i ≡⟨ p' S u ⟩ 74 | sub s (δ S u) ≡⟨ cong (sub s) (sym (sub-id (δ S u))) ⟩ 75 | sub s (sub id-s (δ S u)) ≡⟨ cong (sub s) (sub-ext q' (δ S u)) ⟩ 76 | sub s (sub (s' ∘s δ') (δ S u)) ≡⟨ cong (sub s) (sym (sub-∘ {f = s'} {g = δ'} (δ S u))) ⟩ 77 | sub s (sub s' (sub δ' (δ S u))) ≡⟨ sub-∘ (sub δ' (δ S u)) ⟩ 78 | sub (s ∘s s') (sub δ' (δ S u)) ∎) 79 | 80 | where open ≡-Reasoning 81 | 82 | trans-> : ∀ {m n o} -> m > n -> n > o -> m > o 83 | trans-> (s≤s m≤n) (s≤s z≤n) = s≤s z≤n 84 | trans-> (s≤s m≤n) (s≤s (s≤s m≤n₁)) = s≤s (trans-> m≤n (s≤s m≤n₁)) 85 | 86 | open ≤-Reasoning 87 | 88 | trans-dec : ∀ {Sg G1 G2 G3} -> (s : Sub Sg G2 G3) -> Decreasing s -> (s' : Sub Sg G1 G2) -> Decreasing s' -> Decreasing (s ∘s s') 89 | trans-dec s (inj₁ (G2~G3 , s-is-iso)) s' (inj₁ (G1~G2 , s'-is-iso )) 90 | = inj₁ (trans G1~G2 G2~G3 , IsIso-∘ s s' s-is-iso s'-is-iso ) 91 | trans-dec {Sg} {G1} {G2} {G3} s (inj₁ (G2~G3 , _)) s' (inj₂ G1>G2) 92 | = inj₂ 93 | (begin 94 | suc (Ctx-length G3) ≡⟨ sym (cong suc G2~G3) ⟩ 95 | suc (Ctx-length G2) ≤⟨ G1>G2 ⟩ 96 | Ctx-length G1 ∎) 97 | trans-dec {Sg} {G1} {G2} {G3} s (inj₂ G2>G3) s' (inj₁ (G1~G2 , _)) 98 | = inj₂ 99 | (begin 100 | suc (Ctx-length G3) ≤⟨ G2>G3 ⟩ 101 | Ctx-length G2 ≡⟨ sym G1~G2 ⟩ 102 | Ctx-length G1 ∎) 103 | trans-dec s (inj₂ y) s' (inj₂ y₁) = inj₂ (trans-> y₁ y) 104 | 105 | _∘ds_ : ∀ {Sg G1 G2 G3} -> DSub Sg G2 G3 -> DSub Sg G1 G2 -> DSub Sg G1 G3 106 | (DS σ , G2>G3) ∘ds (DS σ₁ , G1>G2) = DS (σ ∘s σ₁) , trans-dec σ G2>G3 σ₁ G1>G2 107 | 108 | ⟦⟧-∘ : ∀ {Sg g h i} (s : DSub Sg h i) (s₁ : DSub Sg g h) -> ⟦ s ∘ds s₁ ⟧ ≡s (⟦ s ⟧ ∘s ⟦ s₁ ⟧) 109 | ⟦⟧-∘ s s1 S x = refl 110 | 111 | _≅i_ : ∀ {A : Set}{X Y Z : List A} -> Inj X Y -> Inj X Z -> Set 112 | _≅i_ {A} {X} {Y} {Z} i j = Z ≡ Y × i ≅ j 113 | 114 | Decr-i : ∀ {A : Set}{X Y : List A} -> Inj X Y -> Set 115 | Decr-i {A} {E} {S} e = (e ≅i id-i) ⊎ length S > length E 116 | 117 | cons-id-≅i : ∀ {A : Set} {X Y : List A} {x : A} {e : Inj X Y} -> e ≅i id-i -> cons {x = x} e ≅i id-i 118 | cons-id-≅i refl` = refl , ≡-to-≅ cons-id 119 | 120 | equalizer-Decr : ∀ {A : Set}{S T : List A}(f g : Inj S T) -> let open Equalizer (equalizer f g) in Decr-i e 121 | equalizer-Decr [] [] = inj₁ refl` 122 | equalizer-Decr (i ∷ f [ _ ]) ( j ∷ g [ _ ]) with i ≅∋? j | equalizer-Decr f g 123 | equalizer-Decr (i ∷ f [ _ ]) (.i ∷ g [ _ ]) | yes refl` | inj₁ eq = inj₁ (cons-id-≅i eq) 124 | equalizer-Decr (i ∷ f [ _ ]) (.i ∷ g [ _ ]) | yes refl` | inj₂ gt = inj₂ (s≤s gt) 125 | equalizer-Decr (i ∷ f [ _ ]) ( j ∷ g [ _ ]) | no _ | inj₁ (eq , _) = inj₂ (s≤s (begin _ ≡⟨ cong length eq ⟩ _ ∎)) 126 | equalizer-Decr (i ∷ f [ _ ]) ( j ∷ g [ _ ]) | no _ | inj₂ gt = inj₂ (≤-step gt) 127 | 128 | pullback-Decr : ∀ {A : Set} {X Y Z : List A} → (f : Inj X Z)(g : Inj Y Z) -> let open Pullback (pullback f g) in Decr-i p₂ 129 | pullback-Decr f [] = inj₁ refl` 130 | pullback-Decr f (i ∷ g [ _ ]) with invert f i | pullback-Decr f g 131 | pullback-Decr f (.(f $ x) ∷ g [ _ ]) | yes (x , refl) | inj₁ eq = inj₁ (cons-id-≅i eq) 132 | pullback-Decr f (.(f $ x) ∷ g [ _ ]) | yes (x , refl) | inj₂ gt = inj₂ (s≤s gt) 133 | pullback-Decr f (i ∷ g [ _ ]) | no _ | inj₁ eq = inj₂ (s≤s (begin _ ≡⟨ cong length (proj₁ eq) ⟩ _ ∎)) 134 | pullback-Decr f (i ∷ g [ _ ]) | no _ | inj₂ gt = inj₂ (≤-step gt) 135 | 136 | 137 | singleton-Decreasing : ∀ {Sg G E Ss B} (e : Inj E Ss) (u : G ∋ B <<- Ss) -> Decr-i e -> Decreasing {Sg} (toSub (singleton u e)) 138 | singleton-Decreasing {Sg} {G} {.Ss} {Ss} {B} .id-i u (inj₁ refl`) = inj₁ (Ctx-length-lemma u , (δ , eq1) , (λ S u₁ → eq2 S u₁)) where 139 | 140 | δ : (S : MTy) → B <<- Ss ∷ G - u ∋ S → Tm Sg G (ctx S) ([] ->> type S) 141 | δ .(B <<- Ss) zero = mvar u id-i 142 | δ S (suc u₁) = mvar (thin u S u₁) id-i 143 | 144 | eq1 : id-s ≡s (δ ∘s toSub (singleton u id-i)) 145 | eq1 S u₁ with thick u u₁ 146 | eq1 S .(thin u S x) | inj₁ (x , refl) = cong (mvar _) (sym (right-id id-i)) 147 | eq1 .(B <<- Ss) .u | inj₂ refl` = cong (mvar u) (sym (right-id id-i)) 148 | 149 | eq2 : id-s ≡s (toSub (singleton u id-i) ∘s δ) 150 | eq2 ._ (zero {._} {.(_ <<- _)}) rewrite thick-refl u = cong (mvar _) (sym (right-id id-i)) 151 | eq2 S (suc {._} {._} {.(_ <<- _)} v) rewrite thick-thin u v = cong (mvar _) (sym (right-id id-i)) 152 | 153 | singleton-Decreasing {Sg} {G} {E} {Ss} {B} e u (inj₂ Ss>E) 154 | = inj₂ 155 | (begin 156 | suc (suc (length E) + Ctx-length (G - u)) ≤⟨ s≤s (Ss>E +-mono (begin Ctx-length (G - u) ∎)) ⟩ 157 | Ctx-length (G - u <: B <<- Ss) ≡⟨ sym (Ctx-length-lemma u) ⟩ 158 | Ctx-length G ∎) 159 | 160 | rigid-decr : ∀ {G G1}{x}(u : G ∋ x) -> Ctx-length (G - u) ≥ Ctx-length G1 161 | -> Ctx-length G > Ctx-length G1 162 | rigid-decr {G} {G1} {type <<- ctx} u G-u≤G1 = 163 | begin suc (Ctx-length G1) ≤⟨ s≤s G-u≤G1 ⟩ 164 | suc (Ctx-length (G - u)) ≤⟨ s≤s (n≤m+n (length ctx) (Ctx-length (G - u))) ⟩ 165 | Ctx-length (G - u <: type <<- ctx) ≡⟨ sym (Ctx-length-lemma u) ⟩ 166 | Ctx-length G ∎ 167 | -------------------------------------------------------------------------------- /Vars.agda: -------------------------------------------------------------------------------- 1 | module Vars where 2 | 3 | open import Data.Nat 4 | open import Data.Product 5 | open import Data.Maybe 6 | open import Function using (_∘_) 7 | open import Data.Empty 8 | open import Data.Sum 9 | 10 | open import Support.Equality 11 | open ≡-Reasoning 12 | 13 | open import Relation.Nullary public 14 | open import Data.List public hiding ([_]) 15 | 16 | infix 4 _∋_ 17 | 18 | -- Proofs of membership in lists 19 | -- a.k.a. typed de Bruijn indices 20 | -- i.e. our variable names. 21 | data _∋_ {A : Set} : List A → A → Set where 22 | zero : ∀ {G T} → T ∷ G ∋ T 23 | suc : ∀ {G T S} → G ∋ T → S ∷ G ∋ T 24 | 25 | _∈_ : {A : Set} → A → List A → Set 26 | _∈_ x xs = xs ∋ x 27 | 28 | ∋-case : ∀ {A : Set}{xs : List A}{x} {P : ∀ a -> x ∷ xs ∋ a -> Set} -> P x zero -> (∀ a y -> P a (suc y)) -> ∀ a y -> P a y 29 | ∋-case z s a zero = z 30 | ∋-case z s a (suc v) = s a v 31 | 32 | _≅∋_ : ∀ {A : Set}{G S} {T : A} -> G ∋ S -> G ∋ T -> Set 33 | _≅∋_ {A} {G} {S} {T} u v = S ≡ T × u ≅ v 34 | 35 | -- Given G = [T1 , .. , S , .. , Tn], (u : G ∋ S) we have this isomorphism: 36 | -- 37 | -- G ∋ T 38 | -- ≡ [T1 , .. , S , .. , Tn] ∋ T 39 | -- ≅ T1 ≡ T ⊎ .. ⊎ S ≡ T ⊎ .. ⊎ Tn ≡ T 40 | -- ≅ (T1 ≡ T ⊎ .. ⊎ Tn ≡ T) ⊎ S ≡ T 41 | -- ≅ (G - u) ∋ T ⊎ S ≡ T 42 | -- 43 | -- thin and thick below are the witnesses 44 | 45 | _-_ : ∀ {A T} → (G : List A) → G ∋ T → List A 46 | (T ∷ G) - zero = G 47 | (S ∷ G) - suc x = S ∷ (G - x) 48 | 49 | infix 35 _-_ 50 | 51 | length-del : ∀ {A}{S : A}{X} -> (u : X ∋ S) -> length X ≡ suc (length (X - u)) 52 | length-del zero = refl 53 | length-del (suc u) = cong suc (length-del u) 54 | 55 | thin : ∀ {A} {G : List A}{S} → (x : G ∋ S) → ∀ T → (G - x) ∋ T → G ∋ T 56 | thin zero _ y = suc y 57 | thin (suc x) _ zero = zero 58 | thin (suc x) _ (suc y) = suc (thin x _ y) 59 | 60 | suc-inj1 : ∀ {A : Set}{xs : List A}{x z} {i : xs ∋ x}{j : xs ∋ x} → _∋_.suc {S = z} i ≡ suc j → i ≡ j 61 | suc-inj1 refl = refl 62 | 63 | x∉thinx : ∀ {A} {G : List A}{S} → (x : G ∋ S) → (y : (G - x) ∋ S) -> x ≡ thin x S y -> ⊥ 64 | x∉thinx zero y () 65 | x∉thinx (suc x) zero () 66 | x∉thinx (suc x) (suc y) eq = x∉thinx x y (suc-inj1 eq) 67 | 68 | thick : ∀ {A}{G : List A}{S T} → (x : G ∋ S) → (y : G ∋ T) → (∃ \ z → thin x T z ≡ y) ⊎ x ≅∋ y 69 | thick zero zero = inj₂ refl` 70 | thick zero (suc y) = inj₁ (y , refl) 71 | thick (suc x) zero = inj₁ (zero , refl) 72 | thick (suc x) (suc y) with thick x y 73 | thick (suc x) (suc y) | inj₁ (z , eq) = inj₁ (suc z , cong suc eq) 74 | thick (suc .y) (suc y) | inj₂ refl` = inj₂ refl` 75 | 76 | thick-refl : ∀ {A}{G : List A}{S} → (x : G ∋ S) → thick x x ≡ inj₂ refl` 77 | thick-refl zero = refl 78 | thick-refl (suc x) rewrite thick-refl x = refl 79 | 80 | thick-thin : ∀ {A}{G : List A}{S T} (x : G ∋ S) → (y : (G - x) ∋ T) → thick x (thin x T y) ≡ inj₁ (y , refl) 81 | thick-thin zero y = refl 82 | thick-thin (suc x) zero = refl 83 | thick-thin (suc x) (suc y) rewrite thick-thin x y = refl 84 | 85 | thin-inj : ∀ {A : Set}{xs : List A}{x y : A} -> (v : xs ∋ x) -> {i j : (xs - v) ∋ y} -> thin v y i ≡ thin v y j -> i ≡ j 86 | thin-inj v {i} {j} eq with cong (\ x -> maybe′ (\ x -> Maybe.just (proj₁ x)) nothing (isInj₁ (thick v x))) eq 87 | ... | p rewrite thick-thin v i | thick-thin v j with p 88 | thin-inj v {i} {.i} eq | p | refl = refl 89 | 90 | _≅∋?_ : ∀ {A : Set} {G : List A} {S} (u : G ∋ S) {T} (v : G ∋ T) -> Dec (u ≅∋ v) 91 | u ≅∋? v with thick u v 92 | ... | inj₂ eq = yes eq 93 | _≅∋?_ {S = S} u ._ | inj₁ (w , refl) = no (aux w) where 94 | aux : ∀ {T} (w : _ ∋ T) -> Σ (S ≡ T) (λ x → u ≅ thin u T w) → ⊥ 95 | aux w (refl , eq) = x∉thinx u w (≅-to-≡ eq) 96 | 97 | 98 | any? : ∀ {A : Set}{G : List A}{P : ∀ {S} -> G ∋ S -> Set} -> (∀ {S} v -> Dec (P {S} v)) -> Dec (∃ \ S -> ∃ (P {S})) 99 | any? {_} {[]} dec = no (\ {(_ , () , _)}) 100 | any? {_} {S ∷ G} dec 101 | with dec zero | any? (\ v -> dec (suc v)) 102 | ... | yes p | _ = yes (_ , zero , p) 103 | ... | no ¬p | yes (T , v , p) = yes (T , suc v , p) 104 | ... | no ¬p | no ¬q = no λ {(._ , zero , p) → ¬p p; 105 | (_ , suc v , p) → ¬q (_ , v , p)} 106 | -------------------------------------------------------------------------------- /Vars/MatchTwo.agda: -------------------------------------------------------------------------------- 1 | module Vars.MatchTwo where 2 | 3 | open import Data.Product 4 | open import Data.Empty 5 | open import Data.Sum 6 | 7 | open import Support.Equality 8 | open ≡-Reasoning 9 | 10 | open import Vars 11 | 12 | abstract 13 | _-[_,_] : ∀ {A} (G : List A) -> ∀ {S T} -> G ∋ S -> G ∋ T -> List A 14 | G -[ u , v ] with thick u v 15 | ... | inj₂ _ = G - u 16 | ... | inj₁ (v' , _) = (G - u) - v' 17 | 18 | data One⊎Other {A}{G : List A}{S S1} (x : G ∋ S)(z : G ∋ S1) : ∀ {T} -> G ∋ T -> Set where 19 | one : ∀ {y} -> (eq : x ≡ y) -> One⊎Other x z y 20 | other : ∀ {y} -> (neq : (eq : S1 ≡ S) -> ¬ x ≡ subst (_∋_ G) eq y) (eq : z ≡ y) -> One⊎Other x z y 21 | 22 | abstract 23 | thin[_,_] : ∀ {A}{G : List A}{S S1 T} → (x : G ∋ S)(z : G ∋ S1) → (y : (G -[ x , z ]) ∋ T) -> G ∋ T 24 | thin[ x , z ] y with thick x z 25 | thin[ x , z ] y | inj₁ (z' , _) = thin x _ (thin z' _ y) 26 | thin[ x , z ] y | inj₂ _ = thin x _ y 27 | 28 | thin[_,_]-inj : ∀ {A}{G : List A}{S S1 T} → (x : G ∋ S)(z : G ∋ S1) → {y y' : (G -[ x , z ]) ∋ T} 29 | -> thin[ x , z ] y ≡ thin[ x , z ] y' -> y ≡ y' 30 | thin[ x , z ]-inj eq with thick x z 31 | ... | inj₁ (z' , _) = thin-inj z' (thin-inj x eq) 32 | ... | inj₂ _ = thin-inj x eq 33 | 34 | thin[_,_]-disjoint : ∀ {A}{G : List A}{S S1 T} → (x : G ∋ S)(z : G ∋ S1) → {y : (G -[ x , z ]) ∋ T} 35 | -> One⊎Other x z (thin[ x , z ] y) -> ⊥ 36 | thin[ x , z ]-disjoint eq with thick x z 37 | thin[ x , ._ ]-disjoint (one eq) | inj₁ (z' , refl) = x∉thinx x _ eq 38 | thin[ x , ._ ]-disjoint (other _ eq) | inj₁ (z' , refl) = x∉thinx z' _ (thin-inj x eq) 39 | thin[ x , .x ]-disjoint (one eq) | inj₂ refl` = x∉thinx x _ eq 40 | thin[ x , .x ]-disjoint (other neq eq) | inj₂ refl` = x∉thinx x _ eq 41 | 42 | data Thick {A}{G : List A}{S S1} (x : G ∋ S)(z : G ∋ S1) {T : _} (y : G ∋ T) : Set where 43 | one⊎other : (eq : One⊎Other x z y) -> Thick x z y 44 | neither : (w : G -[ x , z ] ∋ T) -> (eq : thin[ x , z ] w ≡ y) -> Thick x z y 45 | 46 | abstract 47 | thick[_,_] : ∀ {A}{G : List A}{S S1 T} → (x : G ∋ S)(z : G ∋ S1) → (y : G ∋ T) → Thick x z y 48 | thick[ x , z ] y with thick x z | thick x y | neither {x = x} {z = z} {y = y} 49 | thick[ x , .x ] ._ | inj₂ refl` | inj₁ (y' , refl) | ne = ne y' refl 50 | thick[ x , .x ] .x | inj₂ refl` | inj₂ refl` | ne = one⊎other (one refl) 51 | thick[ x , z ] .x | inj₁ (z' , eq) | inj₂ refl` | ne = one⊎other (one refl) 52 | thick[ x , z ] y | inj₁ (z' , eq) | inj₁ (y' , eq') | ne with thick z' y' 53 | thick[ x , ._ ] ._ | inj₁ (z' , refl) | inj₁ (._ , refl) | ne | inj₁ (y'' , refl) = ne y'' refl 54 | thick[ x , z ] ._ | inj₁ (z' , eq) | inj₁ (._ , refl) | ne | inj₂ refl` = one⊎other (other neq (sym eq)) 55 | where 56 | neq : ∀ {S v} (eq₁ : S ≡ _) → x ≡ subst (_∋_ _) eq₁ (thin x _ v) → ⊥ 57 | neq refl eq = x∉thinx x _ eq 58 | 59 | thick[_,_]-refl : ∀ {A}{G : List A}{S S1} → (x : G ∋ S)(z : G ∋ S1) → thick[ x , z ] x ≡ one⊎other (one refl) 60 | thick[ x , z ]-refl with thick[ x , z ] x 61 | thick[ x , z ]-refl | one⊎other (one refl) = refl 62 | thick[ x , z ]-refl | one⊎other (other neq eq) = ⊥-elim (neq refl refl) 63 | thick[ x , z ]-refl | neither w eq = ⊥-elim (thin[ x , z ]-disjoint (one (sym eq))) 64 | 65 | 66 | 67 | thick[_,_]-thin : ∀ {A} {G : List A}{S S1 T} → (x : G ∋ S)(z : G ∋ S1)(y : G -[ x , z ] ∋ T) → thick[ x , z ] (thin[ x , z ] y) ≡ neither y refl 68 | thick[ x , z ]-thin y with thick[ x , z ] (thin[ x , z ] y) 69 | ... | one⊎other eq = ⊥-elim (thin[ x , z ]-disjoint eq) 70 | ... | neither y' eq with thin[ x , z ]-inj eq | eq 71 | thick[ x , z ]-thin y | neither .y eq | refl | refl = refl 72 | 73 | thick[_,_]-refl₂ : ∀ {A}{G : List A}{S S1} → (x : G ∋ S)(z : G ∋ S1) → ∃ \ eq -> thick[ x , z ] z ≡ one⊎other eq 74 | thick[ x , z ]-refl₂ with thick[ x , z ] z 75 | thick[ x , z ]-refl₂ | one⊎other eq = _ , refl 76 | thick[ x , z ]-refl₂ | neither w eq with thick x z 77 | thick[ x , ._ ]-refl₂ | neither w eq | inj₁ (z' , refl) = ⊥-elim (thin[ x , _ ]-disjoint (other neq (sym eq))) 78 | where 79 | neq : ∀ {S} {z : _ ∋ S}{v : _ ∋ S} (eq₁ : S ≡ _) → x ≡ subst (_∋_ _) eq₁ (thin[ x , z ] v) → ⊥ 80 | neq refl eq = thin[ x , _ ]-disjoint (one eq) 81 | 82 | thick[ .z , z ]-refl₂ | neither w eq | inj₂ refl` = ⊥-elim (thin[ z , z ]-disjoint (one (sym eq))) -- 83 | 84 | 85 | thick[_,_]-refl₂₂ : ∀ {A}{G : List A}{S S1} → (x : G ∋ S)(z : G ∋ S1) 86 | → ((eq : S1 ≡ S) -> ¬ x ≡ subst (_∋_ G) eq z) 87 | → ∃ \ neq -> thick[ x , z ] z ≡ one⊎other (other neq refl) 88 | thick[ x , z ]-refl₂₂ neq with thick[ x , z ] z | thick[ x , z ]-refl₂ 89 | thick[ x , z ]-refl₂₂ neq | .(one⊎other (one eq)) | one eq , refl = ⊥-elim (neq refl eq) 90 | thick[ x , z ]-refl₂₂ neq₁ | .(one⊎other (other neq refl)) | other neq refl , refl = _ , refl 91 | -------------------------------------------------------------------------------- /Vars/SumIso.agda: -------------------------------------------------------------------------------- 1 | module Vars.SumIso {A : Set} where 2 | 3 | open import Relation.Binary.PropositionalEquality 4 | open import Data.Sum renaming (map to map⊎) 5 | open import Function using (id) 6 | 7 | open import Vars 8 | 9 | split# : ∀ m {n}{T : A} -> _∋_ (m ++ n) T -> _∋_ m T ⊎ _∋_ n T 10 | split# [] i = inj₂ i 11 | split# (_ ∷ n) zero = inj₁ zero 12 | split# (_ ∷ n) (suc i) = map⊎ suc id (split# n i) 13 | 14 | split : ∀ {m} {n}{T} -> _∋_ (m ++ n) T -> _∋_ m T ⊎ _∋_ n T 15 | split = split# _ 16 | 17 | glue# : ∀ m {n} {T : A} -> _∋_ m T ⊎ _∋_ n T -> _∋_ (m ++ n) T 18 | glue# [] (inj₁ ()) 19 | glue# [] (inj₂ i) = i 20 | glue# (_ ∷ n) (inj₁ zero) = zero 21 | glue# (_ ∷ n) (inj₁ (suc i)) = suc (glue# n (inj₁ i)) 22 | glue# (_ ∷ n) (inj₂ i) = suc (glue# n (inj₂ i)) 23 | 24 | glue : ∀ {m n T} -> _∋_ m T ⊎ _∋_ n T -> _∋_ (m ++ n) T 25 | glue = glue# _ 26 | 27 | glue∘split≡id : ∀ {m n T} (i : _∋_ (m ++ n) T) -> glue {m} {n} (split i) ≡ i 28 | glue∘split≡id {[]} i = refl 29 | glue∘split≡id {_ ∷ n} zero = refl 30 | glue∘split≡id {_ ∷ n} (suc i) with split {n} i | glue∘split≡id {n} i 31 | glue∘split≡id {_ ∷ n} (suc i) | inj₁ x | glueinj₁x≡i rewrite glueinj₁x≡i = refl 32 | glue∘split≡id {_ ∷ n'} (suc i) | inj₂ y | glueinj₂y≡i rewrite glueinj₂y≡i = refl 33 | 34 | split∘glue≡id : ∀ {m n T} (i : _∋_ m T ⊎ _∋_ n T) -> split {m} {n} (glue i) ≡ i 35 | split∘glue≡id {[]} (inj₁ ()) 36 | split∘glue≡id {[]} (inj₂ _) = refl 37 | split∘glue≡id {_ ∷ n} (inj₁ zero) = refl 38 | split∘glue≡id {_ ∷ m} {n} (inj₁ (suc i)) rewrite split∘glue≡id {m} {n} (inj₁ i) = refl 39 | split∘glue≡id {_ ∷ m} {n} (inj₂ i) rewrite split∘glue≡id {m} {n} (inj₂ i) = refl 40 | --------------------------------------------------------------------------------