├── .gitignore ├── README.mkd ├── agda └── Lambda │ ├── Confluence.agda │ ├── Core.agda │ └── Properties.agda └── coq ├── CL.v ├── LN └── Untyped.v ├── Makefile ├── Origin └── Untyped.v ├── _CoqProject ├── deBruijn ├── F.v ├── STLC.v └── Untyped.v └── lib ├── Relations_ext.v ├── seq_ext.v ├── seq_ext_base.v └── ssrnat_ext.v /.gitignore: -------------------------------------------------------------------------------- 1 | *.glob 2 | *.vo 3 | *.vok 4 | *.vos 5 | .*.aux 6 | .lia.cache 7 | /coq/.Makefile.coq.d 8 | /coq/.deps 9 | /coq/Makefile.coq 10 | /coq/Makefile.coq.conf 11 | -------------------------------------------------------------------------------- /README.mkd: -------------------------------------------------------------------------------- 1 | # A Formalization of Typed and Untyped λ-Calculi in Coq and Agda2 2 | 3 | By Kazuhiko Sakaguchi, Nov. 2011 - July 2020. 4 | 5 | This is a formalization of the λ-calculus written in Coq and Agda2, and contains the following definitions and theorems: 6 | 7 | * SKI combinator calculus: coq/CL.v 8 | * Confluence of weak reduction 9 | * Untyped lambda calculus: coq/deBruijn/Untyped.v, agda/Lambda/\*.agda 10 | * Confluence of beta reduction 11 | * STLC (simply typed lambda calculus): coq/deBruijn/STLC.v 12 | * Subject reduction (type preservation) theorem 13 | * Strong normalization theorem (three different proofs are available) 14 | * System F (second order typed lambda calculus): coq/deBruijn/F.v 15 | * Subject reduction (type preservation) theorem 16 | * Strong normalization theorem (three different proofs are available) 17 | 18 | ## Useful Literature 19 | 20 | * T. Altenkirch. A Formalization of the Strong Normalization Proof for System F in LEGO. Typed Lambda Calculi and Applications, 13-28, 1993. 21 | * N.G. de Bruijn. Lambda Calculus Notation with Nameless Dummies: A Tool for Automatic Formula Manipulation, with Application to the Church-Rosser Theorem. Indagationes Mathematicae, 75(5), 381–392, 1972. 22 | * J.H. Gallier. On Girard's “Candidats de Réductibilité". In Logic and Computer Science. P. Odifreddi, Editor, Academic Press, 123-203, 1989. 23 | * J.-Y. Girard. Interprétation fonctionnelle et élimination des coupures de l'arithmétique d'ordre supérieur. Thèse d'état, Université de Paris 7. 1972. 24 | * J.-Y. Girard, Y. Lafont and P. Taylor. Proofs and Types. Cambridge University Press, 1989. 25 | * J.R. Hindley and J.P. Seldin. Lambda-Calculus and Combinators: an Introduction. Cambridge University Press, 2008. 3rd edition. 26 | * C.-K. Hur. Heq : a Coq library for Heterogeneous Equality. http://sf.snu.ac.kr/gil.hur/Heq/ 27 | * T. Nipkow. More Church-Rosser Proofs (in Isabelle/HOL). Journal of Automated Reasoning, 26(1), 51–66, 2001. 28 | * A. Popescu, C.J. Osborn and E.L. Gunter. Strong Normalization for System F by HOAS on Top of FOAS. Logic in Computer Science, 31-40, 2010. 29 | * M. Takahashi. Parallel Reductions in λ-Calculus. Information and Computation, 118(1), 120-127, 1995. 30 | * 大堀淳. プログラミング言語の基礎理論, 情報数学講座, 第9巻, 共立出版, 1997. 31 | * 横内寛文. プログラム意味論, 情報数学講座, 第7巻, 共立出版, 1994. 32 | -------------------------------------------------------------------------------- /agda/Lambda/Confluence.agda: -------------------------------------------------------------------------------- 1 | 2 | module Lambda.Confluence where 3 | 4 | open import Function 5 | open import Data.Product 6 | open import Data.Nat 7 | open import Data.Nat.Properties 8 | open import Data.Star 9 | open import Relation.Nullary 10 | open import Relation.Binary 11 | open import Relation.Binary.PropositionalEquality 12 | 13 | open import Lambda.Core 14 | open import Lambda.Properties 15 | 16 | →β*appl : ∀ {t1 t1' t2} → t1 →β* t1' → tapp t1 t2 →β* tapp t1' t2 17 | →β*appl ε = ε 18 | →β*appl (r1 ◅ r2) = →βappl r1 ◅ →β*appl r2 19 | 20 | →β*appr : ∀ {t1 t2 t2'} → t2 →β* t2' → tapp t1 t2 →β* tapp t1 t2' 21 | →β*appr ε = ε 22 | →β*appr (r1 ◅ r2) = →βappr r1 ◅ →β*appr r2 23 | 24 | →β*abs : ∀ {t t'} → t →β* t' → tabs t →β* tabs t' 25 | →β*abs ε = ε 26 | →β*abs (r1 ◅ r2) = →βabs r1 ◅ →β*abs r2 27 | 28 | parRefl : ∀ {t} → t →βP t 29 | parRefl {tvar _} = →βPvar 30 | parRefl {tapp _ _} = →βPapp parRefl parRefl 31 | parRefl {tabs _} = →βPabs parRefl 32 | 33 | →β⊂→βP : _→β_ ⇒ _→βP_ 34 | →β⊂→βP →βbeta = →βPbeta parRefl parRefl 35 | →β⊂→βP (→βappl r) = →βPapp (→β⊂→βP r) parRefl 36 | →β⊂→βP (→βappr r) = →βPapp parRefl (→β⊂→βP r) 37 | →β⊂→βP (→βabs r) = →βPabs (→β⊂→βP r) 38 | 39 | →βP⊂→β* : _→βP_ ⇒ _→β*_ 40 | →βP⊂→β* →βPvar = ε 41 | →βP⊂→β* (→βPapp r1 r2) = →β*appl (→βP⊂→β* r1) ◅◅ →β*appr (→βP⊂→β* r2) 42 | →βP⊂→β* (→βPabs r) = →β*abs (→βP⊂→β* r) 43 | →βP⊂→β* (→βPbeta r1 r2) = →β*appl (→β*abs (→βP⊂→β* r1)) ◅◅ →β*appr (→βP⊂→β* r2) ◅◅ →βbeta ◅ ε 44 | 45 | shiftConservation→β : ∀ {d c} → _→β_ ⇒ ((λ a b → a → b) on Shifted d c) 46 | shiftConservation→β {d} {c} {tapp (tabs t1) t2} →βbeta (sapp (sabs s1) s2) = 47 | betaShifted2 s1 s2 48 | shiftConservation→β (→βappl p) (sapp s1 s2) = sapp (shiftConservation→β p s1) s2 49 | shiftConservation→β (→βappr p) (sapp s1 s2) = sapp s1 (shiftConservation→β p s2) 50 | shiftConservation→β (→βabs p) (sabs s1) = sabs (shiftConservation→β p s1) 51 | 52 | shiftConservation→β* : ∀ {d c} → _→β*_ ⇒ ((λ a b → a → b) on Shifted d c) 53 | shiftConservation→β* ε s = s 54 | shiftConservation→β* (p1 ◅ p2) s = shiftConservation→β* p2 (shiftConservation→β p1 s) 55 | 56 | shiftConservation→βP : ∀ {d c t1 t2} → t1 →βP t2 → Shifted d c t1 → Shifted d c t2 57 | shiftConservation→βP p s = shiftConservation→β* (→βP⊂→β* p) s 58 | 59 | shiftLemma : ∀ {t t' d c} → t →βP t' → shift d c t →βP shift d c t' 60 | shiftLemma →βPvar = parRefl 61 | shiftLemma (→βPapp r1 r2) = →βPapp (shiftLemma r1) (shiftLemma r2) 62 | shiftLemma (→βPabs r) = →βPabs (shiftLemma r) 63 | shiftLemma {d = d} {c} (→βPbeta {t1} {t1'} {t2} {t2'} r1 r2) = r where 64 | open ≡-Reasoning 65 | eq : shift d c (unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ])) ≡ 66 | unshift 1 0 (shift d (suc c) t1' [ 0 ≔ shift 1 0 (shift d c t2') ]) 67 | eq = begin 68 | shift d c (unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ])) 69 | ≡⟨ shiftUnshiftSwap z≤n (betaShifted' 0 t1' t2') ⟩ 70 | unshift 1 0 (shift d (c + 1) (t1' [ 0 ≔ shift 1 0 t2' ])) 71 | ≡⟨ cong (unshift 1 0) $ begin 72 | shift d (c + 1) (t1' [ 0 ≔ shift 1 0 t2' ]) 73 | ≡⟨ cong (λ c → shift d c (t1' [ 0 ≔ shift 1 0 t2' ])) (+-comm c 1) ⟩ 74 | shift d (suc c) (t1' [ 0 ≔ shift 1 0 t2' ]) 75 | ≡⟨ shiftSubstSwap (m≤m+n 1 c) t1' (shift 1 0 t2') ⟩ 76 | shift d (suc c) t1' [ 0 ≔ shift d (suc c) (shift 1 0 t2') ] 77 | ≡⟨ cong (λ t → shift d (suc c) t1' [ 0 ≔ t ]) $ begin 78 | shift d (suc c) (shift 1 0 t2') 79 | ≡⟨ cong (λ c → shift d c (shift 1 0 t2')) (+-comm 1 c) ⟩ 80 | shift d (c + 1) (shift 1 0 t2') 81 | ≡⟨ sym (shiftShiftSwap 1 0 d c t2' z≤n) ⟩ 82 | shift 1 0 (shift d c t2') ∎ 83 | ⟩ 84 | shift d (suc c) t1' [ 0 ≔ shift 1 0 (shift d c t2') ] ∎ 85 | ⟩ 86 | unshift 1 0 (shift d (suc c) t1' [ 0 ≔ shift 1 0 (shift d c t2') ]) ∎ 87 | r : shift d c (tapp (tabs t1) t2) →βP 88 | shift d c (unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ])) 89 | r rewrite eq = →βPbeta (shiftLemma r1) (shiftLemma r2) 90 | 91 | unshiftLemma : ∀ {t t' d c} → t →βP t' → Shifted d c t → unshift d c t →βP unshift d c t' 92 | unshiftLemma →βPvar _ = parRefl 93 | unshiftLemma (→βPapp r1 r2) (sapp s1 s2) = →βPapp (unshiftLemma r1 s1) (unshiftLemma r2 s2) 94 | unshiftLemma (→βPabs r) (sabs s) = →βPabs (unshiftLemma r s) 95 | unshiftLemma {d = d} {c} (→βPbeta {t1} {t1'} {t2} {t2'} r1 r2) (sapp (sabs s1) s2) = r where 96 | open ≡-Reasoning 97 | s1' = shiftConservation→βP r1 s1 98 | s2' = shiftConservation→βP r2 s2 99 | eq : unshift d c (unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ])) ≡ 100 | unshift 1 0 (unshift d (suc c) t1' [ 0 ≔ shift 1 0 (unshift d c t2') ]) 101 | eq = begin 102 | unshift d c (unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ])) 103 | ≡⟨ unshiftUnshiftSwap z≤n (betaShifted' 0 t1' t2') (betaShifted2 s1' s2') ⟩ 104 | unshift 1 0 (unshift d (c + 1) (t1' [ 0 ≔ shift 1 0 t2' ])) 105 | ≡⟨ cong (unshift 1 0) $ begin 106 | unshift d (c + 1) (t1' [ 0 ≔ shift 1 0 t2' ]) 107 | ≡⟨ cong (λ c → unshift d c (t1' [ 0 ≔ shift 1 0 t2' ])) (+-comm c 1) ⟩ 108 | unshift d (suc c) (t1' [ 0 ≔ shift 1 0 t2' ]) 109 | ≡⟨ unshiftSubstSwap2 (s≤s z≤n) s1' (shiftShifted' z≤n s2') ⟩ 110 | unshift d (suc c) t1' [ 0 ≔ unshift d (suc c) (shift 1 0 t2') ] 111 | ≡⟨ cong (λ t → unshift d (suc c) t1' [ 0 ≔ t ]) $ begin 112 | unshift d (suc c) (shift 1 0 t2') 113 | ≡⟨ cong (λ c → unshift d c (shift 1 0 t2')) (+-comm 1 c) ⟩ 114 | unshift d (c + 1) (shift 1 0 t2') 115 | ≡⟨ sym (unshiftShiftSwap z≤n s2') ⟩ 116 | shift 1 0 (unshift d c t2') ∎ 117 | ⟩ 118 | unshift d (suc c) t1' [ 0 ≔ shift 1 0 (unshift d c t2') ] ∎ 119 | ⟩ 120 | unshift 1 0 (unshift d (suc c) t1' [ 0 ≔ shift 1 0 (unshift d c t2') ]) ∎ 121 | r : tapp (tabs (unshift d (suc c) t1)) (unshift d c t2) →βP 122 | unshift d c (unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ])) 123 | r rewrite eq = →βPbeta (unshiftLemma r1 s1) (unshiftLemma r2 s2) 124 | 125 | substLemma : ∀ {n t1 t1' t2 t2'} → 126 | t1 →βP t1' → t2 →βP t2' → t1 [ n ≔ t2 ] →βP t1' [ n ≔ t2' ] 127 | substLemma {n} {tvar n'} →βPvar r1 with n ≟ n' 128 | ...| yes p = r1 129 | ...| no p = →βPvar 130 | substLemma (→βPapp r1 r2) r3 = →βPapp (substLemma r1 r3) (substLemma r2 r3) 131 | substLemma (→βPabs r1) r2 = →βPabs $ substLemma r1 $ shiftLemma r2 132 | substLemma {n} {t2 = t3} {t3'} (→βPbeta {t1} {t1'} {t2} {t2'} r1 r2) r3 = r where 133 | open ≡-Reasoning 134 | eq : unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ]) [ n ≔ t3' ] ≡ 135 | unshift 1 0 ((t1' [ suc n ≔ shift 1 0 t3' ]) [ 0 ≔ shift 1 0 (t2' [ n ≔ t3' ]) ]) 136 | eq = begin 137 | unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ]) [ n ≔ t3' ] 138 | ≡⟨ sym (unshiftSubstSwap' (t1' [ 0 ≔ shift 1 0 t2' ]) t3' (betaShifted' 0 t1' t2')) ⟩ 139 | unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ] [ suc n ≔ shift 1 0 t3' ]) 140 | ≡⟨ cong (unshift 1 0) (substSubstSwap n 0 t1' t2' t3') ⟩ 141 | unshift 1 0 ((t1' [ suc n ≔ shift 1 0 t3' ]) [ 0 ≔ shift 1 0 (t2' [ n ≔ t3' ]) ]) ∎ 142 | r : tapp (tabs (t1 [ suc n ≔ shift 1 0 t3 ])) (t2 [ n ≔ t3 ]) →βP 143 | (unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ])) [ n ≔ t3' ] 144 | r rewrite eq = →βPbeta (substLemma r1 (shiftLemma r3)) (substLemma r2 r3) 145 | 146 | starLemma : ∀ {t t'} → t →βP t' → t' →βP t * 147 | starLemma →βPvar = →βPvar 148 | starLemma {tapp (tvar n) t2} (→βPapp p1 p2) = 149 | →βPapp (starLemma p1) (starLemma p2) 150 | starLemma {tapp (tapp t1l t1r) t2} (→βPapp p1 p2) = 151 | →βPapp (starLemma p1) (starLemma p2) 152 | starLemma {tapp (tabs t1) t2} {tapp (tvar _) t2'} (→βPapp () p2) 153 | starLemma {tapp (tabs t1) t2} {tapp (tapp _ _) t2'} (→βPapp () p2) 154 | starLemma {tapp (tabs t1) t2} {tapp (tabs t1') t2'} (→βPapp (→βPabs p1) p2) = 155 | →βPbeta (starLemma p1) (starLemma p2) 156 | starLemma (→βPabs p1) = →βPabs (starLemma p1) 157 | starLemma {tapp (tabs t1) t2} (→βPbeta {.t1} {t1'} {.t2} {t2'} p1 p2) = 158 | unshiftLemma 159 | (substLemma (starLemma p1) (shiftLemma (starLemma p2))) 160 | (betaShifted 0 t1' t2') 161 | 162 | Diamond : ∀ {ℓ} {A : Set ℓ} (_R_ : Rel A ℓ) → Set ℓ 163 | Diamond _R_ = ∀ {t1 t2 t3} → t1 R t2 → t1 R t3 → ∃ (λ t → t2 R t × t3 R t) 164 | 165 | SemiConfluence : ∀ {ℓ} {A : Set ℓ} → Rel A ℓ → Set ℓ 166 | SemiConfluence R = ∀ {t1 t2 t3} → R t1 t2 → Star R t1 t3 → ∃ (λ t → Star R t2 t × Star R t3 t) 167 | 168 | Confluence : ∀ {ℓ} {A : Set ℓ} → Rel A ℓ → Set ℓ 169 | Confluence R = Diamond (Star R) 170 | 171 | Diamond⇒SemiConfluence : ∀ {ℓ} {A : Set ℓ} {R : Rel A ℓ} → Diamond R → SemiConfluence R 172 | Diamond⇒SemiConfluence diamond {t1} {t2} {.t1} r1 ε = t2 , ε , r1 ◅ ε 173 | Diamond⇒SemiConfluence diamond {t1} {t2} {t3} r1 (r2 ◅ r3) = 174 | proj₁ d' , proj₁ (proj₂ d) ◅ proj₁ (proj₂ d') , proj₂ (proj₂ d') where 175 | d = diamond r1 r2 176 | d' = Diamond⇒SemiConfluence diamond (proj₂ (proj₂ d)) r3 177 | 178 | SemiConfluence⇒Confluence : 179 | ∀ {ℓ} {A : Set ℓ} {R : Rel A ℓ} → SemiConfluence R → Confluence R 180 | SemiConfluence⇒Confluence sconfluence {t1} {.t1} {t3} ε r2 = t3 , r2 , ε 181 | SemiConfluence⇒Confluence sconfluence {t1} {t2} {t3} (r1 ◅ r2) r3 = 182 | proj₁ sc' , proj₁ (proj₂ sc') , proj₂ (proj₂ sc) ◅◅ proj₂ (proj₂ sc') where 183 | sc = sconfluence r1 r3 184 | sc' = SemiConfluence⇒Confluence sconfluence r2 (proj₁ (proj₂ sc)) 185 | 186 | diamond→βP : Diamond _→βP_ 187 | diamond→βP {t1} r1 r2 = (t1 *) , starLemma r1 , starLemma r2 188 | 189 | confluence→βP : Confluence _→βP_ 190 | confluence→βP = SemiConfluence⇒Confluence $ Diamond⇒SemiConfluence diamond→βP 191 | 192 | confluence→β : Confluence _→β_ 193 | confluence→β r1 r2 = 194 | proj₁ c , 195 | Data.Star.concat (Data.Star.map →βP⊂→β* (proj₁ (proj₂ c))) , 196 | Data.Star.concat (Data.Star.map →βP⊂→β* (proj₂ (proj₂ c))) where 197 | c = confluence→βP (Data.Star.map →β⊂→βP r1) (Data.Star.map →β⊂→βP r2) 198 | 199 | -------------------------------------------------------------------------------- /agda/Lambda/Core.agda: -------------------------------------------------------------------------------- 1 | 2 | module Lambda.Core where 3 | 4 | import Level 5 | open import Function 6 | open import Data.Nat 7 | open import Data.Star 8 | open import Relation.Nullary 9 | open import Relation.Binary 10 | 11 | data Term : Set where 12 | tvar : ℕ → Term 13 | tapp : Term → Term → Term 14 | tabs : Term → Term 15 | 16 | shift : ℕ → ℕ → Term → Term 17 | shift d c (tvar n) with c ≤? n 18 | ...| yes p = tvar (n + d) 19 | ...| no p = tvar n 20 | shift d c (tapp t1 t2) = tapp (shift d c t1) (shift d c t2) 21 | shift d c (tabs t) = tabs $ shift d (suc c) t 22 | 23 | unshift : ℕ → ℕ → Term → Term 24 | unshift d c (tvar n) with c ≤? n 25 | ...| yes p = tvar (n ∸ d) 26 | ...| no p = tvar n 27 | unshift d c (tapp t1 t2) = tapp (unshift d c t1) (unshift d c t2) 28 | unshift d c (tabs t) = tabs $ unshift d (suc c) t 29 | 30 | nshift : ℕ → Term → Term 31 | nshift 0 t = t 32 | nshift (suc n) t = shift 1 0 $ nshift n t 33 | 34 | _[_≔_] : Term → ℕ → Term → Term 35 | tvar n' [ n ≔ t ] with n ≟ n' 36 | ...| yes p = t 37 | ...| no p = tvar n' 38 | tapp t1 t2 [ n ≔ t ] = tapp (t1 [ n ≔ t ]) (t2 [ n ≔ t ]) 39 | tabs t1 [ n ≔ t ] = tabs $ t1 [ suc n ≔ shift 1 0 t ] 40 | 41 | data _→β_ : Rel Term Level.zero where 42 | →βbeta : ∀ {t1 t2} → tapp (tabs t1) t2 →β unshift 1 0 (t1 [ 0 ≔ shift 1 0 t2 ]) 43 | →βappl : ∀ {t1 t1' t2} → t1 →β t1' → tapp t1 t2 →β tapp t1' t2 44 | →βappr : ∀ {t1 t2 t2'} → t2 →β t2' → tapp t1 t2 →β tapp t1 t2' 45 | →βabs : ∀ {t t'} → t →β t' → tabs t →β tabs t' 46 | 47 | _→β*_ : Rel Term Level.zero 48 | _→β*_ = Star _→β_ 49 | 50 | data _→βP_ : Rel Term Level.zero where 51 | →βPvar : ∀ {n} → tvar n →βP tvar n 52 | →βPapp : ∀ {t1 t1' t2 t2'} → t1 →βP t1' → t2 →βP t2' → tapp t1 t2 →βP tapp t1' t2' 53 | →βPabs : ∀ {t t'} → t →βP t' → tabs t →βP tabs t' 54 | →βPbeta : ∀ {t1 t1' t2 t2'} → t1 →βP t1' → t2 →βP t2' → 55 | tapp (tabs t1) t2 →βP unshift 1 0 (t1' [ 0 ≔ shift 1 0 t2' ]) 56 | 57 | _* : Term → Term 58 | tvar n * = tvar n 59 | tapp (tabs t1) t2 * = unshift 1 0 (t1 * [ 0 ≔ shift 1 0 (t2 *) ]) 60 | tapp t1 t2 * = tapp (t1 *) (t2 *) 61 | tabs t * = tabs (t *) 62 | 63 | -------------------------------------------------------------------------------- /agda/Lambda/Properties.agda: -------------------------------------------------------------------------------- 1 | 2 | module Lambda.Properties where 3 | 4 | open import Function 5 | open import Algebra 6 | open import Data.Empty 7 | open import Data.Nat 8 | open import Data.Nat.Properties 9 | open import Relation.Nullary 10 | open import Relation.Binary 11 | open import Relation.Binary.PropositionalEquality 12 | 13 | open import Lambda.Core 14 | 15 | +-comm = CommutativeSemiring.+-comm commutativeSemiring 16 | +-assoc = CommutativeSemiring.+-assoc commutativeSemiring 17 | 18 | ≱⇒< : ∀ {n m} → ¬ n ≤ m → m < n 19 | ≱⇒< {zero} f = ⊥-elim $ f z≤n 20 | ≱⇒< {suc _} {zero} f = s≤s z≤n 21 | ≱⇒< {suc _} {suc _} f = s≤s $ ≱⇒< $ λ p → f $ s≤s p 22 | 23 | <⇒≱ : ∀ {n m} → m < n → ¬ n ≤ m 24 | <⇒≱ (s≤s p1) (s≤s p2) = <⇒≱ p1 p2 25 | 26 | a-b+c≡a+c-b : ∀ {a b c} → b ≤ a → a ∸ b + c ≡ a + c ∸ b 27 | a-b+c≡a+c-b z≤n = refl 28 | a-b+c≡a+c-b {suc a} {suc b} {c} (s≤s p) = a-b+c≡a+c-b p 29 | 30 | ≤-addL : ∀ a {b c} → b ≤ c → a + b ≤ a + c 31 | ≤-addL zero p = p 32 | ≤-addL (suc a) p = s≤s (≤-addL a p) 33 | 34 | ≤-addL' : ∀ a {b c} → a ≤ c → b ≤ c ∸ a → a + b ≤ c 35 | ≤-addL' zero p1 p2 = p2 36 | ≤-addL' (suc a) {b} {zero} () p2 37 | ≤-addL' (suc a) {b} {suc n} (s≤s m≤n) p2 = s≤s $ ≤-addL' a m≤n p2 38 | 39 | ≤-addR : ∀ a {b c} → b ≤ c → b + a ≤ c + a 40 | ≤-addR a {b} {c} p rewrite +-comm b a | +-comm c a = ≤-addL a p 41 | 42 | ≤-addR' : ∀ a {b c} → a ≤ c → b ≤ c ∸ a → b + a ≤ c 43 | ≤-addR' a {b} {c} p1 p2 rewrite +-comm b a = ≤-addL' a p1 p2 44 | 45 | ≤-subL : ∀ a {b c} → a + b ≤ a + c → b ≤ c 46 | ≤-subL zero p = p 47 | ≤-subL (suc a) (s≤s p) = ≤-subL a p 48 | 49 | ≤-subL' : ∀ a {b c} → a + b ≤ c → b ≤ c ∸ a 50 | ≤-subL' zero p = p 51 | ≤-subL' (suc a) {zero} p = z≤n 52 | ≤-subL' (suc a) {suc b} {zero} () 53 | ≤-subL' (suc a) {suc b} {suc c} (s≤s p) = ≤-subL' a p 54 | 55 | ≤-subR : ∀ a {b c} → b + a ≤ c + a → b ≤ c 56 | ≤-subR a {b} {c} p = ≤-subL a p' where 57 | p' : a + b ≤ a + c 58 | p' rewrite +-comm a b | +-comm a c = p 59 | 60 | ≤-subR' : ∀ a {b c} → b + a ≤ c → b ≤ c ∸ a 61 | ≤-subR' a {b} {c} p rewrite +-comm b a = ≤-subL' a p 62 | 63 | ≤-sub : ∀ a {b c} → b ≤ c → b ∸ a ≤ c ∸ a 64 | ≤-sub zero p = p 65 | ≤-sub (suc a) {zero} p = z≤n 66 | ≤-sub (suc a) {suc b} (s≤s p) = ≤-sub a p 67 | 68 | ≡-subL' : ∀ a {b c} → a + b ≡ c → b ≡ c ∸ a 69 | ≡-subL' zero p = p 70 | ≡-subL' (suc a) {zero} {zero} p = refl 71 | ≡-subL' (suc a) {suc b} {zero} () 72 | ≡-subL' (suc a) {_} {suc n} p = ≡-subL' a $ cong pred p 73 | 74 | ≡-addL' : ∀ {a b c} → a ≤ c → b ≡ c ∸ a → a + b ≡ c 75 | ≡-addL' z≤n p2 = p2 76 | ≡-addL' {suc a} {b} {suc c} (s≤s p1) p2 = cong suc $ ≡-addL' p1 p2 77 | 78 | a∸b∸c≡a∸c∸b : ∀ a b c → a ∸ b ∸ c ≡ a ∸ c ∸ b 79 | a∸b∸c≡a∸c∸b a b c = 80 | begin 81 | a ∸ b ∸ c ≡⟨ ∸-+-assoc a b c ⟩ 82 | a ∸ (b + c) ≡⟨ cong (_∸_ a) (+-comm b c) ⟩ 83 | a ∸ (c + b) ≡⟨ sym (∸-+-assoc a c b) ⟩ 84 | a ∸ c ∸ b ∎ 85 | where open ≡-Reasoning 86 | 87 | a+b∸c≡a∸c+b : ∀ a b c → c ≤ a → a + b ∸ c ≡ a ∸ c + b 88 | a+b∸c≡a∸c+b a b c p = sym $ ≡-subL' c $ begin 89 | c + (a ∸ c + b) ≡⟨ sym (+-assoc c (a ∸ c) b) ⟩ 90 | c + (a ∸ c) + b ≡⟨ cong (λ n → n + b) (m+n∸m≡n p) ⟩ 91 | a + b ∎ 92 | where open ≡-Reasoning 93 | 94 | shiftVarEq1 : ∀ {d c n} → c ≤ n → shift d c (tvar n) ≡ tvar (n + d) 95 | shiftVarEq1 {d} {c} {n} p1 with c ≤? n 96 | ...| yes p2 = refl 97 | ...| no p2 = ⊥-elim $ p2 p1 98 | 99 | shiftVarEq2 : ∀ {d c n} → c ≰ n → shift d c (tvar n) ≡ tvar n 100 | shiftVarEq2 {d} {c} {n} p1 with c ≤? n 101 | ...| yes p2 = ⊥-elim $ p1 p2 102 | ...| no p2 = refl 103 | 104 | unshiftVarEq1 : ∀ {d c n} → c ≤ n → unshift d c (tvar n) ≡ tvar (n ∸ d) 105 | unshiftVarEq1 {d} {c} {n} p1 with c ≤? n 106 | ...| yes p2 = refl 107 | ...| no p2 = ⊥-elim $ p2 p1 108 | 109 | unshiftVarEq2 : ∀ {d c n} → c ≰ n → unshift d c (tvar n) ≡ tvar n 110 | unshiftVarEq2 {d} {c} {n} p1 with c ≤? n 111 | ...| yes p2 = ⊥-elim $ p1 p2 112 | ...| no p2 = refl 113 | 114 | shiftZero : ∀ c t → t ≡ shift 0 c t 115 | shiftZero c (tvar n) with c ≤? n 116 | ...| yes p = cong tvar $ +-comm 0 n 117 | ...| no p = refl 118 | shiftZero c (tapp t1 t2) = cong₂ tapp (shiftZero c t1) (shiftZero c t2) 119 | shiftZero c (tabs t) = cong tabs (shiftZero (suc c) t) 120 | 121 | shiftAdd : ∀ d d' c t → shift d c (shift d' c t) ≡ shift (d + d') c t 122 | shiftAdd d d' c (tvar n) with c ≤? n 123 | ...| yes p1 = begin 124 | shift d c (tvar (n + d')) ≡⟨ shiftVarEq1 (start c ≤⟨ p1 ⟩ n ≤⟨ m≤m+n n d' ⟩ n + d' □) ⟩ 125 | tvar (n + d' + d) ≡⟨ cong tvar $ begin 126 | n + d' + d ≡⟨ +-assoc n d' d ⟩ 127 | n + (d' + d) ≡⟨ cong (_+_ n) (+-comm d' d) ⟩ 128 | n + (d + d') ∎ ⟩ 129 | tvar (n + (d + d')) ∎ where 130 | open ≡-Reasoning 131 | open ≤-Reasoning renaming (begin_ to start_; _∎ to _□; _≡⟨_⟩_ to _≡⟨_⟩′_) 132 | ...| no p1 = shiftVarEq2 p1 133 | shiftAdd d d' c (tapp t1 t2) = cong₂ tapp (shiftAdd d d' c t1) (shiftAdd d d' c t2) 134 | shiftAdd d d' c (tabs t) = cong tabs (shiftAdd d d' (suc c) t) 135 | 136 | nshiftFold : ∀ n t → nshift n t ≡ shift n 0 t 137 | nshiftFold zero t = shiftZero 0 t 138 | nshiftFold (suc n) t rewrite nshiftFold n t = shiftAdd 1 n 0 t 139 | 140 | data Shifted : ℕ → ℕ → Term → Set where 141 | svar1 : ∀ {d c n} → n < c → Shifted d c (tvar n) 142 | svar2 : ∀ {d c n} → c + d ≤ n → d ≤ n → Shifted d c (tvar n) 143 | sapp : ∀ {d c t1 t2} → Shifted d c t1 → Shifted d c t2 → Shifted d c (tapp t1 t2) 144 | sabs : ∀ {d c t} → Shifted d (suc c) t → Shifted d c (tabs t) 145 | 146 | weakShifted : ∀ {d c t} n → Shifted (d + n) c t → Shifted d (c + n) t 147 | weakShifted {d} {c} zero p rewrite +-comm d 0 | +-comm c 0 = p 148 | weakShifted {d} {c} {tvar n} (suc n') (svar1 p) = 149 | svar1 $ begin suc n ≤⟨ p ⟩ c ≤⟨ m≤m+n c (suc n') ⟩ c + suc n' ∎ 150 | where open ≤-Reasoning 151 | weakShifted {d} {c} {tvar n} (suc n') (svar2 p1 p2) = 152 | svar2 153 | (begin 154 | c + suc n' + d ≡⟨ +-assoc c (suc n') d ⟩ 155 | c + (suc n' + d) ≡⟨ cong (_+_ c) (+-comm (suc n') d) ⟩ 156 | c + (d + suc n') ≤⟨ p1 ⟩ 157 | n ∎) 158 | (begin d ≤⟨ m≤m+n d (suc n') ⟩ d + suc n' ≤⟨ p2 ⟩ n ∎) 159 | where open ≤-Reasoning 160 | weakShifted (suc n) (sapp p1 p2) = sapp (weakShifted (suc n) p1) (weakShifted (suc n) p2) 161 | weakShifted (suc n) (sabs p) = sabs $ weakShifted (suc n) p 162 | 163 | shiftShifted : ∀ d c t → Shifted d c (shift d c t) 164 | shiftShifted d c (tvar n) with c ≤? n 165 | ...| no p = svar1 $ ≱⇒< p 166 | ...| yes p = svar2 (≤-addR d p) (n≤m+n n d) 167 | shiftShifted d c (tapp t1 t2) = sapp (shiftShifted d c t1) (shiftShifted d c t2) 168 | shiftShifted d c (tabs t) = sabs (shiftShifted d (suc c) t) 169 | 170 | shiftShifted' : 171 | ∀ {d c d' c' t} → c' ≤ d + c → Shifted d c t → Shifted d (d' + c) (shift d' c' t) 172 | shiftShifted' {d} {c} {d'} {c'} {tvar n} p s1 = r where 173 | open ≤-Reasoning 174 | r : Shifted d (d' + c) (shift d' c' (tvar n)) 175 | r with c' ≤? n | s1 176 | r | yes p1 | svar1 p2 = svar1 $ 177 | begin suc n + d' ≤⟨ ≤-addR d' p2 ⟩ c + d' ≡⟨ +-comm c d' ⟩ d' + c ∎ 178 | r | yes p1 | svar2 p2 p3 = svar2 179 | (begin 180 | d' + c + d ≡⟨ +-assoc d' c d ⟩ 181 | d' + (c + d) ≤⟨ ≤-addL d' p2 ⟩ 182 | d' + n ≡⟨ +-comm d' n ⟩ 183 | n + d' ∎) 184 | (begin d ≤⟨ p3 ⟩ n ≤⟨ m≤m+n n d' ⟩ n + d' ∎) 185 | r | no p1 | svar1 p2 = svar1 $ 186 | begin suc n ≤⟨ p2 ⟩ c ≤⟨ n≤m+n d' c ⟩ d' + c ∎ 187 | r | no p1 | svar2 p2 p3 = ⊥-elim $ p1 $ 188 | begin c' ≤⟨ p ⟩ d + c ≡⟨ +-comm d c ⟩ c + d ≤⟨ p2 ⟩ n ∎ 189 | shiftShifted' p (sapp s1 s2) = sapp (shiftShifted' p s1) (shiftShifted' p s2) 190 | shiftShifted' {d} {c} {d'} {c'} {tabs t} p (sabs s1) = 191 | sabs $ 192 | subst (λ c → Shifted d c (shift d' (suc c') t)) 193 | (let open ≡-Reasoning in begin 194 | d' + suc c ≡⟨ +-comm d' (suc c) ⟩ 195 | suc (c + d') ≡⟨ cong suc (+-comm c d') ⟩ 196 | suc (d' + c) ∎) 197 | (shiftShifted' 198 | (let open ≤-Reasoning in begin 199 | suc c' ≤⟨ s≤s p ⟩ 200 | suc (d + c) ≡⟨ cong suc (+-comm d c) ⟩ 201 | suc (c + d) ≡⟨ +-comm (suc c) d ⟩ 202 | d + suc c ∎) s1) 203 | 204 | zeroShifted : ∀ c t → Shifted 0 c t 205 | zeroShifted c t = subst (Shifted 0 c) (sym (shiftZero c t)) (shiftShifted 0 c t) 206 | 207 | nshiftShifted : ∀ n t → Shifted n 0 (nshift n t) 208 | nshiftShifted n t rewrite nshiftFold n t = shiftShifted n 0 t 209 | 210 | unshiftShiftSetoff : ∀ {d c d' c'} t → c ≤ c' → c' ≤ d' + d + c → 211 | unshift d' c' (shift (d' + d) c t) ≡ shift d c t 212 | unshiftShiftSetoff {d} {c} {d'} {c'} (tvar n) p1 p2 = r where 213 | open ≡-Reasoning 214 | open ≤-Reasoning renaming (begin_ to start_; _∎ to _□; _≡⟨_⟩_ to _≡⟨_⟩′_) 215 | r : unshift d' c' (shift (d' + d) c (tvar n)) ≡ shift d c (tvar n) 216 | r with c ≤? n 217 | r | yes p3 with c' ≤? (n + (d' + d)) 218 | r | yes p3 | yes p4 = cong tvar $ begin 219 | n + (d' + d) ∸ d' ≡⟨ cong (λ a → n + a ∸ d') (+-comm d' d) ⟩ 220 | n + (d + d') ∸ d' ≡⟨ sym (cong (λ a → a ∸ d') (+-assoc n d d')) ⟩ 221 | n + d + d' ∸ d' ≡⟨ m+n∸n≡m (n + d) d' ⟩ 222 | n + d ∎ 223 | r | yes p3 | no p4 = ⊥-elim $ p4 $ start 224 | c' ≤⟨ p2 ⟩ 225 | d' + d + c ≤⟨ ≤-addL (d' + d) p3 ⟩ 226 | d' + d + n ≡⟨ +-comm (d' + d) n ⟩′ 227 | n + (d' + d) □ 228 | r | no p3 = unshiftVarEq2 $ λ p4 → p3 $ start c ≤⟨ p1 ⟩ c' ≤⟨ p4 ⟩ n □ 229 | unshiftShiftSetoff (tapp t1 t2) p1 p2 = 230 | cong₂ tapp (unshiftShiftSetoff t1 p1 p2) (unshiftShiftSetoff t2 p1 p2) 231 | unshiftShiftSetoff {d} {c} {d'} {c'} (tabs t) p1 p2 = 232 | cong tabs $ unshiftShiftSetoff t (s≤s p1) $ begin 233 | suc c' ≤⟨ s≤s p2 ⟩ 234 | suc d' + d + c ≡⟨ cong suc (+-comm (d' + d) c) ⟩ 235 | suc c + (d' + d) ≡⟨ +-comm (suc c) (d' + d) ⟩ 236 | d' + d + suc c ∎ 237 | where open ≤-Reasoning 238 | 239 | betaShifted : ∀ n t1 t2 → Shifted 1 n (t1 [ n ≔ nshift (suc n) t2 ]) 240 | betaShifted n (tvar n') t2 with n ≟ n' | StrictTotalOrder.compare strictTotalOrder n n' 241 | ...| yes p | _ = weakShifted n $ nshiftShifted (suc n) t2 242 | ...| no p1 | tri< p2 _ _ = 243 | svar2 (subst (λ n → n ≤ n') (+-comm 1 n) p2) (begin 1 ≤⟨ s≤s z≤n ⟩ suc n ≤⟨ p2 ⟩ n' ∎) 244 | where open ≤-Reasoning 245 | ...| no p1 | tri≈ _ p2 _ = ⊥-elim $ p1 p2 246 | ...| no p1 | tri> _ _ p2 = svar1 p2 247 | betaShifted n (tapp t1 t2) t3 = sapp (betaShifted n t1 t3) (betaShifted n t2 t3) 248 | betaShifted n (tabs t1) t2 = sabs $ betaShifted (suc n) t1 t2 249 | 250 | betaShifted' : ∀ n t1 t2 → Shifted 1 n (t1 [ n ≔ shift (suc n) 0 t2 ]) 251 | betaShifted' n t1 t2 rewrite sym (nshiftFold (suc n) t2) = betaShifted n t1 t2 252 | 253 | betaShifted2 : ∀ {d c n t1 t2} → Shifted d (suc n + c) t1 → Shifted d c t2 → 254 | Shifted d (n + c) (unshift 1 n (t1 [ n ≔ shift (suc n) 0 t2 ])) 255 | betaShifted2 {d} {c} {n} {tvar n'} {t2} s1 s2 = r where 256 | open ≤-Reasoning 257 | r : Shifted d (n + c) (unshift 1 n (tvar n' [ n ≔ shift (suc n) 0 t2 ])) 258 | r with n ≟ n' 259 | r | yes p1 rewrite unshiftShiftSetoff {n} {0} {1} t2 z≤n $ 260 | begin n ≤⟨ n≤1+n n ⟩ suc n ≤⟨ s≤s (m≤m+n n 0) ⟩ suc (n + 0) ∎ = 261 | shiftShifted' z≤n s2 262 | r | no p1 with n ≤? n' | s1 263 | r | no p1 | yes p2 | svar1 (s≤s p3) with n' | n 264 | r | no p1 | yes p2 | svar1 (s≤s p3) | 0 | 0 = ⊥-elim $ p1 refl 265 | r | no p1 | yes p2 | svar1 (s≤s p3) | 0 | suc _ = svar1 (s≤s z≤n) 266 | r | no p1 | yes p2 | svar1 (s≤s p3) | suc _ | _ = svar1 p3 267 | r | no p1 | yes p2 | svar2 p3 p4 = svar2 (≤-sub 1 p3) $ begin 268 | d ≤⟨ n≤m+n (n + c) d ⟩ 269 | n + c + d ≤⟨ ≤-sub 1 p3 ⟩ 270 | n' ∸ 1 ∎ 271 | r | no p1 | no p2 | svar1 p3 = svar1 $ begin 272 | suc n' ≤⟨ ≱⇒< p2 ⟩ 273 | n ≤⟨ m≤m+n n c ⟩ 274 | n + c ∎ 275 | r | no p1 | no p2 | svar2 p3 p4 = ⊥-elim $ p2 $ begin 276 | n ≤⟨ n≤m+n (suc (c + d)) n ⟩ 277 | suc c + d + n ≡⟨ cong suc (+-comm (c + d) n) ⟩ 278 | suc n + (c + d) ≡⟨ cong suc (sym (+-assoc n c d)) ⟩ 279 | suc n + c + d ≤⟨ p3 ⟩ 280 | n' ∎ 281 | betaShifted2 (sapp s1 s2) s3 = sapp (betaShifted2 s1 s3) (betaShifted2 s2 s3) 282 | betaShifted2 {n = n} {t2 = t2} (sabs s1) s2 rewrite shiftAdd 1 (suc n) 0 t2 = 283 | sabs (betaShifted2 s1 s2) 284 | 285 | substShiftedCancel : 286 | ∀ {d c n t1 t2} → c ≤ n → n < c + d → Shifted d c t1 → t1 ≡ t1 [ n ≔ t2 ] 287 | substShiftedCancel {d} {c} {n} {tvar n'} {t2} p1 p2 p3 with n ≟ n' | p3 288 | ...| yes p4 | svar1 p5 rewrite p4 = ⊥-elim $ 1+n≰n $ 289 | begin suc n' ≤⟨ p5 ⟩ c ≤⟨ p1 ⟩ n' ∎ 290 | where open ≤-Reasoning 291 | ...| yes p4 | svar2 p5 p6 rewrite p4 = ⊥-elim $ 1+n≰n $ 292 | begin suc c + d ≤⟨ s≤s p5 ⟩ suc n' ≤⟨ p2 ⟩ c + d ∎ 293 | where open ≤-Reasoning 294 | ...| no p4 | _ = refl 295 | substShiftedCancel p1 p2 (sapp p3 p4) = 296 | cong₂ tapp (substShiftedCancel p1 p2 p3) (substShiftedCancel p1 p2 p4) 297 | substShiftedCancel p1 p2 (sabs p3) = 298 | cong tabs $ substShiftedCancel (s≤s p1) (s≤s p2) p3 299 | 300 | shiftShiftSwap : 301 | ∀ d c d' c' t → c ≤ c' → shift d c (shift d' c' t) ≡ shift d' (c' + d) (shift d c t) 302 | shiftShiftSwap d c d' c' (tvar n) p = r where 303 | open ≤-Reasoning 304 | r : shift d c (shift d' c' (tvar n)) ≡ shift d' (c' + d) (shift d c (tvar n)) 305 | r with c ≤? n | c' ≤? n 306 | r | yes p1 | yes p2 with c ≤? (n + d') | (c' + d) ≤? (n + d) 307 | r | yes p1 | yes p2 | yes p3 | yes p4 308 | rewrite +-assoc n d' d | +-assoc n d d' | +-comm d d' = refl 309 | r | yes p1 | yes p2 | _ | no p4 = ⊥-elim $ p4 $ ≤-addR d p2 310 | r | yes p1 | yes p2 | no p3 | _ = ⊥-elim $ p3 $ 311 | begin c ≤⟨ p1 ⟩ n ≤⟨ m≤m+n n d' ⟩ n + d' ∎ 312 | r | yes p1 | no p2 with c ≤? n | (c' + d) ≤? (n + d) 313 | r | yes p1 | no p2 | yes p3 | yes p4 = ⊥-elim $ p2 $ ≤-subR d p4 314 | r | yes p1 | no p2 | yes p3 | no p4 = refl 315 | r | yes p1 | no p2 | no p3 | _ = ⊥-elim $ p3 p1 316 | r | no p1 | yes p2 = ⊥-elim $ p1 $ 317 | begin c ≤⟨ p ⟩ c' ≤⟨ p2 ⟩ n ∎ 318 | r | no p1 | no p2 with c ≤? n | (c' + d) ≤? n 319 | r | no p1 | no p2 | yes p3 | _ = ⊥-elim $ p1 p3 320 | r | no p1 | no p2 | no p3 | yes p4 = ⊥-elim $ p2 $ 321 | begin c' ≤⟨ m≤m+n c' d ⟩ c' + d ≤⟨ p4 ⟩ n ∎ 322 | r | no p1 | no p2 | no p3 | no p4 = refl 323 | shiftShiftSwap d c d' c' (tapp t1 t2) p = 324 | cong₂ tapp (shiftShiftSwap d c d' c' t1 p) (shiftShiftSwap d c d' c' t2 p) 325 | shiftShiftSwap d c d' c' (tabs t) p = 326 | cong tabs (shiftShiftSwap d (suc c) d' (suc c') t (s≤s p)) 327 | 328 | shiftUnshiftSwap : ∀ {d c d' c' t} → c' ≤ c → Shifted d' c' t → 329 | shift d c (unshift d' c' t) ≡ unshift d' c' (shift d (c + d') t) 330 | shiftUnshiftSwap {d} {c} {d'} {c'} {tvar n} p1 p2 = r where 331 | open ≤-Reasoning 332 | r : shift d c (unshift d' c' (tvar n)) ≡ unshift d' c' (shift d (c + d') (tvar n)) 333 | r with c' ≤? n | (c + d') ≤? n 334 | r | yes p3 | yes p4 with c' ≤? (n + d) | c ≤? (n ∸ d') 335 | r | yes p3 | yes p4 | yes p5 | yes p6 = cong tvar $ a-b+c≡a+c-b $ 336 | begin d' ≤⟨ n≤m+n c d' ⟩ c + d' ≤⟨ p4 ⟩ n ∎ 337 | r | yes p3 | yes p4 | yes p5 | no p6 = ⊥-elim $ p6 $ ≤-subR' d' p4 338 | r | yes p3 | yes p4 | no p5 | _ = ⊥-elim $ p5 $ 339 | begin c' ≤⟨ p3 ⟩ n ≤⟨ m≤m+n n d ⟩ n + d ∎ 340 | r | yes p3 | no p4 with c' ≤? n | c ≤? (n ∸ d') | p2 341 | r | yes p3 | no p4 | yes p5 | yes p6 | svar1 p7 = ⊥-elim $ 1+n≰n $ 342 | begin suc n ≤⟨ p7 ⟩ c' ≤⟨ p3 ⟩ n ∎ 343 | r | yes p3 | no p4 | yes p5 | yes p6 | svar2 p7 p8 = ⊥-elim $ p4 $ ≤-addR' d' p8 p6 344 | r | yes p3 | no p4 | yes p5 | no p6 | _ = refl 345 | r | yes p3 | no p4 | no p5 | _ | _ = ⊥-elim $ p5 p3 346 | r | no p3 | yes p4 with c ≤? n | c' ≤? (n + d) 347 | r | no p3 | yes p4 | yes p5 | yes p6 = ⊥-elim $ p3 $ 348 | begin c' ≤⟨ p1 ⟩ c ≤⟨ p5 ⟩ n ∎ 349 | r | no p3 | yes p4 | yes p5 | no p6 = refl 350 | r | no p3 | yes p4 | no p5 | _ = ⊥-elim $ p5 $ 351 | begin c ≤⟨ m≤m+n c d' ⟩ c + d' ≤⟨ p4 ⟩ n ∎ 352 | r | no p3 | no p4 with c' ≤? n | c ≤? n 353 | r | no p3 | no p4 | yes p5 | _ = ⊥-elim $ p3 p5 354 | r | no p3 | no p4 | no p5 | yes p6 = ⊥-elim $ p5 $ 355 | begin c' ≤⟨ p1 ⟩ c ≤⟨ p6 ⟩ n ∎ 356 | r | no p3 | no p4 | no p5 | no p6 = refl 357 | shiftUnshiftSwap p1 (sapp p2 p3) = 358 | cong₂ tapp (shiftUnshiftSwap p1 p2) (shiftUnshiftSwap p1 p3) 359 | shiftUnshiftSwap p1 (sabs p2) = 360 | cong tabs (shiftUnshiftSwap (s≤s p1) p2) 361 | 362 | shiftSubstSwap : ∀ {d c n} → n < c → ∀ t1 t2 → 363 | shift d c (t1 [ n ≔ t2 ]) ≡ shift d c t1 [ n ≔ shift d c t2 ] 364 | shiftSubstSwap {d} {c} {n} p (tvar n') t = r where 365 | open ≤-Reasoning 366 | r : shift d c (tvar n' [ n ≔ t ]) ≡ shift d c (tvar n') [ n ≔ shift d c t ] 367 | r with n ≟ n' | c ≤? n' 368 | r | yes p1 | yes p2 with n ≟ n' + d 369 | r | yes p1 | yes p2 | yes p3 = refl 370 | r | yes p1 | yes p2 | no p3 rewrite p1 = ⊥-elim $ 1+n≰n $ 371 | begin suc n ≤⟨ p ⟩ c ≤⟨ p2 ⟩ n' ≡⟨ sym p1 ⟩ n ∎ 372 | r | yes p1 | no p2 with n ≟ n' 373 | r | yes p1 | no p2 | yes p3 = refl 374 | r | yes p1 | no p2 | no p3 = ⊥-elim $ p3 p1 375 | r | no p1 | yes p2 with c ≤? n' | n ≟ n' + d 376 | r | no p1 | yes p2 | yes p3 | yes p4 = ⊥-elim $ 1+n≰n $ begin 377 | suc n' ≤⟨ m≤m+n (suc n') d ⟩ 378 | suc n' + d ≡⟨ cong suc (sym p4) ⟩ 379 | suc n ≤⟨ p ⟩ 380 | c ≤⟨ p2 ⟩ 381 | n' ∎ 382 | r | no p1 | yes p2 | yes p3 | no p4 = refl 383 | r | no p1 | yes p2 | no p3 | _ = ⊥-elim $ p3 p2 384 | r | no p1 | no p2 with n ≟ n' | c ≤? n' 385 | r | no p1 | no p2 | yes p3 | _ = ⊥-elim $ p1 p3 386 | r | no p1 | no p2 | _ | yes p4 = ⊥-elim $ p2 p4 387 | r | no p1 | no p2 | no p3 | no p4 = refl 388 | shiftSubstSwap p (tapp t1 t2) t3 = 389 | cong₂ tapp (shiftSubstSwap p t1 t3) (shiftSubstSwap p t2 t3) 390 | shiftSubstSwap {d} {c} {n} p (tabs t1) t2 391 | rewrite shiftShiftSwap 1 0 d c t2 z≤n | +-comm c 1 = 392 | cong tabs $ shiftSubstSwap (s≤s p) t1 (shift 1 0 t2) 393 | 394 | shiftSubstSwap' : ∀ {d c n} → c ≤ n → ∀ t1 t2 → 395 | shift d c (t1 [ n ≔ t2 ]) ≡ shift d c t1 [ n + d ≔ shift d c t2 ] 396 | shiftSubstSwap' {d} {c} {n} p1 (tvar n') t = r where 397 | r : shift d c (tvar n' [ n ≔ t ]) ≡ shift d c (tvar n') [ n + d ≔ shift d c t ] 398 | r with n ≟ n' | c ≤? n' 399 | r | yes p2 | yes p3 with n + d ≟ n' + d 400 | r | yes p2 | yes p3 | yes p4 = refl 401 | r | yes p2 | yes p3 | no p4 = ⊥-elim $ p4 $ cong (λ n → n + d) p2 402 | r | yes p2 | no p3 rewrite sym p2 = ⊥-elim $ p3 p1 403 | r | no p2 | yes p3 with c ≤? n' | n + d ≟ n' + d 404 | r | no p2 | yes p3 | yes p4 | yes p5 = ⊥-elim $ p2 $ begin 405 | n ≡⟨ sym (m+n∸n≡m n d) ⟩ 406 | n + d ∸ d ≡⟨ cong (λ n → n ∸ d) p5 ⟩ 407 | n' + d ∸ d ≡⟨ m+n∸n≡m n' d ⟩ 408 | n' ∎ 409 | where open ≡-Reasoning 410 | r | no p2 | yes p3 | yes p4 | no p5 = refl 411 | r | no p2 | yes p3 | no p4 | _ = ⊥-elim $ p4 p3 412 | r | no p2 | no p3 with c ≤? n' | n + d ≟ n' 413 | r | no p2 | no p3 | yes p4 | _ = ⊥-elim $ p3 p4 414 | r | no p2 | no p3 | no p4 | yes p5 = ⊥-elim $ p3 $ begin 415 | c ≤⟨ p1 ⟩ n ≤⟨ m≤m+n n d ⟩ n + d ≡⟨ p5 ⟩ n' ∎ 416 | where open ≤-Reasoning 417 | r | no p2 | no p3 | no p4 | no p5 = refl 418 | shiftSubstSwap' p1 (tapp t1 t2) t3 = 419 | cong₂ tapp (shiftSubstSwap' p1 t1 t3) (shiftSubstSwap' p1 t2 t3) 420 | shiftSubstSwap' {d} {c} {n} p1 (tabs t1) t2 421 | rewrite shiftShiftSwap 1 0 d c t2 z≤n | +-comm c 1 = 422 | cong tabs $ shiftSubstSwap' (s≤s p1) t1 (shift 1 0 t2) 423 | 424 | unshiftShiftSwap : ∀ {d c d' c' t} → c' ≤ c → Shifted d c t → 425 | shift d' c' (unshift d c t) ≡ unshift d (c + d') (shift d' c' t) 426 | unshiftShiftSwap {d} {c} {d'} {c'} {tvar n} p s1 = r where 427 | open ≤-Reasoning 428 | r : shift d' c' (unshift d c (tvar n)) ≡ unshift d (c + d') (shift d' c' (tvar n)) 429 | r with c ≤? n | c' ≤? n 430 | r | yes p1 | yes p2 with c' ≤? (n ∸ d) | (c + d') ≤? (n + d') | s1 431 | r | yes p1 | yes p2 | _ | _ | svar1 p5 = 432 | ⊥-elim $ 1+n≰n $ begin suc n ≤⟨ p5 ⟩ c ≤⟨ p1 ⟩ n ∎ 433 | r | yes p1 | yes p2 | yes p3 | yes p4 | svar2 p5 p6 = 434 | cong tvar $ sym $ a+b∸c≡a∸c+b n d' d p6 435 | r | yes p1 | yes p2 | no p3 | yes p4 | svar2 p5 p6 = ⊥-elim $ p3 $ 436 | begin c' ≤⟨ p ⟩ c ≤⟨ ≤-subR' d p5 ⟩ n ∸ d ∎ 437 | r | yes p1 | yes p2 | _ | no p4 | _ = ⊥-elim $ p4 $ ≤-addR d' p1 438 | r | yes p1 | no p2 = ⊥-elim $ p2 $ begin c' ≤⟨ p ⟩ c ≤⟨ p1 ⟩ n ∎ 439 | r | no p1 | yes p2 with c' ≤? n | (c + d') ≤? (n + d') 440 | r | no p1 | yes p2 | yes _ | yes p3 = ⊥-elim $ p1 $ ≤-subR d' p3 441 | r | no p1 | yes p2 | yes _ | no p3 = refl 442 | r | no p1 | yes p2 | no p3 | _ = ⊥-elim $ p3 p2 443 | r | no p1 | no p2 with c' ≤? n | (c + d') ≤? n 444 | r | no p1 | no p2 | _ | yes p4 = 445 | ⊥-elim $ p1 $ begin c ≤⟨ m≤m+n c d' ⟩ c + d' ≤⟨ p4 ⟩ n ∎ 446 | r | no p1 | no p2 | yes p3 | no p4 = ⊥-elim $ p2 p3 447 | r | no p1 | no p2 | no p3 | no p4 = refl 448 | unshiftShiftSwap p (sapp s1 s2) = 449 | cong₂ tapp (unshiftShiftSwap p s1) (unshiftShiftSwap p s2) 450 | unshiftShiftSwap p (sabs s1) = cong tabs (unshiftShiftSwap (s≤s p) s1) 451 | 452 | unshiftSubstSwap : 453 | ∀ {c n} t1 t2 → c ≤ n → Shifted 1 c t1 → 454 | unshift 1 c (t1 [ suc n ≔ shift (suc c) 0 t2 ]) ≡ unshift 1 c t1 [ n ≔ shift c 0 t2 ] 455 | unshiftSubstSwap {c} {n} (tvar n') t2 p1 p2 = r where 456 | open ≤-Reasoning 457 | r : unshift 1 c (tvar n' [ suc n ≔ shift (suc c) 0 t2 ]) ≡ 458 | unshift 1 c (tvar n') [ n ≔ shift c 0 t2 ] 459 | r with suc n ≟ n' | c ≤? n' 460 | r | yes p3 | yes p4 with n ≟ n' ∸ 1 461 | r | yes p3 | yes p4 | yes p5 = unshiftShiftSetoff t2 z≤n $ 462 | begin c ≡⟨ +-comm 0 c ⟩ c + 0 ≤⟨ n≤1+n (c + 0) ⟩ suc (c + 0) ∎ 463 | r | yes p3 | yes p4 | no p5 = ⊥-elim $ p5 $ ≡-subL' 1 p3 464 | r | yes p3 | no p4 with n ≟ n' | p2 465 | r | yes p3 | no p4 | yes p5 | _ = ⊥-elim $ 1+n≰n $ 466 | begin suc n ≡⟨ p3 ⟩ n' ≡⟨ sym p5 ⟩ n ∎ 467 | r | yes p3 | no p4 | no p5 | svar1 p6 = ⊥-elim $ 1+n≰n $ 468 | begin suc n ≡⟨ p3 ⟩ n' ≤⟨ n≤1+n n' ⟩ suc n' ≤⟨ p6 ⟩ c ≤⟨ p1 ⟩ n ∎ 469 | r | yes p3 | no p4 | no p5 | svar2 p6 p7 = ⊥-elim $ p4 $ 470 | begin c ≤⟨ m≤m+n c 1 ⟩ c + 1 ≤⟨ p6 ⟩ n' ∎ 471 | r | no p3 | yes p4 with c ≤? n' | n ≟ n' ∸ 1 | p2 472 | r | no p3 | yes p4 | yes p5 | yes p6 | svar1 p7 = ⊥-elim $ 1+n≰n $ 473 | begin suc n' ≤⟨ p7 ⟩ c ≤⟨ p4 ⟩ n' ∎ 474 | r | no p3 | yes p4 | yes p5 | yes p6 | svar2 p7 p8 = ⊥-elim $ p3 $ ≡-addL' p8 p6 475 | r | no p3 | yes p4 | yes p5 | no p6 | _ = refl 476 | r | no p3 | yes p4 | no p5 | _ | _ = ⊥-elim $ p5 p4 477 | r | no p3 | no p4 with c ≤? n' | n ≟ n' 478 | r | no p3 | no p4 | yes p5 | _ = ⊥-elim $ p4 p5 479 | r | no p3 | no p4 | no p5 | yes p6 = ⊥-elim $ p5 $ 480 | begin c ≤⟨ p1 ⟩ n ≡⟨ p6 ⟩ n' ∎ 481 | r | no p3 | no p4 | no p5 | no p6 = refl 482 | unshiftSubstSwap (tapp t1 t2) t3 p1 (sapp p2 p3) = 483 | cong₂ tapp (unshiftSubstSwap t1 t3 p1 p2) (unshiftSubstSwap t2 t3 p1 p3) 484 | unshiftSubstSwap {c} {n} (tabs t1) t2 p1 (sabs p2) 485 | rewrite shiftAdd 1 (suc c) 0 t2 | shiftAdd 1 c 0 t2 = 486 | cong tabs $ unshiftSubstSwap t1 t2 (s≤s p1) p2 487 | 488 | unshiftSubstSwap' : 489 | ∀ {n} t1 t2 → Shifted 1 0 t1 → 490 | unshift 1 0 (t1 [ suc n ≔ shift 1 0 t2 ]) ≡ unshift 1 0 t1 [ n ≔ t2 ] 491 | unshiftSubstSwap' {n} t1 t2 p 492 | rewrite cong (λ t2 → unshift 1 0 t1 [ n ≔ t2 ]) (shiftZero 0 t2) = 493 | unshiftSubstSwap t1 t2 z≤n p 494 | 495 | unshiftSubstSwap2 : 496 | ∀ {d c n t1 t2} → n < c → Shifted d c t1 → Shifted d c t2 → 497 | unshift d c (t1 [ n ≔ t2 ]) ≡ unshift d c t1 [ n ≔ unshift d c t2 ] 498 | unshiftSubstSwap2 {d} {c} {n} {tvar n'} {t2} p s1 s2 = r where 499 | open ≤-Reasoning 500 | r : unshift d c (tvar n' [ n ≔ t2 ]) ≡ unshift d c (tvar n') [ n ≔ unshift d c t2 ] 501 | r with n ≟ n' | c ≤? n' 502 | r | yes p1 | yes p2 = ⊥-elim $ 1+n≰n $ 503 | begin suc n ≤⟨ p ⟩ c ≤⟨ p2 ⟩ n' ≡⟨ sym p1 ⟩ n ∎ 504 | r | yes p1 | no p2 with n ≟ n' 505 | r | yes p1 | no p2 | yes p3 = refl 506 | r | yes p1 | no p2 | no p3 = ⊥-elim $ p3 p1 507 | r | no p1 | yes p2 with c ≤? n' | n ≟ n' ∸ d 508 | r | no p1 | yes p2 | yes p3 | yes p4 with s1 509 | r | no p1 | yes p2 | yes p3 | yes p4 | svar1 p5 = 510 | ⊥-elim $ 1+n≰n $ begin suc n' ≤⟨ p5 ⟩ c ≤⟨ p3 ⟩ n' ∎ 511 | r | no p1 | yes p2 | yes p3 | yes p4 | svar2 p5 p6 = ⊥-elim $ 1+n≰n $ begin 512 | suc n' ≡⟨ cong suc $ sym $ ≡-addL' p6 p4 ⟩ 513 | suc d + n ≡⟨ cong suc (+-comm d n) ⟩ 514 | suc n + d ≤⟨ ≤-addR d p ⟩ 515 | c + d ≤⟨ p5 ⟩ 516 | n' ∎ 517 | r | no p1 | yes p2 | yes p3 | no p4 = refl 518 | r | no p1 | yes p2 | no p3 | _ = ⊥-elim $ p3 p2 519 | r | no p1 | no p2 with c ≤? n' | n ≟ n' 520 | r | no p1 | no p2 | yes p3 | _ = ⊥-elim $ p2 p3 521 | r | no p1 | no p2 | _ | yes p4 = ⊥-elim $ p1 p4 522 | r | no p1 | no p2 | no p3 | no p4 = refl 523 | unshiftSubstSwap2 p (sapp s1 s2) s3 = 524 | cong₂ tapp (unshiftSubstSwap2 p s1 s3) (unshiftSubstSwap2 p s2 s3) 525 | unshiftSubstSwap2 {d} {c} {n} {tabs t1} {t2} p (sabs s1) s2 526 | rewrite unshiftShiftSwap {d} {c} {1} {0} z≤n s2 | +-comm c 1 = 527 | cong tabs $ unshiftSubstSwap2 (s≤s p) s1 $ shiftShifted' z≤n s2 528 | 529 | substSubstSwap : 530 | ∀ n m t1 t2 t3 → 531 | t1 [ m ≔ shift (suc m) 0 t2 ] [ suc m + n ≔ shift (suc m) 0 t3 ] ≡ 532 | t1 [ suc m + n ≔ shift (suc m) 0 t3 ] [ m ≔ shift (suc m) 0 (t2 [ n ≔ t3 ])] 533 | substSubstSwap n m (tvar n') t2 t3 = r where 534 | r : tvar n' [ m ≔ shift (suc m) 0 t2 ] [ suc m + n ≔ shift (suc m) 0 t3 ] ≡ 535 | tvar n' [ suc m + n ≔ shift (suc m) 0 t3 ] [ m ≔ shift (suc m) 0 (t2 [ n ≔ t3 ])] 536 | r with m ≟ n' | suc (m + n) ≟ n' 537 | r | yes p1 | yes p2 rewrite p1 = ⊥-elim $ m≢1+m+n n' $ sym p2 538 | r | yes p1 | no p2 with m ≟ n' 539 | r | yes p1 | no p2 | yes p3 rewrite +-comm (suc m) n = 540 | sym $ shiftSubstSwap' z≤n t2 t3 541 | r | yes p1 | no p2 | no p3 = ⊥-elim $ p3 p1 542 | r | no p1 | yes p2 with suc (m + n) ≟ n' 543 | r | no p1 | yes p2 | yes p3 = 544 | substShiftedCancel z≤n (≤′⇒≤ ≤′-refl) (shiftShifted (suc m) 0 t3) 545 | r | no p1 | yes p2 | no p3 = ⊥-elim $ p3 p2 546 | r | no p1 | no p2 with m ≟ n' | suc (m + n) ≟ n' 547 | r | no p1 | no p2 | yes p3 | _ = ⊥-elim $ p1 p3 548 | r | no p1 | no p2 | _ | yes p4 = ⊥-elim $ p2 p4 549 | r | no p1 | no p2 | no p3 | no p4 = refl 550 | substSubstSwap n m (tapp t1l t1r) t2 t3 = 551 | cong₂ tapp (substSubstSwap n m t1l t2 t3) (substSubstSwap n m t1r t2 t3) 552 | substSubstSwap n m (tabs t1) t2 t3 rewrite 553 | shiftAdd 1 (suc m) 0 t2 | 554 | shiftAdd 1 (suc m) 0 t3 | 555 | shiftAdd 1 (suc m) 0 (t2 [ n ≔ t3 ]) = 556 | cong tabs $ substSubstSwap n (suc m) t1 t2 t3 557 | 558 | unshiftUnshiftSwap : 559 | ∀ {d c d' c' t} → c' ≤ c → Shifted d' c' t → Shifted d c (unshift d' c' t) → 560 | unshift d c (unshift d' c' t) ≡ unshift d' c' (unshift d (c + d') t) 561 | unshiftUnshiftSwap {d} {c} {d'} {c'} {t = tvar n} p1 p2 p3 = r where 562 | open ≤-Reasoning 563 | r : unshift d c (unshift d' c' (tvar n)) ≡ 564 | unshift d' c' (unshift d (c + d') (tvar n)) 565 | r with c' ≤? n | (c + d') ≤? n 566 | r | yes p4 | yes p5 with c ≤? (n ∸ d') | c' ≤? (n ∸ d) 567 | r | yes p4 | yes p5 | yes p6 | yes p7 = cong tvar $ a∸b∸c≡a∸c∸b n d' d 568 | r | yes p4 | yes p5 | yes p6 | no p7 with subst (Shifted d c) (unshiftVarEq1 p4) p3 569 | r | yes p4 | yes p5 | yes p6 | no p7 | svar1 p8 = 570 | ⊥-elim $ 1+n≰n $ begin suc (n ∸ d') ≤⟨ p8 ⟩ c ≤⟨ p6 ⟩ n ∸ d' ∎ 571 | r | yes p4 | yes p5 | yes p6 | no p7 | svar2 p8 p9 = ⊥-elim $ p7 $ begin 572 | c' ≤⟨ p1 ⟩ 573 | c ≤⟨ ≤-subR' d p8 ⟩ 574 | n ∸ d' ∸ d ≡⟨ a∸b∸c≡a∸c∸b n d' d ⟩ 575 | n ∸ d ∸ d' ≤⟨ n∸m≤n d' (n ∸ d) ⟩ 576 | n ∸ d ∎ 577 | r | yes p4 | yes p5 | no p6 | _ = ⊥-elim $ p6 $ ≤-subR' d' p5 578 | r | yes p4 | no p5 with c' ≤? n | c ≤? (n ∸ d') 579 | r | yes p4 | no p5 | yes p6 | yes p7 with p2 580 | r | yes p4 | no p5 | yes p6 | yes p7 | svar1 p8 = 581 | ⊥-elim $ 1+n≰n $ begin suc n ≤⟨ p8 ⟩ c' ≤⟨ p4 ⟩ n ∎ 582 | r | yes p4 | no p5 | yes p6 | yes p7 | svar2 p8 p9 = ⊥-elim $ p5 $ ≤-addR' d' p9 p7 583 | r | yes p4 | no p5 | yes p6 | no p7 = refl 584 | r | yes p4 | no p5 | no p6 | _ = ⊥-elim $ p6 p4 585 | r | no p4 | yes p5 with c ≤? n | c' ≤? (n ∸ d) 586 | r | no p4 | yes p5 | yes p6 | yes p7 = 587 | ⊥-elim $ p4 $ begin c' ≤⟨ p1 ⟩ c ≤⟨ p6 ⟩ n ∎ 588 | r | no p4 | yes p5 | yes p6 | no p7 = refl 589 | r | no p4 | yes p5 | no p6 | yes p7 = 590 | ⊥-elim $ p6 $ begin c ≤⟨ m≤m+n c d' ⟩ c + d' ≤⟨ p5 ⟩ n ∎ 591 | r | no p4 | yes p5 | no p6 | no p7 with p2 592 | r | no p4 | yes p5 | no p6 | no p7 | svar1 p8 = ⊥-elim $ 1+n≰n $ 593 | begin suc c + d' ≤⟨ s≤s p5 ⟩ suc n ≤⟨ p8 ⟩ c' ≤⟨ p1 ⟩ c ≤⟨ m≤m+n c d' ⟩ c + d' ∎ 594 | r | no p4 | yes p5 | no p6 | no p7 | svar2 p8 p9 = 595 | ⊥-elim $ ⊥-elim $ p4 $ begin c' ≤⟨ m≤m+n c' d' ⟩ c' + d' ≤⟨ p8 ⟩ n ∎ 596 | r | no p4 | no p5 with c' ≤? n | c ≤? n 597 | r | no p4 | no p5 | yes p6 | _ = ⊥-elim $ p4 p6 598 | r | no p4 | no p5 | no p6 | yes p7 = 599 | ⊥-elim $ p4 $ begin c' ≤⟨ p1 ⟩ c ≤⟨ p7 ⟩ n ∎ 600 | r | no p4 | no p5 | no p6 | no p7 = refl 601 | unshiftUnshiftSwap p1 (sapp p2 p3) (sapp p4 p5) = 602 | cong₂ tapp (unshiftUnshiftSwap p1 p2 p4) (unshiftUnshiftSwap p1 p3 p5) 603 | unshiftUnshiftSwap p1 (sabs p2) (sabs p3) = cong tabs (unshiftUnshiftSwap (s≤s p1) p2 p3) 604 | 605 | -------------------------------------------------------------------------------- /coq/CL.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import ssreflect ssrfun ssrbool ssrnat seq. 2 | From LCAC Require Import Relations_ext. 3 | Require FunInd. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Printing Implicit Defensive. 8 | 9 | (* Definition 2.1: Combinatory logic terms, or CL-terms *) 10 | 11 | Inductive clterm : Set := 12 | | clvar of nat 13 | | clapp of clterm & clterm 14 | | clatoms | clatomk | clatomi. 15 | 16 | Lemma eq_clterm_dec (t1 t2 : clterm) : {t1 = t2}+{t1 <> t2}. 17 | Proof. do ?decide equality. Qed. 18 | 19 | Local Infix "@" := clapp (at level 20, left associativity). 20 | 21 | Check ((clatoms @ (clatomk @ clatoms)) @ clatomk). 22 | Check ((clatoms @ (clatomk @ clvar 0)) @ ((clatoms @ clatomk) @ clatomk)). 23 | 24 | (* Definition 2.3: Length of CL-term *) 25 | 26 | Fixpoint cl_length (t : clterm) : nat := 27 | match t with 28 | | clapp t1 t2 => cl_length t1 + cl_length t2 29 | | _ => 1 30 | end. 31 | 32 | (* Definition 2.4: Binary relation "occurs in" on CL-terms *) 33 | 34 | Inductive cl_occurs : relation clterm := 35 | | cl_occurs_refl t : cl_occurs t t 36 | | cl_occurs_left t1 t2 t3 : cl_occurs t1 t2 -> cl_occurs t1 (t2 @ t3) 37 | | cl_occurs_right t1 t2 t3 : cl_occurs t1 t3 -> cl_occurs t1 (t2 @ t3). 38 | 39 | (* Definition 2.6, Exercise 2.8: Substitution *) 40 | 41 | Fixpoint subst_lookup (l : seq (nat * clterm)) (v : nat) : option clterm := 42 | match l with 43 | | nil => None 44 | | (v', t) :: l' => if eqn v v' then Some t else subst_lookup l' v 45 | end. 46 | 47 | Fixpoint substitute (l : list (nat * clterm)) (t : clterm) : clterm := 48 | match t with 49 | | clvar v => 50 | match subst_lookup l v with 51 | | None => clvar v 52 | | Some t' => t' 53 | end 54 | | clapp tl tr => clapp (substitute l tl) (substitute l tr) 55 | | bc => bc 56 | end. 57 | 58 | (* Example 2.7 *) 59 | 60 | Eval compute in substitute 61 | [:: (1, clatoms @ clatomk)] (clvar 0 @ clvar 1 @ clvar 1). 62 | Eval compute in substitute [:: (1, clatoms @ clatomk); (0, clatomk @ clatomi)] 63 | (clvar 0 @ clvar 1 @ clvar 1). 64 | 65 | (* Definition 2.9: Weak reduction *) 66 | 67 | Inductive cl_weakred : relation clterm := 68 | | weakred_left t1 t2 t3 : cl_weakred t1 t2 -> cl_weakred (t1 @ t3) (t2 @ t3) 69 | | weakred_right t1 t2 t3 : cl_weakred t1 t2 -> cl_weakred (t3 @ t1) (t3 @ t2) 70 | | weakred_s t1 t2 t3 : 71 | cl_weakred (clatoms @ t1 @ t2 @ t3) (t1 @ t3 @ (t2 @ t3)) 72 | | weakred_k t1 t2 : cl_weakred (clatomk @ t1 @ t2) t1 73 | | weakred_i t : cl_weakred (clatomi @ t) t. 74 | 75 | Notation cl_weakred_rtc := (clos_refl_trans_1n clterm cl_weakred). 76 | 77 | Infix "->1w" := cl_weakred (at level 50, no associativity). 78 | Infix "->w" := cl_weakred_rtc (at level 50, no associativity). 79 | 80 | Lemma weakred_rtc_left t1 t2 t3 : t1 ->w t2 -> t1 @ t3 ->w t2 @ t3. 81 | Proof. 82 | elim => // {t1 t2} t1 t2 t2' H H0 H1. 83 | by apply rt1n_trans with (t2 @ t3) => //; constructor. 84 | Qed. 85 | 86 | Lemma weakred_rtc_right t1 t2 t3 : t1 ->w t2 -> t3 @ t1 ->w t3 @ t2. 87 | Proof. 88 | elim => // {t1 t2} t1 t1' t2 H H0 H1. 89 | by apply rt1n_trans with (t3 @ t1') => //; constructor. 90 | Qed. 91 | 92 | Hint Resolve weakred_rtc_left weakred_rtc_right. 93 | 94 | Lemma weakred_rtc_app t1 t1' t2 t2' : 95 | t1 ->w t1' -> t2 ->w t2' -> (t1 @ t2) ->w (t1' @ t2'). 96 | Proof. move => H H0; apply rtc_trans with (t1 @ t2'); auto. Qed. 97 | 98 | Hint Resolve weakred_rtc_app. 99 | 100 | (* Definition 2.10: Weak normal form *) 101 | 102 | Definition cl_weaknf (t : clterm) : Prop := ~(exists t' : clterm, t ->1w t'). 103 | Definition cl_weaknf_of (t1 t2 : clterm) : Prop := t2 ->w t1 /\ cl_weaknf t1. 104 | 105 | (* Example 2.11 *) 106 | 107 | Definition cl_comb_b : clterm := clatoms @ (clatomk @ clatoms) @ clatomk. 108 | 109 | Example example_2_11 t1 t2 t3 : cl_comb_b @ t1 @ t2 @ t3 ->w t1 @ (t2 @ t3). 110 | Proof. 111 | rewrite /cl_comb_b. 112 | eapply rt1n_trans; first apply weakred_left, weakred_left, weakred_s. 113 | eapply rt1n_trans; 114 | first apply weakred_left, weakred_left, weakred_left, weakred_k. 115 | eapply rt1n_trans; first apply weakred_s. 116 | apply rtc_step; apply weakred_left, weakred_k. 117 | Qed. 118 | 119 | (* Example 2.12 *) 120 | 121 | Definition cl_comb_c : clterm := 122 | clatoms @ (cl_comb_b @ cl_comb_b @ clatoms) @ (clatomk @ clatomk). 123 | 124 | Example example_2_12 t1 t2 t3 : cl_comb_c @ t1 @ t2 @ t3 ->w t1 @ t3 @ t2. 125 | Proof. 126 | rewrite /cl_comb_c. 127 | eapply rt1n_trans; first apply weakred_left, weakred_left, weakred_s. 128 | eapply rtc_trans; first (do 3 apply weakred_rtc_left; apply example_2_11). 129 | eapply rtc_trans; first apply weakred_rtc_left, example_2_11. 130 | eapply rt1n_trans; first apply weakred_s. 131 | apply weakred_rtc_right. 132 | eapply rt1n_trans; first apply weakred_left, weakred_left, weakred_k. 133 | apply rtc_step; apply weakred_k. 134 | Qed. 135 | 136 | (* Exercise 2.13 *) 137 | 138 | 139 | 140 | (* Lemma 2.14: Substitution lemma for ->w *) 141 | 142 | Lemma substlemma_a' t1 t2 v : 143 | t1 ->1w t2 -> cl_occurs (clvar v) t2 -> cl_occurs (clvar v) t1. 144 | Proof. 145 | move => H H0. 146 | induction H. 147 | inversion H0. 148 | apply cl_occurs_left, IHcl_weakred, H3. 149 | apply cl_occurs_right, H3. 150 | inversion H0. 151 | apply cl_occurs_left, H3. 152 | apply cl_occurs_right, IHcl_weakred, H3. 153 | inversion H0; inversion H2. 154 | apply cl_occurs_left, cl_occurs_left, cl_occurs_right, H6. 155 | apply cl_occurs_right, H6. 156 | apply cl_occurs_left, cl_occurs_right, H6. 157 | apply cl_occurs_right, H6. 158 | apply cl_occurs_left, cl_occurs_right, H0. 159 | apply cl_occurs_right, H0. 160 | Qed. 161 | 162 | Lemma substlemma_a t1 t2 v : 163 | t1 ->w t2 -> cl_occurs (clvar v) t2 -> cl_occurs (clvar v) t1. 164 | Proof. elim => // x y z H H0 H1 H2; eapply substlemma_a'; eauto. Qed. 165 | 166 | Lemma substlemma_b t1 t2 t3 v : 167 | t1 ->w t2 -> substitute [:: (v, t1)] t3 ->w substitute [:: (v, t2)] t3. 168 | Proof. 169 | move => H; elim: t3 => //=. 170 | - by move => n; case: ifP. 171 | - auto. 172 | Qed. 173 | 174 | Lemma substlemma_c t1 t2 l : t1 ->1w t2 -> substitute l t1 ->1w substitute l t2. 175 | Proof. by elim => /=; constructor. Qed. 176 | 177 | (* Theorem 2.15: Church-Rosser theorem for ->w *) 178 | 179 | Inductive cl_parred : relation clterm := 180 | | parred_app t1 t1' t2 t2' : 181 | cl_parred t1 t1' -> cl_parred t2 t2' -> cl_parred (t1 @ t2) (t1' @ t2') 182 | | parred_s t1 t1' t2 t2' t3 t3' t3'' : 183 | cl_parred t1 t1' -> cl_parred t2 t2' -> 184 | cl_parred t3 t3' -> cl_parred t3 t3'' -> 185 | cl_parred (clatoms @ t1 @ t2 @ t3) (t1' @ t3' @ (t2' @ t3'')) 186 | | parred_k t1 t1' t2 : cl_parred t1 t1' -> cl_parred (clatomk @ t1 @ t2) t1' 187 | | parred_i t1 t1' : cl_parred t1 t1' -> cl_parred (clatomi @ t1) t1' 188 | | parred_var n : cl_parred (clvar n) (clvar n) 189 | | parred_atoms : cl_parred clatoms clatoms 190 | | parred_atomk : cl_parred clatomk clatomk 191 | | parred_atomi : cl_parred clatomi clatomi. 192 | 193 | Infix "->p" := cl_parred (at level 50, no associativity). 194 | 195 | Lemma cl_parred_refl t : t ->p t. 196 | Proof. by elim: t; constructor. Qed. 197 | 198 | Hint Resolve cl_parred_refl. 199 | 200 | Lemma cl_weakred_in_parred : inclusion cl_weakred cl_parred. 201 | Proof. by apply cl_weakred_ind; constructor. Qed. 202 | 203 | Lemma cl_parred_in_weakred_rtc : inclusion cl_parred cl_weakred_rtc. 204 | Proof. 205 | apply cl_parred_ind => //=; auto. 206 | - move => t1 t1' t2 t2' t3 t3' t3'' H H0 H1 H2 H3 H4 H5 H6. 207 | apply rt1n_trans with (t1 @ t3 @ (t2 @ t3)). 208 | + constructor. 209 | + auto. 210 | - move => t1 t1' t2 H H0; apply rt1n_trans with t1 => //=; constructor. 211 | - move => t1 t1' H H0; apply rt1n_trans with t1 => //=; constructor. 212 | Qed. 213 | 214 | Function cl_parred_all (t : clterm) : clterm := 215 | match t with 216 | | clatoms @ t1 @ t2 @ t3 => 217 | (cl_parred_all t1 @ cl_parred_all t3) @ 218 | (cl_parred_all t2 @ cl_parred_all t3) 219 | | clatomk @ t1 @ t2 => cl_parred_all t1 220 | | clatomi @ t1 => cl_parred_all t1 221 | | t1 @ t2 => cl_parred_all t1 @ cl_parred_all t2 222 | | _ => t 223 | end. 224 | 225 | Lemma cl_parred_all_lemma t1 t2 : t1 ->p t2 -> t2 ->p cl_parred_all t1. 226 | Proof. 227 | move => H; induction H; try by do ?constructor. 228 | destruct t1; try by constructor. 229 | - destruct t1_1; try by constructor. 230 | + destruct t1_1_1; try by constructor. 231 | inversion H; inversion H3; inversion H8 ;subst. 232 | inversion IHcl_parred1; subst; simpl in *. 233 | by inversion H6; apply parred_s. 234 | + inversion H; inversion H3; subst. 235 | by inversion IHcl_parred1; apply parred_k. 236 | - inversion H; subst. 237 | by inversion IHcl_parred1; apply parred_i. 238 | Qed. 239 | 240 | Theorem cl_parred_confluent : confluent cl_parred. 241 | Proof. by move => t1; exists (cl_parred_all t1); apply cl_parred_all_lemma. Qed. 242 | 243 | Theorem cl_weakred_confluent : confluent cl_weakred_rtc. 244 | Proof. 245 | apply (rtc_confluent' 246 | cl_weakred_in_parred cl_parred_in_weakred_rtc cl_parred_confluent). 247 | Qed. 248 | 249 | (* Corollary 2.15.1: Uniqueness of nf *) 250 | 251 | Corollary cl_weaknf_uniqueness t1 t2 t3 : 252 | cl_weaknf_of t2 t1 -> cl_weaknf_of t3 t1 -> t2 = t3. 253 | Proof. 254 | case => H H0; case => H1 H2. 255 | case: (cl_weakred_confluent H H1) => [t4 H3 H4]. 256 | inversion H3. 257 | - inversion H4. 258 | + auto. 259 | + apply False_ind, H2; eauto. 260 | - apply False_ind, H0; eauto. 261 | Qed. 262 | 263 | (* Exercise 2.16 *) 264 | 265 | Lemma exercise_2_16 t : clatoms @ clatomk @ clatomk @ t ->w t. 266 | Proof. 267 | eapply rt1n_trans. 268 | - apply weakred_s. 269 | - apply rtc_step, weakred_k. 270 | Qed. 271 | 272 | (* Exercise 2.17 *) 273 | -------------------------------------------------------------------------------- /coq/LN/Untyped.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. 2 | From LCAC Require Import Relations_ext. 3 | 4 | Set Implicit Arguments. 5 | 6 | Notation nopts n T := (iter n option T). 7 | 8 | Inductive term (V : Type) : Type := 9 | | var of V 10 | | app of term V & term V 11 | | abs of term (option V). 12 | 13 | Notation term' n V := (term (nopts n V)). 14 | 15 | Fixpoint tmap (V1 V2 : Type) (f : V1 -> V2) (t : term V1) : term V2 := 16 | match t with 17 | | var a => var (f a) 18 | | app tl tr => app (tmap f tl) (tmap f tr) 19 | | abs t => abs (tmap (Option.map f) t) 20 | end. 21 | 22 | Fixpoint somenth (T : Type) (n : nat) (a : T) : nopts n T := 23 | match n with 24 | | 0 => a 25 | | n.+1 => Some (somenth n a) 26 | end. 27 | 28 | Fixpoint substitute_var 29 | (V : Type) (n : nat) (t : term V) : nopts n (option V) -> term' n V := 30 | match n with 31 | | 0 => fun v => match v with 32 | | None => t 33 | | Some v => var v 34 | end 35 | | n.+1 => fun v => match v with 36 | | None => var None 37 | | Some v => tmap some (substitute_var n t v) 38 | end 39 | end. 40 | 41 | Fixpoint substitute 42 | (V : Type) (n : nat) (t1 : term V) (t2 : term' n (option V)) : term' n V := 43 | match t2 with 44 | | var v => substitute_var n t1 v 45 | | app t2l t2r => app (substitute n t1 t2l) (substitute n t1 t2r) 46 | | abs t2' => abs (substitute n.+1 t1 t2') 47 | end. 48 | 49 | Fixpoint substitute_seq_var2 50 | (V : Type) (ts : seq (term V)) : nopts (size ts) V -> term V := 51 | match ts with 52 | | [::] => fun v => var v 53 | | t :: ts => fun v => match v with 54 | | None => t 55 | | Some v => substitute_seq_var2 ts v 56 | end 57 | end. 58 | 59 | Fixpoint substitute_seq_var1 60 | (V : Type) (n : nat) (ts : seq (term V)) : 61 | nopts (n + size ts) V -> term' (n + 0) V := 62 | match n with 63 | | 0 => fun v => substitute_seq_var2 ts v 64 | | n.+1 => fun v => match v with 65 | | None => var None 66 | | Some v => tmap some (substitute_seq_var1 n ts v) 67 | end 68 | end. 69 | 70 | Fixpoint substitute_seq 71 | (V : Type) (n : nat) (ts : seq (term V)) (t : term' (n + size ts) V) : 72 | term' (n + 0) V := 73 | match t with 74 | | var v => substitute_seq_var1 n ts v 75 | | app tl tr => app (substitute_seq n ts tl) (substitute_seq n ts tr) 76 | | abs t => abs (substitute_seq n.+1 ts t) 77 | end. 78 | 79 | Lemma substitute_seq_nil V n (t : term' (n + 0) V) : 80 | substitute_seq n [::] t = t. 81 | Proof. 82 | move: n t; fix IH 2 => n; case => //=. 83 | - elim: n => // {IH} n IH //=; case => // v. 84 | by rewrite (IH v). 85 | - move => tl tr; f_equal; apply IH. 86 | - move => t; f_equal; apply IH. 87 | Qed. 88 | 89 | Reserved Notation "t ->1b t'" (at level 70, no associativity). 90 | Reserved Notation "t ->bp t'" (at level 70, no associativity). 91 | 92 | Inductive betared1 : forall V, relation (term V) := 93 | | betared1beta V (t1 : term (option V)) (t2 : term V) : 94 | app (abs t1) t2 ->1b substitute 0 t2 t1 95 | | betared1appl V (t1 t1' t2 : term V) : 96 | t1 ->1b t1' -> app t1 t2 ->1b app t1' t2 97 | | betared1appr V (t1 t2 t2' : term V) : 98 | t2 ->1b t2' -> app t1 t2 ->1b app t1 t2' 99 | | betared1abs V (t t' : term (option V)) : t ->1b t' -> abs t ->1b abs t' 100 | where "t ->1b t'" := (betared1 t t'). 101 | 102 | Inductive parred : forall V, relation (term V) := 103 | | parredvar V (x : V) : var x ->bp var x 104 | | parredapp V (t1 t1' t2 t2' : term V) : 105 | t1 ->bp t1' -> t2 ->bp t2' -> app t1 t2 ->bp app t1' t2' 106 | | parredabs V (t t' : term (option V)) : t ->bp t' -> abs t ->bp abs t' 107 | | parredbeta V (t1 t1' : term (option V)) (t2 t2' : term V) : 108 | t1 ->bp t1' -> t2 ->bp t2' -> app (abs t1) t2 ->bp substitute 0 t2' t1' 109 | where "t ->bp t'" := (parred t t'). 110 | -------------------------------------------------------------------------------- /coq/Makefile: -------------------------------------------------------------------------------- 1 | # KNOWNTARGETS will not be passed along to CoqMakefile 2 | KNOWNTARGETS := Makefile.coq 3 | # KNOWNFILES will not get implicit targets from the final rule, and so 4 | # depending on them won't invoke the submake 5 | # Warning: These files get declared as PHONY, so any targets depending 6 | # on them always get rebuilt 7 | KNOWNFILES := Makefile _CoqProject 8 | 9 | .DEFAULT_GOAL := invoke-coqmakefile 10 | 11 | Makefile.coq: Makefile _CoqProject 12 | $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq 13 | 14 | invoke-coqmakefile: Makefile.coq 15 | $(MAKE) --no-print-directory -f Makefile.coq $(filter-out $(KNOWNTARGETS),$(MAKECMDGOALS)) 16 | 17 | .PHONY: invoke-coqmakefile $(KNOWNFILES) 18 | 19 | #################################################################### 20 | ## Your targets here ## 21 | #################################################################### 22 | 23 | # This should be the last rule, to handle any targets not declared above 24 | %: invoke-coqmakefile 25 | @true 26 | -------------------------------------------------------------------------------- /coq/Origin/Untyped.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import ssreflect ssrnat. 2 | From LCAC Require Import Relations_ext. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | (* Definition 1.1: Lambda-terms *) 9 | 10 | Inductive lcterm : Set := 11 | | lcbvar of nat 12 | | lcubvar of nat 13 | | lclam of lcterm 14 | | lcapp of lcterm & lcterm. 15 | 16 | Lemma eq_lcterm_dec (t1 t2 : lcterm) : {t1 = t2}+{t1 <> t2}. 17 | Proof. do !decide equality. Qed. 18 | 19 | Infix "@" := lcapp (at level 20, left associativity). 20 | 21 | (* Example 1.2 *) 22 | 23 | Check (lclam (lcapp (lcbvar 0) (lcubvar 0))). 24 | Check (lclam (lcapp (lcbvar 0) (lcubvar 0))). 25 | Check (lcapp (lcubvar 0) (lclam (lclam (lcbvar 0)))). 26 | Check (lcapp (lclam (lcbvar 0)) (lclam (lcapp (lcbvar 0) (lcubvar 0)))). 27 | Check (lclam (lcapp (lcubvar 0) (lcbvar 0))). 28 | 29 | (* Notation 1.3 *) 30 | 31 | Fixpoint lclams (n : nat) (t : lcterm) := 32 | match n with 33 | | 0 => t 34 | | n.+1 => lclams n (lclam t) 35 | end. 36 | 37 | (* Definition 1.6: Length of lambda-term *) 38 | 39 | Fixpoint lc_length (t : lcterm) : nat := 40 | match t with 41 | | lclam t' => (lc_length t').+1 42 | | lcapp t1 t2 => lc_length t1 + lc_length t2 43 | | _ => 1 44 | end. 45 | 46 | (* Definition 1.7: Binary relation "occurs in" on lambda-terms *) 47 | 48 | Inductive lc_occurs : relation lcterm := 49 | | lc_occurs_refl t : lc_occurs t t 50 | | lc_occurs_left t1 t2 t3 : lc_occurs t1 t2 -> lc_occurs t1 (t2 @ t3) 51 | | lc_occurs_right t1 t2 t3 : lc_occurs t1 t3 -> lc_occurs t1 (t2 @ t3) 52 | | lc_occurs_lam t1 t2 : lc_occurs t1 t2 -> lc_occurs t1 (lclam t2). 53 | 54 | (* Definition 1.12: Substitution *) 55 | -------------------------------------------------------------------------------- /coq/_CoqProject: -------------------------------------------------------------------------------- 1 | CL.v 2 | LN/Untyped.v 3 | Origin/Untyped.v 4 | deBruijn/F.v 5 | deBruijn/STLC.v 6 | deBruijn/Untyped.v 7 | lib/Relations_ext.v 8 | lib/seq_ext.v 9 | lib/seq_ext_base.v 10 | lib/ssrnat_ext.v 11 | 12 | -R . LCAC 13 | -------------------------------------------------------------------------------- /coq/deBruijn/F.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. 2 | From LCAC Require Import Relations_ext seq_ext_base ssrnat_ext seq_ext. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | Inductive typ := tyvar of nat | tyfun of typ & typ | tyabs of typ. 9 | 10 | Inductive term 11 | := var of nat (* variable *) 12 | | app of term & term (* value application *) 13 | | abs of typ & term (* value abstraction *) 14 | | uapp of term & typ (* universal application *) 15 | | uabs of term. (* universal abstraction *) 16 | 17 | Coercion tyvar : nat >-> typ. 18 | Coercion var : nat >-> term. 19 | 20 | Notation "ty ->t ty'" := (tyfun ty ty') (at level 55, right associativity). 21 | Notation "t @ t'" := (app t t') (at level 65, left associativity). 22 | Notation "t @' ty" := (uapp t ty) (at level 65, left associativity). 23 | 24 | Fixpoint eqtyp t1 t2 := 25 | match t1, t2 with 26 | | tyvar n, tyvar m => n == m 27 | | t1l ->t t1r, t2l ->t t2r => eqtyp t1l t2l && eqtyp t1r t2r 28 | | tyabs t1, tyabs t2 => eqtyp t1 t2 29 | | _, _ => false 30 | end. 31 | 32 | Lemma eqtypP : Equality.axiom eqtyp. 33 | Proof. 34 | move => t1 t2; apply: (iffP idP) => [| <-]; last by elim: t1 => //= t ->. 35 | by elim: t1 t2 => [n | t1l IHt1l t1r IHt1r | t1 IHt1] 36 | [// m /eqP -> | //= t2l t2r /andP [] /IHt1l -> /IHt1r -> | // t2 /IHt1 ->]. 37 | Defined. 38 | 39 | Canonical typ_eqMixin := EqMixin eqtypP. 40 | Canonical typ_eqType := Eval hnf in EqType typ typ_eqMixin. 41 | 42 | Fixpoint eqterm t1 t2 := 43 | match t1, t2 with 44 | | var n, var m => n == m 45 | | t1l @ t1r, t2l @ t2r => eqterm t1l t2l && eqterm t1r t2r 46 | | abs ty1 t1, abs ty2 t2 => (ty1 == ty2) && eqterm t1 t2 47 | | t1 @' ty1, t2 @' ty2 => eqterm t1 t2 && (ty1 == ty2) 48 | | uabs t1, uabs t2 => eqterm t1 t2 49 | | _, _ => false 50 | end. 51 | 52 | Lemma eqtermP : Equality.axiom eqterm. 53 | Proof. 54 | move => t1 t2; apply: (iffP idP) => [| <-]; 55 | last by elim: t1 => //= [t -> t' -> | ty t -> | t ->] *; rewrite ?eqxx. 56 | by elim: t1 t2 => 57 | [n | t1l IHt1l t1r IHt1r | ty1 t1 IHt1 | t1 IHt1 ty1 | t1 IHt1] 58 | [// m /eqP -> | //= t2l t2r /andP [] /IHt1l -> /IHt1r -> | 59 | //= ty2 t2 /andP [] /eqP -> /IHt1 -> | //= t2 ty2 /andP [] /IHt1 -> /eqP -> | 60 | //= t2 /IHt1 ->]. 61 | Defined. 62 | 63 | Canonical term_eqMixin := EqMixin eqtermP. 64 | Canonical term_eqType := Eval hnf in EqType term term_eqMixin. 65 | 66 | Fixpoint shift_typ d c t := 67 | match t with 68 | | tyvar n => tyvar (if c <= n then n + d else n) 69 | | tl ->t tr => shift_typ d c tl ->t shift_typ d c tr 70 | | tyabs t => tyabs (shift_typ d c.+1 t) 71 | end. 72 | 73 | Notation subst_typv ts m n := 74 | (shift_typ n 0 (nth (tyvar (m - n - size ts)) ts (m - n))) (only parsing). 75 | 76 | Fixpoint subst_typ n ts t := 77 | match t with 78 | | tyvar m => if n <= m then subst_typv ts m n else m 79 | | tl ->t tr => subst_typ n ts tl ->t subst_typ n ts tr 80 | | tyabs t => tyabs (subst_typ n.+1 ts t) 81 | end. 82 | 83 | Fixpoint typemap (f : nat -> typ -> typ) n t := 84 | match t with 85 | | var m => var m 86 | | tl @ tr => typemap f n tl @ typemap f n tr 87 | | abs ty t => abs (f n ty) (typemap f n t) 88 | | t @' ty => typemap f n t @' f n ty 89 | | uabs t => uabs (typemap f n.+1 t) 90 | end. 91 | 92 | Fixpoint shift_term d c t := 93 | match t with 94 | | var n => var (if c <= n then n + d else n) 95 | | tl @ tr => shift_term d c tl @ shift_term d c tr 96 | | abs ty t => abs ty (shift_term d c.+1 t) 97 | | t @' ty => shift_term d c t @' ty 98 | | uabs t => uabs (shift_term d c t) 99 | end. 100 | 101 | Notation subst_termv ts m n n' := 102 | (typemap (shift_typ n') 0 103 | (shift_term n 0 (nth (var (m - n - size ts)) ts (m - n)))) (only parsing). 104 | 105 | Fixpoint subst_term n n' ts t := 106 | match t with 107 | | var m => if n <= m then subst_termv ts m n n' else m 108 | | tl @ tr => subst_term n n' ts tl @ subst_term n n' ts tr 109 | | abs ty t => abs ty (subst_term n.+1 n' ts t) 110 | | t @' ty => subst_term n n' ts t @' ty 111 | | uabs t => uabs (subst_term n n'.+1 ts t) 112 | end. 113 | 114 | Reserved Notation "t ->r1 t'" (at level 70, no associativity). 115 | 116 | Inductive reduction1 : relation term := 117 | | red1fst ty t1 t2 : abs ty t1 @ t2 ->r1 subst_term 0 0 [:: t2] t1 118 | | red1snd t ty : uabs t @' ty ->r1 typemap (subst_typ^~ [:: ty]) 0 t 119 | | red1appl t1 t1' t2 : t1 ->r1 t1' -> t1 @ t2 ->r1 t1' @ t2 120 | | red1appr t1 t2 t2' : t2 ->r1 t2' -> t1 @ t2 ->r1 t1 @ t2' 121 | | red1abs ty t t' : t ->r1 t' -> abs ty t ->r1 abs ty t' 122 | | red1uapp t t' ty : t ->r1 t' -> t @' ty ->r1 t' @' ty 123 | | red1uabs t t' : t ->r1 t' -> uabs t ->r1 uabs t' 124 | where "t ->r1 t'" := (reduction1 t t'). 125 | 126 | Notation reduction := [* reduction1]. 127 | Infix "->r" := reduction (at level 70, no associativity). 128 | 129 | Hint Constructors reduction1. 130 | 131 | Fixpoint typing_rec (ctx : context typ) (t : term) : option typ := 132 | match t with 133 | | var n => nth None ctx n 134 | | tl @ tr => 135 | if typing_rec ctx tl is Some (tyl ->t tyr) 136 | then (if typing_rec ctx tr == Some tyl then Some tyr else None) 137 | else None 138 | | abs ty t => omap (tyfun ty) (typing_rec (Some ty :: ctx) t) 139 | | t @' ty => 140 | if typing_rec ctx t is Some (tyabs ty') 141 | then Some (subst_typ 0 [:: ty] ty') 142 | else None 143 | | uabs t => omap tyabs (typing_rec (ctxmap (shift_typ 1 0) ctx) t) 144 | end. 145 | 146 | Definition typing := nosimpl typing_rec. 147 | 148 | Notation "ctx \|- t \: ty" := (Some ty == typing ctx t) 149 | (at level 69, no associativity). 150 | Notation "ctx \|- t \: ty" := (Some (ty : typ) == typing ctx t) 151 | (at level 69, no associativity, only parsing). 152 | 153 | Lemma typing_varE ctx (n : nat) (ty : typ) : 154 | ctx \|- n \: ty = ctxindex ctx n ty. 155 | Proof. by rewrite /typing /=. Qed. 156 | 157 | Lemma typing_appP ctx t1 t2 ty : 158 | reflect (exists2 tyl, ctx \|- t1 \: tyl ->t ty & ctx \|- t2 \: tyl) 159 | (ctx \|- t1 @ t2 \: ty). 160 | Proof. 161 | apply: (iffP idP); rewrite /typing /=. 162 | - by move: (typing_rec ctx t1) => [] // [] // tyl tyr; 163 | case: ifP => // /eqP -> /eqP [] ->; exists tyl. 164 | - by case => tyl /eqP <- /eqP <-; rewrite eqxx. 165 | Qed. 166 | 167 | Lemma typing_absP ctx t tyl ty : 168 | reflect (exists2 tyr, ty = tyl ->t tyr & Some tyl :: ctx \|- t \: tyr) 169 | (ctx \|- abs tyl t \: ty). 170 | Proof. 171 | apply: (iffP idP); rewrite /typing /=. 172 | - by case: typing_rec => //= tyr /eqP [] ->; exists tyr. 173 | - by case => tyr ->; case: typing_rec => // tyr' /eqP [] <-. 174 | Qed. 175 | 176 | Lemma typing_absE ctx t tyl tyr : 177 | ctx \|- abs tyl t \: tyl ->t tyr = Some tyl :: ctx \|- t \: tyr. 178 | Proof. 179 | by rewrite /typing /=; case: typing_rec => //= tyr'; 180 | rewrite /eq_op /= /eq_op /= -/eq_op eqxx. 181 | Qed. 182 | 183 | Lemma typing_absE' ctx t tyl tyl' tyr : 184 | ctx \|- abs tyl t \: tyl' ->t tyr = 185 | (tyl == tyl') && (Some tyl :: ctx \|- t \: tyr). 186 | Proof. 187 | rewrite /typing /=; case: typing_rec => //=. 188 | - by move => tyr'; rewrite !/eq_op /= /eq_op /= -/eq_op eq_sym. 189 | - by rewrite andbF. 190 | Qed. 191 | 192 | Lemma typing_uappP ctx t ty1 ty2 : 193 | reflect 194 | (exists2 ty3, ty2 = subst_typ 0 [:: ty1] ty3 & ctx \|- t \: tyabs ty3) 195 | (ctx \|- t @' ty1 \: ty2). 196 | Proof. 197 | apply: (iffP idP); rewrite /typing /=. 198 | - by move: (typing_rec ctx t) => [] // [] // ty3 /eqP [->]; exists ty3. 199 | - by case => t3 -> /eqP <-. 200 | Qed. 201 | 202 | Lemma typing_uabsP ctx t ty : 203 | reflect 204 | (exists2 ty', ty = tyabs ty' & ctxmap (shift_typ 1 0) ctx \|- t \: ty') 205 | (ctx \|- uabs t \: ty). 206 | Proof. 207 | apply: (iffP idP); rewrite /typing /=. 208 | - by case: typing_rec => // ty' /eqP [] ->; exists ty'. 209 | - by case => ty' -> /eqP; case: typing_rec => // ty'' <-. 210 | Qed. 211 | 212 | Lemma typing_uabsE ctx t ty : 213 | ctx \|- uabs t \: tyabs ty = ctxmap (shift_typ 1 0) ctx \|- t \: ty. 214 | Proof. by rewrite /typing /=; case: typing_rec. Qed. 215 | 216 | Notation SN := (Acc (fun x y => reduction1 y x)). 217 | 218 | Definition neutral (t : term) : bool := 219 | match t with abs _ _ => false | uabs _ => false | _ => true end. 220 | 221 | Lemma inj_shift_typ d c ty ty' : 222 | (shift_typ d c ty == shift_typ d c ty') = (ty == ty'). 223 | Proof. 224 | elim: ty ty' d c => 225 | [n | tyl IHl tyr IHr | ty IH] [m | tyl' tyr' | ty'] d c //=; 226 | try by rewrite !eqE /= -!eqE ?IHl ?IHr ?IH. 227 | apply/eqP; elimif; case: ifP => H0; apply/eqP; move: H0; 228 | rewrite !eqE /= ?(eqn_add2l, eqn_add2r) //; 229 | (move => -> // || move/eqP => H; apply/eqP; ssromega). 230 | Qed. 231 | 232 | Lemma shift_zero_ty n t : shift_typ 0 n t = t. 233 | Proof. by elim: t n; congruence' => m n; rewrite addn0 if_same. Qed. 234 | 235 | Lemma shift_add_ty d c d' c' t : 236 | c <= c' <= c + d -> 237 | shift_typ d' c' (shift_typ d c t) = shift_typ (d' + d) c t. 238 | Proof. case/andP; do 2 elimleq; elim: t c; congruence' => *; elimif_omega. Qed. 239 | 240 | Lemma shift_shift_distr_ty d c d' c' t : 241 | c' <= c -> 242 | shift_typ d' c' (shift_typ d c t) = shift_typ d (d' + c) (shift_typ d' c' t). 243 | Proof. elimleq; elim: t c'; congruence' => *; elimif_omega. Qed. 244 | 245 | Lemma shift_subst_distr_ty n d c ts t : 246 | c <= n -> 247 | shift_typ d c (subst_typ n ts t) = subst_typ (d + n) ts (shift_typ d c t). 248 | Proof. 249 | elimleq; elim: t n c; congruence' => v n c; 250 | elimif; rewrite shift_add_ty; elimif_omega. 251 | Qed. 252 | 253 | Lemma subst_shift_distr_ty n d c ts t : 254 | n <= c -> 255 | shift_typ d c (subst_typ n ts t) = 256 | subst_typ n (map (shift_typ d (c - n)) ts) (shift_typ d (size ts + c) t). 257 | Proof. 258 | elimleq; elim: t n; congruence' => v n; rewrite size_map; elimif. 259 | - rewrite !nth_default ?size_map /=; elimif_omega. 260 | - rewrite -shift_shift_distr_ty // nth_map' /=. 261 | congr shift_typ; apply nth_equal; rewrite size_map; elimif_omega. 262 | Qed. 263 | 264 | Lemma subst_nil_ty n t : subst_typ n [::] t = t. 265 | Proof. elim: t n; congruence' => v n; rewrite nth_nil /=; elimif_omega. Qed. 266 | 267 | Lemma subst_shift_cancel_ty1 n d c ts t : 268 | c <= n -> 269 | subst_typ n ts (shift_typ d c t) = 270 | subst_typ n (drop (c + d - n) ts) 271 | (shift_typ (d - minn (c + d - n) (size ts)) c t). 272 | Proof. 273 | elimleq; elim: t c; congruence' => v c; rewrite size_drop nth_drop. 274 | case (leqP' d n); elimif_omega; elimleq. 275 | case (leqP' d.+1 (size ts)) => H0; elimif_omega. 276 | rewrite !nth_default //; ssromega. 277 | Qed. 278 | 279 | Lemma subst_shift_cancel_ty2 n d c ts t : 280 | c <= n -> size ts + n <= d + c -> 281 | subst_typ n ts (shift_typ d c t) = shift_typ (d - size ts) c t. 282 | Proof. 283 | move => H H0; rewrite subst_shift_cancel_ty1 ?drop_oversize ?subst_nil_ty; 284 | first congr shift_typ; ssromega. 285 | Qed. 286 | 287 | Lemma subst_shift_cancel_ty3 n d c ts t : 288 | n <= c <= n + size ts -> 289 | subst_typ n ts (shift_typ d c t) = 290 | subst_typ n (take (c - n) ts ++ drop (c + d - n) ts) 291 | (shift_typ (c + d - (n + size ts)) c t). 292 | Proof. 293 | case/andP; elimleq => H; elim: t n; congruence' => v n. 294 | rewrite size_cat size_take size_drop nth_cat nth_drop 295 | size_take (minn_idPl H); elimif_omega; congr shift_typ. 296 | - case (leqP' (c + d) (size ts)) => H0; 297 | [rewrite addn0 | rewrite !nth_default]; elimif_omega. 298 | - rewrite nth_take; elimif_omega. 299 | Qed. 300 | 301 | Lemma subst_shift_cancel_ty4 n c ts ts' t t' : 302 | c <= size ts -> 303 | subst_typ n (insert ts' ts t' c) (shift_typ (size ts') (n + c) t) = 304 | subst_typ n ts t. 305 | Proof. 306 | rewrite subst_shift_cancel_ty3 ?size_insert; last ssromega. 307 | rewrite -addnA subnDA !addKn subnDA addnK take_insert drop_insert /leq. 308 | move/eqP => -> /=; 309 | rewrite cats0 cat_take_drop (_ : _ - _ = 0) ?shift_zero_ty //; ssromega. 310 | Qed. 311 | 312 | Lemma subst_app_ty n xs ys t : 313 | subst_typ n xs (subst_typ (size xs + n) ys t) = subst_typ n (xs ++ ys) t. 314 | Proof. 315 | elim: t n; congruence' => v n; rewrite size_cat nth_cat; elimif_omega. 316 | rewrite subst_shift_cancel_ty2; elimif_omega. 317 | Qed. 318 | 319 | Lemma subst_subst_distr_ty n m xs ys t : 320 | m <= n -> 321 | subst_typ n xs (subst_typ m ys t) = 322 | subst_typ m (map (subst_typ (n - m) xs) ys) (subst_typ (size ys + n) xs t). 323 | Proof. 324 | elimleq; elim: t m; congruence' => v m; elimif_omega. 325 | - rewrite (@subst_shift_cancel_ty2 m) ?size_map 1?nth_default //=; 326 | elimif_omega. 327 | - rewrite size_map -shift_subst_distr_ty //= nth_map' /=. 328 | congr shift_typ; apply nth_equal; rewrite size_map; elimif_omega. 329 | Qed. 330 | 331 | Lemma subst_subst_compose_ty n m xs ys t : 332 | m <= size xs -> 333 | subst_typ n xs (subst_typ (n + m) ys t) = 334 | subst_typ n (insert (map (subst_typ 0 (drop m xs)) ys) xs 0 m) t. 335 | Proof. 336 | move => H; elim: t n; congruence' => v n. 337 | rewrite size_insert nth_insert size_map; elimif_omega. 338 | - by rewrite (maxn_idPr H) nth_default /= 1?addnCA ?leq_addr //= addKn addnC. 339 | - rewrite (nth_map (tyvar (v - size ys))) // shift_subst_distr_ty // 340 | addn0 subst_shift_cancel_ty1 //=; elimif_omega. 341 | Qed. 342 | 343 | Lemma typemap_compose f g n m t : 344 | typemap f n (typemap g m t) = 345 | typemap (fun i ty => f (i + n) (g (i + m) ty)) 0 t. 346 | Proof. 347 | elim: t n m => //=. 348 | - by move => tl IHtl tr IHtr n m; rewrite IHtl IHtr. 349 | - by move => ty t IHt n m; rewrite IHt. 350 | - by move => t IHt ty n m; rewrite IHt. 351 | - move => t IHt n m; rewrite {}IHt; congr uabs; elim: t 0 => //; congruence'. 352 | Qed. 353 | 354 | Lemma typemap_eq' f g n m t : 355 | (forall o t, f (o + n) t = g (o + m) t) -> typemap f n t = typemap g m t. 356 | Proof. 357 | move => H; elim: t n m H => //=. 358 | - move => tl IHtl tr IHtr n m H; congr app; auto; apply (H 0). 359 | - by move => ty t IHt n m H; congr abs; auto; rewrite -(add0n n) -(add0n m). 360 | - by move => t IHt ty n m H; congr uapp; auto; apply (H 0). 361 | - by move => t IHt n m H; congr uabs; apply IHt => o ty; rewrite -!addSnnS. 362 | Qed. 363 | 364 | Lemma typemap_eq f g n t : 365 | (forall n t, f n t = g n t) -> typemap f n t = typemap g n t. 366 | Proof. move => H; apply typemap_eq' => {t} o; apply H. Qed. 367 | 368 | Lemma shifttyp_zero c t : typemap (shift_typ 0) c t = t. 369 | Proof. elim: t c => /=; move: shift_zero_ty; congruence. Qed. 370 | 371 | Lemma shifttyp_add d c d' c' t : 372 | c <= c' <= c + d -> 373 | typemap (shift_typ d') c' (typemap (shift_typ d) c t) = 374 | typemap (shift_typ (d' + d)) c t. 375 | Proof. 376 | rewrite typemap_compose => H; apply typemap_eq' => {t} n t. 377 | by rewrite addn0 shift_add_ty // -addnA !leq_add2l. 378 | Qed. 379 | 380 | Lemma shifttyp_shifttyp_distr d c d' c' t : 381 | c' <= c -> 382 | typemap (shift_typ d') c' (typemap (shift_typ d) c t) = 383 | typemap (shift_typ d) (d' + c) (typemap (shift_typ d') c' t). 384 | Proof. 385 | rewrite !typemap_compose => H; apply typemap_eq => {t} n t. 386 | by rewrite shift_shift_distr_ty ?leq_add2l // addnCA. 387 | Qed. 388 | 389 | Lemma shifttyp_substtyp_distr n d c ts t : 390 | c <= n -> 391 | typemap (shift_typ d) c (typemap (subst_typ^~ ts) n t) = 392 | typemap (subst_typ^~ ts) (d + n) (typemap (shift_typ d) c t). 393 | Proof. 394 | rewrite !typemap_compose => H; apply typemap_eq => {t} n' t. 395 | by rewrite shift_subst_distr_ty ?leq_add2l // addnCA. 396 | Qed. 397 | 398 | Lemma substtyp_shifttyp_distr n d c ts t : 399 | n <= c -> 400 | typemap (shift_typ d) c (typemap (subst_typ^~ ts) n t) = 401 | typemap (subst_typ^~ (map (shift_typ d (c - n)) ts)) n 402 | (typemap (shift_typ d) (size ts + c) t). 403 | Proof. 404 | rewrite !typemap_compose => H; apply typemap_eq => {t} n' t. 405 | by rewrite subst_shift_distr_ty ?leq_add2l // subnDl addnCA addnA. 406 | Qed. 407 | 408 | Lemma substtyp_substtyp_distr n m xs ys t : 409 | m <= n -> 410 | typemap (subst_typ^~ xs) n (typemap (subst_typ^~ ys) m t) = 411 | typemap (subst_typ^~ (map (subst_typ (n - m) xs) ys)) m 412 | (typemap (subst_typ^~ xs) (size ys + n) t). 413 | Proof. 414 | rewrite !typemap_compose => H; apply typemap_eq => {t} n' t. 415 | by rewrite subst_subst_distr_ty ?leq_add2l // subnDl addnCA addnA. 416 | Qed. 417 | 418 | Lemma shift_typemap_distr f d c n t : 419 | shift_term d c (typemap f n t) = typemap f n (shift_term d c t). 420 | Proof. elim: t c n => /=; congruence. Qed. 421 | 422 | Lemma subst_shifttyp_distr n m d c ts t : 423 | c + d <= m -> 424 | subst_term n m ts (typemap (shift_typ d) c t) = 425 | typemap (shift_typ d) c (subst_term n (m - d) ts t). 426 | Proof. 427 | elimleq; elim: t n c; congruence' => v n c; elimif_omega. 428 | rewrite -!shift_typemap_distr shifttyp_add; elimif_omega. 429 | Qed. 430 | 431 | Lemma shifttyp_subst_distr d c n m ts t : 432 | m <= c -> 433 | typemap (shift_typ d) c (subst_term n m ts t) = 434 | subst_term n m (map (typemap (shift_typ d) (c - m)) ts) 435 | (typemap (shift_typ d) c t). 436 | Proof. 437 | by elimleq; elim: t n m; congruence' => v n m; elimif_omega; 438 | rewrite size_map -!shift_typemap_distr -shifttyp_shifttyp_distr // nth_map'. 439 | Qed. 440 | 441 | Lemma subst_substtyp_distr n m m' tys ts t : 442 | m' <= m -> 443 | subst_term n m ts (typemap (subst_typ^~ tys) m' t) = 444 | typemap (subst_typ^~ tys) m' (subst_term n (size tys + m) ts t). 445 | Proof. 446 | elimleq; elim: t n m'; congruence' => v n m'; elimif_omega. 447 | rewrite -!shift_typemap_distr typemap_compose. 448 | congr shift_term; apply typemap_eq => {v n ts} n t. 449 | rewrite subst_shift_cancel_ty2; first congr shift_typ; ssromega. 450 | Qed. 451 | 452 | Lemma substtyp_subst_distr n m m' tys ts t : 453 | m <= m' -> 454 | typemap (subst_typ^~ tys) m' (subst_term n m ts t) = 455 | subst_term n m (map (typemap (subst_typ^~ tys) (m' - m)) ts) 456 | (typemap (subst_typ^~ tys) m' t). 457 | Proof. 458 | elimleq; elim: t n m; congruence' => v n m; elimif_omega. 459 | by rewrite size_map -!shift_typemap_distr 460 | -shifttyp_substtyp_distr // (nth_map' (typemap _ _)). 461 | Qed. 462 | 463 | Lemma shift_zero n t : shift_term 0 n t = t. 464 | Proof. by elim: t n => /=; congruence' => m n; rewrite addn0 if_same. Qed. 465 | 466 | Lemma shift_add d c d' c' t : 467 | c <= c' <= c + d -> 468 | shift_term d' c' (shift_term d c t) = shift_term (d' + d) c t. 469 | Proof. case/andP; do 2 elimleq; elim: t c; congruence' => *; elimif_omega. Qed. 470 | 471 | Lemma shift_shift_distr d c d' c' t : 472 | c' <= c -> 473 | shift_term d' c' (shift_term d c t) = 474 | shift_term d (d' + c) (shift_term d' c' t). 475 | Proof. elimleq; elim: t c'; congruence' => v c'; elimif_omega. Qed. 476 | 477 | Lemma subst_shift_distr n m d c ts t : 478 | n <= c -> 479 | shift_term d c (subst_term n m ts t) = 480 | subst_term n m (map (shift_term d (c - n)) ts) (shift_term d (size ts + c) t). 481 | Proof. 482 | elimleq; elim: t n m; congruence' => v n m; rewrite size_map; elimif_omega. 483 | - rewrite !nth_default ?size_map /=; elimif_omega. 484 | - rewrite shift_typemap_distr -shift_shift_distr // nth_map' /=. 485 | congr (typemap _ _ (shift_term _ _ _)); 486 | apply nth_equal; rewrite size_map; elimif_omega. 487 | Qed. 488 | 489 | Lemma shift_subst_distr n m d c ts t : 490 | c <= n -> 491 | shift_term d c (subst_term n m ts t) = 492 | subst_term (d + n) m ts (shift_term d c t). 493 | Proof. 494 | elimleq; elim: t m c; congruence' => v m c; elimif_omega. 495 | rewrite shift_typemap_distr shift_add; elimif_omega. 496 | Qed. 497 | 498 | Lemma subst_shift_cancel n m d c ts t : 499 | c <= n -> size ts + n <= d + c -> 500 | subst_term n m ts (shift_term d c t) = shift_term (d - size ts) c t. 501 | Proof. 502 | do 2 elimleq; elim: t m c; congruence' => v m c; elimif_omega. 503 | rewrite nth_default /=; first congr var; ssromega. 504 | Qed. 505 | 506 | Lemma subst_subst_distr n n' m m' xs ys t : 507 | n' <= n -> m' <= m -> 508 | subst_term n m xs (subst_term n' m' ys t) = 509 | subst_term n' m' (map (subst_term (n - n') (m - m') xs) ys) 510 | (subst_term (size ys + n) m xs t). 511 | Proof. 512 | do 2 elimleq; elim: t n' m'; congruence' => v n' m'; elimif_omega. 513 | - rewrite nth_default /=; elimif_omega. 514 | rewrite -!shift_typemap_distr (@subst_shift_cancel n') // size_map; 515 | elimif_omega. 516 | - rewrite -!shift_typemap_distr -shift_subst_distr // 517 | subst_shifttyp_distr; elimif_omega. 518 | rewrite nth_map' /=; congr (shift_term _ _ (typemap _ _ _)); 519 | apply nth_equal; rewrite !size_map; elimif_omega. 520 | Qed. 521 | 522 | Lemma subst_app n m xs ys t : 523 | subst_term n m xs (subst_term (size xs + n) m ys t) = 524 | subst_term n m (xs ++ ys) t. 525 | Proof. 526 | elim: t n m; congruence' => v n m; rewrite nth_cat size_cat; elimif_omega. 527 | rewrite -!shift_typemap_distr subst_shift_cancel; elimif_omega. 528 | Qed. 529 | 530 | Lemma subst_nil n m t : subst_term n m [::] t = t. 531 | Proof. 532 | elim: t n m; congruence' => v n m; rewrite nth_nil /= -fun_if; elimif_omega. 533 | Qed. 534 | 535 | Lemma subst_shifttyp_app n m m' ts t : 536 | subst_term n m (map (typemap (shift_typ m') 0) ts) t = 537 | subst_term n (m' + m) ts t. 538 | Proof. 539 | elim: t n m; congruence' => v n m; rewrite size_map; elimif_omega. 540 | rewrite -!shift_typemap_distr !(nth_map' (typemap _ _)) /=. 541 | congr (shift_term _ _ (nth _ _ _)); elim: ts => //= t ts ->; congr cons. 542 | rewrite typemap_compose; apply typemap_eq => {n t ts} n t. 543 | rewrite addn0 shift_add_ty; elimif_omega. 544 | Qed. 545 | 546 | Lemma redappl t1 t1' t2 : t1 ->r t1' -> t1 @ t2 ->r t1' @ t2. 547 | Proof. apply (rtc_map' (fun x y => @red1appl x y t2)). Qed. 548 | 549 | Lemma redappr t1 t2 t2' : t2 ->r t2' -> t1 @ t2 ->r t1 @ t2'. 550 | Proof. apply (rtc_map' (@red1appr t1)). Qed. 551 | 552 | Lemma redapp t1 t1' t2 t2' : 553 | t1 ->r t1' -> t2 ->r t2' -> t1 @ t2 ->r t1' @ t2'. 554 | Proof. 555 | by move => H H0; eapply rtc_trans with (t1' @ t2); 556 | [apply redappl | apply redappr]. 557 | Qed. 558 | 559 | Lemma redabs ty t t' : t ->r t' -> abs ty t ->r abs ty t'. 560 | Proof. apply (rtc_map' (red1abs ty)). Qed. 561 | 562 | Lemma reduapp t t' ty : t ->r t' -> t @' ty ->r t' @' ty. 563 | Proof. apply (rtc_map' (fun x y => @red1uapp x y ty)). Qed. 564 | 565 | Lemma reduabs t t' : t ->r t' -> uabs t ->r uabs t'. 566 | Proof. apply (rtc_map' red1uabs). Qed. 567 | 568 | Hint Resolve redappl redappr redapp redabs reduapp reduabs. 569 | 570 | Lemma shifttyp_reduction1 t t' d c : 571 | t ->r1 t' -> typemap (shift_typ d) c t ->r1 typemap (shift_typ d) c t'. 572 | Proof. 573 | move => H; elim/reduction1_ind: t t' / H d c => /=; try by constructor. 574 | - by move => ty t1 t2 d c; rewrite shifttyp_subst_distr //= subn0. 575 | - by move => t ty d c; rewrite substtyp_shifttyp_distr //= subn0. 576 | Qed. 577 | 578 | Lemma shift_reduction1 t t' d c : 579 | t ->r1 t' -> shift_term d c t ->r1 shift_term d c t'. 580 | Proof. 581 | move => H; elim/reduction1_ind: t t' / H d c => /=; try by constructor. 582 | - by move => ty t1 t2 d c; rewrite subst_shift_distr //= subn0. 583 | - by move => t ty d c; rewrite shift_typemap_distr. 584 | Qed. 585 | 586 | Lemma substtyp_reduction1 t t' tys n : 587 | t ->r1 t' -> 588 | typemap (subst_typ^~ tys) n t ->r1 typemap (subst_typ^~ tys) n t'. 589 | Proof. 590 | move => H; elim/reduction1_ind: t t' / H tys n => /=; try by constructor. 591 | - by move => ty t1 t2 tys n; 592 | rewrite substtyp_subst_distr // subn0; constructor. 593 | - by move => t ty tys n; rewrite substtyp_substtyp_distr //= subn0. 594 | Qed. 595 | 596 | Lemma subst_reduction1 t t' n m ts : 597 | t ->r1 t' -> subst_term n m ts t ->r1 subst_term n m ts t'. 598 | Proof. 599 | move => H; elim/reduction1_ind: t t' / H n m ts => /=; try by constructor. 600 | - by move => ty t1 t2 n m ts; rewrite subst_subst_distr //= !subn0. 601 | - by move => t ty n m ts; rewrite subst_substtyp_distr. 602 | Qed. 603 | 604 | Lemma subst_reduction t n m ts : 605 | Forall (fun t => t.1 ->r t.2) ts -> 606 | subst_term n m (unzip1 ts) t ->r subst_term n m (unzip2 ts) t. 607 | Proof. 608 | move => H; elim: t n m => //=; auto => v n m; rewrite !size_map; elimif_omega. 609 | elim: ts v H => //= [[t t']] ts IH [[H _] | v] //=. 610 | - move: H. 611 | apply (rtc_map' (f := fun x => 612 | typemap (shift_typ m) 0 (shift_term n 0 x))) => t1 t2 H. 613 | by apply shifttyp_reduction1, shift_reduction1. 614 | - by case => _ H; rewrite subSS; apply IH. 615 | Qed. 616 | 617 | Module confluence_proof. 618 | 619 | Reserved Notation "t ->rp t'" (at level 70, no associativity). 620 | 621 | Inductive parred : relation term := 622 | | parredfst ty t1 t1' t2 t2' : 623 | t1 ->rp t1' -> t2 ->rp t2' -> 624 | abs ty t1 @ t2 ->rp subst_term 0 0 [:: t2'] t1' 625 | | parredsnd t t' ty : 626 | t ->rp t' -> uabs t @' ty ->rp typemap (subst_typ^~ [:: ty]) 0 t' 627 | | parredvar n : var n ->rp var n 628 | | parredapp t1 t1' t2 t2' : 629 | t1 ->rp t1' -> t2 ->rp t2' -> t1 @ t2 ->rp t1' @ t2' 630 | | parredabs ty t t' : t ->rp t' -> abs ty t ->rp abs ty t' 631 | | parreduapp t t' ty : t ->rp t' -> t @' ty ->rp t' @' ty 632 | | parreduabs t t' : t ->rp t' -> uabs t ->rp uabs t' 633 | where "t ->rp t'" := (parred t t'). 634 | 635 | Fixpoint reduce_all_redex t : term := 636 | match t with 637 | | var _ => t 638 | | abs ty t1 @ t2 => 639 | subst_term 0 0 [:: reduce_all_redex t2] (reduce_all_redex t1) 640 | | t1 @ t2 => reduce_all_redex t1 @ reduce_all_redex t2 641 | | abs ty t' => abs ty (reduce_all_redex t') 642 | | uabs t @' ty => 643 | typemap (subst_typ^~ [:: ty]) 0 (reduce_all_redex t) 644 | | t @' ty => reduce_all_redex t @' ty 645 | | uabs t => uabs (reduce_all_redex t) 646 | end. 647 | 648 | Lemma parred_refl t : parred t t. 649 | Proof. by elim: t; constructor. Qed. 650 | 651 | Hint Constructors parred. 652 | Hint Resolve parred_refl. 653 | 654 | Lemma betared1_in_parred : inclusion reduction1 parred. 655 | Proof. by apply reduction1_ind; constructor. Qed. 656 | 657 | Lemma parred_in_betared : inclusion parred reduction. 658 | Proof. 659 | apply parred_ind => //; auto. 660 | - move => ty t1 t1' t2 t2' H H0 H1 H2. 661 | apply rtc_trans with (abs ty t1' @ t2); auto. 662 | apply rtc_trans with (abs ty t1' @ t2'); auto. 663 | by apply rtc_step. 664 | - move => t t' ty H H0. 665 | apply rtc_trans with (uabs t' @' ty); auto. 666 | by apply rtc_step. 667 | Qed. 668 | 669 | Lemma shift_parred t t' d c : 670 | t ->rp t' -> shift_term d c t ->rp shift_term d c t'. 671 | Proof. 672 | move => H; elim/parred_ind: t t' / H d c => //=; auto. 673 | - by move => t1 t1' t2 t2' ty H H0 H1 H2 d c; 674 | rewrite subst_shift_distr //= subn0; auto. 675 | - by move => t t' ty H H0 d c; rewrite shift_typemap_distr; auto. 676 | Qed. 677 | 678 | Lemma shifttyp_parred t t' d c : 679 | t ->rp t' -> typemap (shift_typ d) c t ->rp typemap (shift_typ d) c t'. 680 | Proof. 681 | move => H; elim/parred_ind: t t' / H d c => /=; auto. 682 | - by move => t1 t1' t2 t2' ty H H0 H1 H2 d c; 683 | rewrite shifttyp_subst_distr //= subn0; auto. 684 | - by move => t t' ty H H0 n m; 685 | rewrite substtyp_shifttyp_distr //= subn0; auto. 686 | Qed. 687 | 688 | Lemma substtyp_parred n tys t t' : 689 | t ->rp t' -> 690 | typemap (subst_typ^~ tys) n t ->rp typemap (subst_typ^~ tys) n t'. 691 | Proof. 692 | move => H; elim/parred_ind: t t' / H n => /=; auto. 693 | - by move => t1 t1' t2 t2' ty H H0 H1 H2 n; 694 | rewrite substtyp_subst_distr //= subn0; auto. 695 | - by move => t t' ty H H0 n; 696 | rewrite substtyp_substtyp_distr //= subn0; auto. 697 | Qed. 698 | 699 | Lemma subst_parred n m ps t t' : 700 | Forall (prod_curry parred) ps -> t ->rp t' -> 701 | subst_term n m (unzip1 ps) t ->rp subst_term n m (unzip2 ps) t'. 702 | Proof. 703 | move => H H0; elim/parred_ind: t t' / H0 n m => /=; auto. 704 | - by move => tl tl' tr tr' ty H0 H1 H2 H3 n m; 705 | rewrite subst_subst_distr //= !subn0; auto. 706 | - by move => t t' ty H0 H1 n m; rewrite subst_substtyp_distr //=; auto. 707 | - move => v n m; rewrite !size_map; case: ifP => //; elimleq. 708 | elim: ps H v => //= [[t t']] ts IH [H H0] [| v] //=. 709 | - by apply shifttyp_parred, shift_parred. 710 | - by rewrite subSS; apply IH. 711 | Qed. 712 | 713 | Lemma parred_all_lemma t t' : t ->rp t' -> t' ->rp reduce_all_redex t. 714 | Proof with auto. 715 | move: t t'; fix IH 3 => t t' H; inversion H; subst => {H} /=... 716 | - apply (@subst_parred 0 0 [:: (t2', reduce_all_redex t2)]) => /=... 717 | - exact/substtyp_parred/IH. 718 | - destruct t1; auto; inversion H0... 719 | - destruct t0; auto; inversion H0... 720 | Qed. 721 | 722 | Lemma parred_confluent : confluent parred. 723 | Proof. 724 | by move => t1 t2 t3 H H0; exists (reduce_all_redex t1); apply parred_all_lemma. 725 | Qed. 726 | 727 | Theorem betared_confluent : confluent reduction. 728 | Proof. 729 | apply (rtc_confluent' betared1_in_parred parred_in_betared parred_confluent). 730 | Qed. 731 | 732 | End confluence_proof. 733 | 734 | Module subject_reduction_proof. 735 | 736 | Lemma ctxleq_preserves_typing ctx1 ctx2 t ty : 737 | ctx1 <=c ctx2 -> ctx1 \|- t \: ty -> ctx2 \|- t \: ty. 738 | Proof. 739 | elim: t ty ctx1 ctx2 => 740 | [n | tl IHtl tr IHtr | tyl t IHt | t IHt tyr | t IHt] ty ctx1 ctx2. 741 | - move/ctxleqP; apply. 742 | - by move => H /typing_appP [tyl H0 H1]; apply/typing_appP; 743 | exists tyl; [apply (IHtl _ ctx1) | apply (IHtr _ ctx1)]. 744 | - by move => H /typing_absP [tyr -> H0]; rewrite typing_absE; 745 | move: H0; apply IHt; rewrite ctxleqE eqxx. 746 | - by move => H /typing_uappP [tyl -> H0]; apply/typing_uappP; 747 | exists tyl => //; move: H0; apply IHt. 748 | - by move => H /typing_uabsP [ty' -> H0]; rewrite typing_uabsE; 749 | move: H0; apply IHt, ctxleq_map. 750 | Qed. 751 | 752 | Lemma typing_shifttyp d c ctx t : 753 | typing (ctxmap (shift_typ d c) ctx) (typemap (shift_typ d) c t) = 754 | omap (shift_typ d c) (typing ctx t). 755 | Proof. 756 | rewrite /typing. 757 | elim: t d c ctx => /= 758 | [n | tl IHtl tr IHtr | ty t IHt | t IHt ty | t IHt] d c ctx. 759 | - by rewrite (nth_map' (omap _)) /=. 760 | - rewrite IHtl /=; move: (typing_rec ctx tl) => [] // [] //= tyl tyr. 761 | by rewrite IHtr; case: typing_rec => //= tyl'; 762 | rewrite !eqE /= inj_shift_typ; case: ifP. 763 | - by move: (IHt d c (Some ty :: ctx)) => /= ->; case: typing_rec. 764 | - rewrite IHt; move: (typing_rec ctx t) => [] // [] //= t'. 765 | by rewrite subst_shift_distr_ty //= subn0. 766 | - have H ty: omap (shift_typ d c) (omap tyabs ty) = 767 | omap tyabs (omap (shift_typ d c.+1) ty) by case: ty. 768 | by rewrite {}H -IHt; congr (omap tyabs (typing_rec _ _)); 769 | rewrite -!map_comp; apply eq_map; case => //= ty'; 770 | rewrite shift_shift_distr_ty. 771 | Qed. 772 | 773 | Lemma substtyp_preserves_typing n tys ctx t ty : 774 | ctx \|- t \: ty -> 775 | ctxmap (subst_typ n tys) ctx \|- 776 | typemap (subst_typ^~ tys) n t \: subst_typ n tys ty. 777 | Proof. 778 | elim: t ty n ctx => 779 | [m | tl IHtl tr IHtr | tyl t IHt | t IHt tyr | t IHt] ty n ctx /=. 780 | - by apply ctxindex_map. 781 | - by case/typing_appP => tyl H H0; apply/typing_appP; 782 | exists (subst_typ n tys tyl); [apply (IHtl (tyl ->t ty)) | apply IHtr]. 783 | - by case/typing_absP => tyr -> H; rewrite /= typing_absE; move: H; apply IHt. 784 | - case/typing_uappP => tyl -> H; apply/typing_uappP; 785 | exists (subst_typ n.+1 tys tyl). 786 | + by rewrite subst_subst_distr_ty //= subn0. 787 | + by move: H; apply IHt. 788 | - case/typing_uabsP => ty' -> H; rewrite /= typing_uabsE. 789 | suff ->: ctxmap (shift_typ 1 0) (ctxmap (subst_typ n tys) ctx) = 790 | ctxmap (subst_typ n.+1 tys) (ctxmap (shift_typ 1 0) ctx) by 791 | apply IHt. 792 | by rewrite -!map_comp; apply eq_map; 793 | case => //= ty''; rewrite shift_subst_distr_ty. 794 | Qed. 795 | 796 | Lemma typing_shift t c ctx1 ctx2 : 797 | typing (ctxinsert ctx2 ctx1 c) (shift_term (size ctx2) c t) = typing ctx1 t. 798 | Proof. 799 | rewrite /typing; elim: t c ctx1 ctx2 => 800 | /= [n | tl IHtl tr IHtr | tyl t IHt | t IHt tyr | t IHt] c ctx1 ctx2. 801 | - rewrite nth_insert; elimif_omega. 802 | - by rewrite IHtl IHtr. 803 | - by rewrite -(IHt c.+1 _ ctx2). 804 | - by rewrite IHt. 805 | - by rewrite map_insert /= -(size_map (omap (shift_typ 1 0)) ctx2) IHt. 806 | Qed. 807 | 808 | Lemma subject_subst t ty n m ctx ctx' : 809 | all (fun p => drop n ctx \|- typemap (shift_typ m) 0 p.1 \: p.2) ctx' -> 810 | ctxinsert [seq Some p.2 | p <- ctx'] ctx n \|- t \: ty -> 811 | ctx \|- subst_term n m (unzip1 ctx') t \: ty. 812 | Proof. 813 | elim: t ty n m ctx ctx' => 814 | [v | tl IHtl tr IHtr | tyl t IHt | t IHt tyr | t IHt] ty n m ctx ctx' H. 815 | - rewrite /typing /= nth_insert !size_map => /eqP ->; elimif. 816 | + apply/eqP/esym; rewrite nth_default ?size_map /=; elimif_omega. 817 | + rewrite -/typing !(nth_map (var 0, tyvar 0)) //. 818 | case: {ty v H0} (nth _ _ _) (all_nthP (var 0, tyvar 0) H v H0) => t ty. 819 | rewrite /= -(typing_shift _ 0 _ (ctxinsert [::] (take n ctx) n)) 820 | size_insert size_take minnC minKn add0n shift_typemap_distr. 821 | apply ctxleq_preserves_typing. 822 | rewrite /insert take0 sub0n take_minn minnn size_take minnE subKn 823 | ?leq_subr //= drop_take_nil cats0 drop0 -catA 824 | -{4}(cat_take_drop n ctx) ctxleq_appl. 825 | by case (leqP' n (size ctx)) => //= H0; 826 | rewrite drop_oversize ?(ltnW H0) //= cats0; 827 | apply/ctxleqP => /= i ty' /eqP; rewrite nth_nseq if_same. 828 | - by case/typing_appP => /= tyl H0 H1; apply/typing_appP; exists tyl; auto. 829 | - by case/typing_absP => /= tyr -> H0; rewrite typing_absE; apply IHt. 830 | - by case/typing_uappP => /= tyl -> H0; apply/typing_uappP; exists tyl; auto. 831 | - case/typing_uabsP => /= ty' -> H0; 832 | rewrite typing_uabsE -(addn0 m.+1) -subst_shifttyp_app. 833 | set ctx'' := (map _ (map _ _)). 834 | have {ctx''} ->: ctx'' = unzip1 835 | [seq (typemap (shift_typ m.+1) 0 p.1, shift_typ 1 0 p.2) | p <- ctx'] 836 | by rewrite /unzip1 /ctx'' -!map_comp /comp /=. 837 | apply IHt. 838 | + move: H {t ty IHt H0}; rewrite all_map; apply sub_all. 839 | rewrite /subpred /preim /pred_of_simpl; case => /= t ty. 840 | by rewrite -map_drop shifttyp_zero -(@shifttyp_add m _ 1 0) // 841 | typing_shifttyp; case: typing => // ty'' /eqP [] <-. 842 | + by move: H0; rewrite map_insert -!map_comp. 843 | Qed. 844 | 845 | Lemma subject_subst0 t ty ctx ctx' : 846 | all (fun p => ctx \|- p.1 \: p.2) ctx' -> 847 | [seq Some p.2 | p <- ctx'] ++ ctx \|- t \: ty -> 848 | ctx \|- subst_term 0 0 (unzip1 ctx') t \: ty. 849 | Proof. 850 | move => H; move: (@subject_subst t ty 0 0 ctx ctx'); 851 | rewrite /insert take0 sub0n drop0 /=; apply; move: H. 852 | by apply sub_all; case => t' ty' /=; rewrite shifttyp_zero. 853 | Qed. 854 | 855 | Lemma subject_subst0' ctx tyl tyr t t' : 856 | Some tyl :: ctx \|- t \: tyr -> ctx \|- t' \: tyl -> 857 | ctx \|- subst_term 0 0 [:: t'] t \: tyr. 858 | Proof. 859 | by move => H H0; 860 | apply (@subject_subst0 _ _ _ [:: (t', tyl)]); rewrite //= andbT. 861 | Qed. 862 | 863 | Theorem subject_reduction ctx t t' ty : 864 | t ->r1 t' -> ctx \|- t \: ty -> ctx \|- t' \: ty. 865 | Proof. 866 | move => H; elim/reduction1_ind: t t' / H ctx ty => /=. 867 | - by move => tyl t1 t2 ctx ty 868 | /typing_appP [tyl']; rewrite typing_absE' => /andP [] /eqP <- H H0; 869 | apply (subject_subst0' H). 870 | - move => t tyr ctx ty /typing_uappP [tyl] ->; rewrite typing_uabsE. 871 | move/(substtyp_preserves_typing 0 [:: tyr]). 872 | set ctx' := ctxmap _ _; have {ctx'} -> //: ctx' = ctx. 873 | rewrite {}/ctx' -map_comp -{2}(map_id ctx). 874 | apply eq_map; case => //= ty'. 875 | by rewrite subst_shift_cancel_ty2 //= shift_zero_ty. 876 | - by move => t1 t1' t2 H H0 ctx ty /typing_appP [tyl H1 H2]; 877 | apply/typing_appP; exists tyl; auto. 878 | - by move => t1 t2 t2' H H0 ctx ty /typing_appP [tyl H1 H2]; 879 | apply/typing_appP; exists tyl; auto. 880 | - by move => tyl t t' H H0 ctx ty /typing_absP [tyr -> H1]; 881 | rewrite typing_absE; auto. 882 | - by move => t t' tyr H H0 ctx ty /typing_uappP [tyl -> H1]; 883 | apply/typing_uappP; exists tyl; auto. 884 | - by move => t t' H H0 ctx ty /typing_uabsP [ty' -> H1]; 885 | rewrite typing_uabsE; auto. 886 | Qed. 887 | 888 | End subject_reduction_proof. 889 | 890 | Module strong_normalization_proof_typefree. 891 | 892 | Section CRs. 893 | Variable (P : term -> Prop). 894 | Definition CR1 := forall t, P t -> SN t. 895 | Definition CR2 := forall t t', t ->r1 t' -> P t -> P t'. 896 | Definition CR3 := forall t, neutral t -> (forall t', t ->r1 t' -> P t') -> P t. 897 | Record RC : Prop := reducibility_candidate 898 | { rc_cr1 : CR1 ; rc_cr2 : CR2 ; rc_cr3 : CR3 }. 899 | End CRs. 900 | 901 | Lemma CR2' P t t' : RC P -> t ->r t' -> P t -> P t'. 902 | Proof. 903 | move => H; elim/clos_refl_trans_1n_ind: t t' / => //= x y z H0 _ H1 H2. 904 | by apply H1, (rc_cr2 H H0). 905 | Qed. 906 | 907 | Lemma CR4 P t : RC P -> neutral t -> (forall t', ~ t ->r1 t') -> P t. 908 | Proof. by move/rc_cr3 => H H0 H1; apply H => // t' /H1. Qed. 909 | 910 | Lemma CR4' P (v : nat) : RC P -> P v. 911 | Proof. by move/CR4; apply => // t' H; inversion H. Qed. 912 | 913 | Lemma snorm_isrc : RC SN. 914 | Proof. 915 | constructor; move => /=; try tauto. 916 | - by move => t t' H []; apply. 917 | - by move => t H H0; constructor => t' H1; apply H0. 918 | Qed. 919 | 920 | Lemma rcfun_isrc P Q : 921 | RC P -> RC Q -> RC (fun u => forall v, P v -> Q (u @ v)). 922 | Proof. 923 | move => H H0; constructor; move => /=. 924 | - by move => t /(_ 0 (CR4' _ H)) /(rc_cr1 H0); 925 | apply (acc_preservation (f := app^~_)); constructor. 926 | - by move => t t' H1 H2 v /H2; apply (rc_cr2 H0); constructor. 927 | - move => t1 H1 H2 t2 H3. 928 | elim/Acc_ind': t2 / {H3} (rc_cr1 H H3) (H3) => t2 H3 H4. 929 | apply (rc_cr3 H0) => //= t3 H5; move: H1; inversion H5; 930 | subst => //= _; auto; apply H3 => //; apply (rc_cr2 H H8 H4). 931 | Qed. 932 | 933 | Fixpoint reducible ty (preds : seq (term -> Prop)) : term -> Prop := 934 | match ty with 935 | | tyvar v => nth SN preds v 936 | | tyfun tyl tyr => fun t => 937 | forall t', reducible tyl preds t' -> reducible tyr preds (t @ t') 938 | | tyabs ty => fun t => 939 | forall ty' P, RC P -> reducible ty (P :: preds) (t @' ty') 940 | end. 941 | 942 | Lemma reducibility_isrc ty preds : Forall RC preds -> RC (reducible ty preds). 943 | Proof. 944 | elim: ty preds => /= [n | tyl IHl tyr IHr | ty IH] preds. 945 | - move/Forall_nth /(_ _ n); case: (ltnP n (size preds)). 946 | + by move => _; apply. 947 | + by move => H _; rewrite (nth_default _ H); apply snorm_isrc. 948 | - by move => H; apply rcfun_isrc; [apply IHl | apply IHr]. 949 | - move => H; constructor. 950 | + move => /= t /(_ 0 SN snorm_isrc) 951 | /(rc_cr1 (IH (_ :: _) (conj snorm_isrc H))). 952 | by apply (acc_preservation (f := uapp^~_)) => x y H0; constructor. 953 | + move => /= t t' H0 H1 ty' P H2; move: (H1 ty' _ H2). 954 | by apply (rc_cr2 (IH (_ :: _) (conj H2 H))); constructor. 955 | + move => /= t H0 H1 ty' P H2. 956 | apply (rc_cr3 (IH (_ :: _) (conj H2 H))) => //= t' H3. 957 | by move: H0; inversion H3; subst => //= _; apply H1. 958 | Qed. 959 | 960 | Lemma shift_reducibility c ty preds preds' t : 961 | c <= size preds -> 962 | (reducible (shift_typ (size preds') c ty) (insert preds' preds SN c) t <-> 963 | reducible ty preds t). 964 | Proof. 965 | elim: ty c preds t => [v | tyl IHl tyr IHr | ty IH] c preds t H. 966 | - rewrite /= nth_insert; elimif_omega. 967 | - by split => /= H0 t' /(IHl c _ _ H) /H0 /(IHr c _ _ H). 968 | - by split => H0 ty' P H1; apply (IH c.+1 (P :: preds)) => //; apply H0. 969 | Qed. 970 | 971 | Lemma subst_reducibility ty preds n tys t : 972 | n <= size preds -> 973 | (reducible (subst_typ n tys ty) preds t <-> 974 | reducible ty 975 | (insert [seq reducible ty (drop n preds) | ty <- tys] preds SN n) t). 976 | Proof. 977 | elim: ty preds n tys t => [v | tyl IHl tyr IHr | ty IH] preds n tys t. 978 | - rewrite /= nth_insert size_map; elimif_omega. 979 | + by rewrite nth_default ?leq_addr //= addnC. 980 | + move => H0. 981 | rewrite (nth_map (tyvar (v - size tys))) //=. 982 | move: (shift_reducibility (nth (tyvar (v - size tys)) tys v) 983 | (take n preds) t (leq0n (size (drop n preds)))). 984 | by rewrite /insert take0 drop0 /= cat_take_drop size_takel. 985 | - by move => H /=; split => H0 t' /IHl => /(_ H) /H0 /IHr => /(_ H). 986 | - move => H /=; split => H0 ty' P /(H0 ty') {H0}. 987 | + by move/IH => /= /(_ H); rewrite /insert /= subSS. 988 | + by move => H0; apply IH. 989 | Qed. 990 | 991 | Lemma abs_reducibility t tyl tyr tys preds : 992 | Forall RC preds -> 993 | (forall t', 994 | reducible tyl preds t' -> 995 | reducible tyr preds (subst_term 0 0 [:: t'] t)) -> 996 | reducible (tyl ->t tyr) preds (abs (subst_typ 0 tys tyl) t). 997 | Proof. 998 | move => /= H. 999 | move: (reducible tyl preds) (reducible tyr preds) 1000 | (reducibility_isrc tyl H) (reducibility_isrc tyr H) 1001 | => {H} P Q HP HQ H t' H0. 1002 | have H1: SN t by 1003 | move: (rc_cr1 HQ (H _ H0)); 1004 | apply acc_preservation => x y; apply subst_reduction1. 1005 | move: t t' {H H0} H1 (rc_cr1 HP H0) (H0) (H _ H0); 1006 | refine (Acc_ind2 _) => t t' IHt IHt' H H0. 1007 | apply rc_cr3 => // t'' H1; inversion H1; subst => //. 1008 | - by inversion H5; subst; apply IHt => //; 1009 | apply (rc_cr2 HQ (subst_reduction1 0 0 [:: t'] H6)). 1010 | - apply IHt' => //; first by apply (rc_cr2 HP H5). 1011 | move: H0; apply (CR2' HQ). 1012 | by apply (@subst_reduction t 0 0 [:: (t', t2')]) => /=; 1013 | intuition apply rtc_step. 1014 | Qed. 1015 | 1016 | Lemma uabs_reducibility t ty preds : 1017 | Forall RC preds -> 1018 | (forall v P, RC P -> 1019 | reducible ty (P :: preds) (typemap (subst_typ^~ [:: v]) 0 t)) -> 1020 | reducible (tyabs ty) preds (uabs t). 1021 | Proof. 1022 | move => /= H H0 ty' P H1. 1023 | move: (reducible _ _) (@reducibility_isrc ty (P :: preds)) 1024 | (H0 ty' P H1) => P' /= /(_ (conj H1 H)) {H H0 H1} H H0. 1025 | have H1: SN t by 1026 | move: (rc_cr1 H H0); apply acc_preservation => x y; apply substtyp_reduction1. 1027 | elim/Acc_ind': t / H1 H0 => t H0 H1. 1028 | apply (rc_cr3 H) => //= t' H2; inversion H2; subst => //. 1029 | inversion H6; subst; apply H0 => //. 1030 | apply (rc_cr2 H (substtyp_reduction1 _ _ H4) H1). 1031 | Qed. 1032 | 1033 | Lemma uapp_reducibility t ty ty' tys preds : 1034 | Forall RC preds -> reducible (tyabs ty) preds t -> 1035 | reducible (subst_typ 0 [:: ty'] ty) preds (t @' subst_typ 0 tys ty'). 1036 | Proof. 1037 | move => /= H H0. 1038 | apply subst_reducibility => //=. 1039 | rewrite /insert take0 sub0n !drop0 /=. 1040 | by apply H0, reducibility_isrc. 1041 | Qed. 1042 | 1043 | Lemma reduce_lemma ctx preds t ty : 1044 | [seq Some c.2 | c <- ctx] \|- t \: ty -> 1045 | Forall RC (unzip2 preds) -> 1046 | Forall (fun c => reducible c.2 (unzip2 preds) c.1) ctx -> 1047 | reducible ty (unzip2 preds) 1048 | (subst_term 0 0 (unzip1 ctx) (typemap (subst_typ^~ (unzip1 preds)) 0 t)). 1049 | Proof. 1050 | elim: t ty ctx preds => /=. 1051 | - move => v ty ctx preds H H0 H1. 1052 | rewrite shifttyp_zero shift_zero subn0 size_map. 1053 | elim: ctx v H H1 => //=; first by case. 1054 | case => t ty' ctx IH [] //=. 1055 | + by case/eqP => <- []. 1056 | + by move => v H [H1 H2]; rewrite subSS; apply IH. 1057 | - move => tl IHtl tr IHtr ty ctx preds /typing_appP [tyl H H0] H1 H2. 1058 | by move: (IHtl (tyl ->t ty) ctx preds H H1 H2) => /=; apply; apply IHtr. 1059 | - move => tyl t IHt ty ctx preds /typing_absP [tyr -> H] H0 H1. 1060 | apply abs_reducibility => // t' H2. 1061 | by rewrite subst_app //=; apply (IHt tyr ((t', tyl) :: ctx)). 1062 | - move => t IHt tyr ty ctx preds /typing_uappP [tyl -> H] H0 H1. 1063 | by apply uapp_reducibility => //; apply IHt. 1064 | - move => t IHt ty ctx preds /typing_uabsP [ty' ->]; 1065 | rewrite -map_comp /comp /= => H H0 H1. 1066 | apply uabs_reducibility => // v P H2. 1067 | rewrite -(subst_substtyp_distr 0 [:: v]) // typemap_compose /=. 1068 | have /(typemap_eq 0 t) -> : forall i ty, 1069 | subst_typ (i + 0) [:: v] (subst_typ (i + 1) (unzip1 preds) ty) = 1070 | subst_typ i (v :: unzip1 preds) ty by 1071 | move => i ty''; rewrite addn0 addn1 subst_app_ty. 1072 | move: (IHt ty' [seq (c.1, shift_typ 1 0 c.2) | c <- ctx] ((v, P) :: preds)). 1073 | rewrite /unzip1 -!map_comp /comp /=; apply => //=. 1074 | apply Forall_map; move: H1; rewrite /comp /=; apply Forall_impl => 1075 | {t ty IHt H H0 H2} [[t ty]] /= 1076 | /(proj2 (shift_reducibility ty [:: P] t (leq0n _))). 1077 | by rewrite /insert take0 drop0 sub0n /=; apply. 1078 | Qed. 1079 | 1080 | Theorem typed_term_is_snorm ctx t ty : ctx \|- t \: ty -> SN t. 1081 | Proof. 1082 | move => H. 1083 | move: (@reduce_lemma 1084 | (map (oapp (pair (var 0)) (var 0, tyvar 0)) ctx) [::] t ty) => /=. 1085 | rewrite -map_comp /comp /=. 1086 | have {H} H: 1087 | [seq Some (oapp (pair (var 0)) (var 0, tyvar 0) c).2 | c <- ctx] \|- t \: ty 1088 | by move: H; apply subject_reduction_proof.ctxleq_preserves_typing; 1089 | elim: ctx {t ty} => //; case => // ty ctx H; rewrite ctxleqE eqxx /=. 1090 | move/(_ H I) {H}. 1091 | have H: Forall (fun c => reducible c.2 [::] c.1) 1092 | (map (oapp (pair (var 0)) (var 0, tyvar 0)) ctx) by 1093 | apply Forall_map, Forall_nth => /= x m _; rewrite /oapp /snd; 1094 | case: nth => [ty' |]; apply CR4', reducibility_isrc. 1095 | move/(_ H) => {H} /(rc_cr1 (reducibility_isrc _ _)) /= /(_ I). 1096 | set f := subst_term _ _ _; set g := typemap _ _. 1097 | apply (acc_preservation (f := f \o g)) => x y H. 1098 | by apply subst_reduction1, substtyp_reduction1. 1099 | Qed. 1100 | 1101 | End strong_normalization_proof_typefree. 1102 | 1103 | Module strong_normalization_proof_typed. 1104 | 1105 | Import subject_reduction_proof. 1106 | 1107 | Notation bottom := (tyabs (tyvar 0)). 1108 | Notation "# ctx" := (rcons ctx (Some bottom)) (at level 65, no associativity). 1109 | 1110 | Section CRs. 1111 | Variable (ctx : context typ) (ty : typ) (P : term -> Prop). 1112 | Definition CR1 := forall t, #ctx \|- t \: ty -> P t -> SN t. 1113 | Definition CR2 := forall t t', #ctx \|- t \: ty -> t ->r1 t' -> P t -> P t'. 1114 | Definition CR3 := forall t, 1115 | #ctx \|- t \: ty -> neutral t -> (forall t', t ->r1 t' -> P t') -> P t. 1116 | Record RC : Prop := reducibility_candidate 1117 | { rc_cr1 : CR1 ; rc_cr2 : CR2 ; rc_cr3 : CR3 }. 1118 | End CRs. 1119 | 1120 | Lemma CR2' ctx ty P t t' : 1121 | RC ctx ty P -> #ctx \|- t \: ty -> t ->r t' -> P t -> P t'. 1122 | Proof. 1123 | move => H H0 H1. 1124 | elim/clos_refl_trans_1n_ind: t t' / H1 H0 => //= x y z H0 H1 H2 H3 H4. 1125 | by apply H2; [apply (subject_reduction H0) | apply (rc_cr2 H H3)]. 1126 | Qed. 1127 | 1128 | Lemma CR4 ctx ty P t : RC ctx ty P -> 1129 | #ctx \|- t \: ty -> neutral t -> (forall t', ~ t ->r1 t') -> P t. 1130 | Proof. by move/rc_cr3 => H H0 H1 H2; apply H => // t' /H2. Qed. 1131 | 1132 | Lemma CR4' ctx ty P (v : nat) : RC ctx ty P -> ctxindex (#ctx) v ty -> P v. 1133 | Proof. move/CR4 => H H0; apply H => // => t' H1; inversion H1. Qed. 1134 | 1135 | Lemma snorm_isrc ctx ty : RC ctx ty SN. 1136 | Proof. 1137 | constructor; move => /=; auto. 1138 | - by move => t t' H H0 []; apply. 1139 | - by move => t _ _ /(Acc_intro t). 1140 | Qed. 1141 | 1142 | Lemma rcfun_isrc ctx tyl tyr P Q : 1143 | RC ctx tyl P -> RC ctx tyr Q -> 1144 | RC ctx (tyl ->t tyr) 1145 | (fun u => forall v, #ctx \|- v \: tyl -> P v -> Q (u @ v)). 1146 | Proof. 1147 | move => HP HQ; constructor; move => /=. 1148 | - move => t H /(_ (size ctx @' tyl)). 1149 | have H0 : #ctx \|- size ctx @' tyl \: tyl by 1150 | rewrite /typing /= nth_rcons ltnn eqxx /= shift_zero_ty. 1151 | have H1 t': ~ size ctx @' tyl ->r1 t' by 1152 | move => H1; inversion H1; subst; inversion H5. 1153 | move/(_ H0 (CR4 HP H0 erefl H1))/(rc_cr1 HQ) => {H1}. 1154 | have -> : #ctx \|- t @ (size ctx @' tyl) \: tyr by 1155 | apply/typing_appP; exists tyl. 1156 | by move/(_ erefl); apply (acc_preservation (f := app^~_)); auto. 1157 | - move => t t' H H0 H1 v H2 /(H1 _ H2); apply (rc_cr2 HQ). 1158 | + by apply/typing_appP; exists tyl. 1159 | + by constructor. 1160 | - move => t H H0 H1 v H2 H3. 1161 | elim/Acc_ind': v / {H2 H3} (rc_cr1 HP H2 H3) (H2) (H3) => v IH H2 H3. 1162 | apply (rc_cr3 HQ) => //=; first by apply/typing_appP; exists tyl. 1163 | move => t'' H4; move: H0; inversion H4; subst => //= _. 1164 | + by apply (H1 _ H7). 1165 | + apply IH => //. 1166 | * by apply (subject_reduction H7). 1167 | * by apply (rc_cr2 HP H2 H7). 1168 | Qed. 1169 | 1170 | Fixpoint reducible ctx ty (preds : seq (typ * (term -> Prop))) t : Prop := 1171 | match ty with 1172 | | tyvar v => nth SN (unzip2 preds) v t 1173 | | tyfun tyl tyr => forall v, 1174 | #ctx \|- v \: subst_typ 0 (unzip1 preds) tyl -> 1175 | reducible ctx tyl preds v -> reducible ctx tyr preds (t @ v) 1176 | | tyabs ty => 1177 | forall ty' P, RC ctx ty' P -> 1178 | reducible ctx ty ((ty', P) :: preds) (t @' ty') 1179 | end. 1180 | 1181 | Lemma reducibility_isrc ctx ty preds : 1182 | Forall (fun p => RC ctx p.1 p.2) preds -> 1183 | RC ctx (subst_typ 0 (unzip1 preds) ty) (reducible ctx ty preds). 1184 | Proof. 1185 | elim: ty preds => /= [n | tyl IHl tyr IHr | ty IH] preds. 1186 | - rewrite shift_zero_ty subn0. 1187 | elim: preds n => [| P preds IH []] /=. 1188 | + move => n _; rewrite !nth_nil /= !subn0; apply snorm_isrc. 1189 | + by case. 1190 | + by move => n [H] /(IH n) {IH}; rewrite !subSS. 1191 | - by move => H; apply rcfun_isrc; [apply IHl | apply IHr]. 1192 | - move => H; constructor. 1193 | + move => /= t H0 /(_ 0 _ (snorm_isrc _ _)) 1194 | /(rc_cr1 (IH ((_, _) :: _) (conj (snorm_isrc _ _) H))) /=. 1195 | suff: #ctx \|- t @' 0 \: subst_typ 0 (tyvar 0 :: unzip1 preds) ty by 1196 | move => -> /(_ erefl); 1197 | apply (acc_preservation (f := uapp^~_)) => x y H1; constructor. 1198 | by apply/typing_uappP; 1199 | exists (subst_typ 1 (unzip1 preds) ty) => //; rewrite subst_app_ty. 1200 | + move => /= t t' H0 H1 H2 ty' P H3. 1201 | move: (H2 _ _ H3); apply (rc_cr2 (IH ((_, _) :: _) (conj H3 H))). 1202 | * by apply/typing_uappP; 1203 | exists (subst_typ 1 (unzip1 preds) ty) => //; rewrite subst_app_ty. 1204 | * by constructor. 1205 | + move => /= t H0 H1 H2 ty' P H3. 1206 | apply (rc_cr3 (IH ((_, _) :: _) (conj H3 H))) => //=. 1207 | + by apply/typing_uappP; 1208 | exists (subst_typ 1 (unzip1 preds) ty) => //; rewrite subst_app_ty. 1209 | * by move => t' H4; move: H0; inversion H4; subst => //= _; apply H2. 1210 | Qed. 1211 | 1212 | Lemma shift_reducibility c ty preds preds' ctx t : 1213 | c <= size preds -> 1214 | (reducible ctx (shift_typ (size preds') c ty) 1215 | (insert preds' preds (tyvar 0, SN) c) t <-> 1216 | reducible ctx ty preds t). 1217 | Proof. 1218 | elim: ty c preds t => [v | tyl IHl tyr IHr | ty IH] c preds t H. 1219 | - rewrite /= /unzip2 map_insert nth_insert size_map; elimif_omega. 1220 | - rewrite /= /unzip1 map_insert -(size_map (@fst _ _)) /= 1221 | !subst_shift_cancel_ty4 ?size_map //. 1222 | by split => H0 v H1 /(IHl c _ _ H) /(H0 _ H1) /(IHr c _ _ H). 1223 | - by split => /= H0 ty' P H1; 1224 | apply (IH c.+1 ((ty', P) :: preds)) => //; apply H0. 1225 | Qed. 1226 | 1227 | Lemma subst_reducibility ty preds n tys ctx t : 1228 | n <= size preds -> 1229 | (reducible ctx (subst_typ n tys ty) preds t <-> 1230 | reducible ctx ty 1231 | (insert [seq (subst_typ 0 (drop n (unzip1 preds)) ty, 1232 | reducible ctx ty (drop n preds)) | ty <- tys] 1233 | preds (tyvar 0, SN) n) 1234 | t). 1235 | Proof. 1236 | elim: ty preds n tys t => [v | tyl IHl tyr IHr | ty IH] preds n tys t. 1237 | - rewrite /= /unzip2 map_insert /= -map_comp /comp /= 1238 | nth_insert size_map -/unzip2; elimif_omega. 1239 | + by rewrite nth_default ?leq_addr //= addnC. 1240 | + move => H0. 1241 | rewrite (nth_map (tyvar (v - size tys))) //=. 1242 | move: (shift_reducibility (nth (tyvar (v - size tys)) tys v) 1243 | (take n preds) ctx t (leq0n (size (drop n preds)))). 1244 | by rewrite /insert take0 drop0 /= cat_take_drop size_takel. 1245 | - by move => H /=; 1246 | split => H0 t' H1 /IHl => /(_ H) /H0 /IHr => /(_ H); apply; move: H1; 1247 | rewrite /unzip1 map_insert -map_comp /= subst_subst_compose_ty ?size_map. 1248 | - move => H /=; split => H0 ty' P /H0. 1249 | + by move/IH => /(_ H); rewrite /insert /= subSS. 1250 | + by move => H1; apply IH. 1251 | Qed. 1252 | 1253 | Lemma abs_reducibility tyl tyr preds ctx t : 1254 | #ctx \|- abs (subst_typ 0 (unzip1 preds) tyl) t 1255 | \: subst_typ 0 (unzip1 preds) (tyl ->t tyr) -> 1256 | Forall (fun p => RC ctx p.1 p.2) preds -> 1257 | (forall t', 1258 | #ctx \|- t' \: subst_typ 0 (unzip1 preds) tyl -> 1259 | reducible ctx tyl preds t' -> 1260 | reducible ctx tyr preds (subst_term 0 0 [:: t'] t)) -> 1261 | reducible ctx (tyl ->t tyr) preds (abs (subst_typ 0 (unzip1 preds) tyl) t). 1262 | Proof. 1263 | rewrite /= typing_absE => H H0. 1264 | move: (reducible ctx tyl preds) (reducible ctx tyr preds) 1265 | (reducibility_isrc tyl H0) (reducibility_isrc tyr H0) => 1266 | {H0} P Q HP HQ H0 t' H1 H2. 1267 | have H3: SN t by 1268 | move: (rc_cr1 HQ (subject_subst0' H H1) (H0 t' H1 H2)); 1269 | apply acc_preservation => x y; apply subst_reduction1. 1270 | move: t t' H3 {H0 H1 H2} (rc_cr1 HP H1 H2) H (H1) (H2) (H0 _ H1 H2). 1271 | refine (Acc_ind2 _) => t t' IHt IHt' H H0 H1 H2; apply (rc_cr3 HQ) => //=. 1272 | - by apply/typing_appP; 1273 | exists (subst_typ 0 (unzip1 preds) tyl) => //; rewrite typing_absE. 1274 | - move => t'' H3; inversion H3; subst => //. 1275 | + inversion H7; subst; apply IHt => //. 1276 | * by move: H; apply subject_reduction. 1277 | * by apply (@rc_cr2 _ _ _ HQ (subst_term 0 0 [:: t'] t)) => //; 1278 | [apply (subject_subst0' H) | apply subst_reduction1]. 1279 | + apply IHt' => //. 1280 | * by apply (subject_reduction H7). 1281 | * by apply (rc_cr2 HP H0 H7). 1282 | * move: H2; apply (CR2' HQ); first by apply (subject_subst0' H). 1283 | by apply (@subst_reduction t 0 0 [:: (t', t2')]) => /=; 1284 | intuition apply rtc_step. 1285 | Qed. 1286 | 1287 | Lemma uabs_reducibility ty preds ctx t : 1288 | #ctx \|- uabs t \: subst_typ 0 (unzip1 preds) (tyabs ty) -> 1289 | Forall (fun p => RC ctx p.1 p.2) preds -> 1290 | (forall v P, RC ctx v P -> 1291 | reducible ctx ty ((v, P) :: preds) (typemap (subst_typ^~ [:: v]) 0 t)) -> 1292 | reducible ctx (tyabs ty) preds (uabs t). 1293 | Proof. 1294 | rewrite /= typing_uabsE => H H0 H1 ty' P H2. 1295 | move: (substtyp_preserves_typing 0 [:: ty'] H) (H1 ty' P H2). 1296 | have ->: ctxmap (subst_typ 0 [:: ty']) (ctxmap (shift_typ 1 0) (#ctx)) = #ctx 1297 | by rewrite !map_rcons /=; congr rcons; 1298 | elim: ctx {H H0 H1 H2} => //= ty'' ctx {2}<-; case: ty'' => //= ty''; 1299 | rewrite subst_shift_cancel_ty2 //= subnn shift_zero_ty. 1300 | move: (reducible _ _) (@reducibility_isrc ctx ty ((ty', P) :: preds)) 1301 | => P' /= /(_ (conj H2 H0)) {H0 H1 H2}. 1302 | rewrite -typing_uabsE -/(subst_typ 0 _ (tyabs _)) in H. 1303 | rewrite subst_app_ty /= => H0 H1 H2. 1304 | have H3: SN t by 1305 | move: (rc_cr1 H0 H1 H2); 1306 | apply acc_preservation => x y; apply substtyp_reduction1. 1307 | elim/Acc_ind': t / H3 H H1 H2 => t IH H H1 H2; apply (rc_cr3 H0) => //=. 1308 | - by apply/typing_uappP; 1309 | exists (subst_typ 1 (unzip1 preds) ty) => //; rewrite subst_app_ty. 1310 | - move => t' H4; inversion H4; subst => //; inversion H7; subst; apply IH => //. 1311 | + by apply (subject_reduction H7). 1312 | + by move: H1; apply subject_reduction, substtyp_reduction1. 1313 | + by apply (rc_cr2 H0 H1 (substtyp_reduction1 _ _ H5) H2). 1314 | Qed. 1315 | 1316 | Lemma uapp_reducibility ty ty' preds ctx t : 1317 | Forall (fun p => RC ctx p.1 p.2) preds -> 1318 | reducible ctx (tyabs ty) preds t -> 1319 | reducible ctx (subst_typ 0 [:: ty'] ty) preds 1320 | (t @' subst_typ 0 (unzip1 preds) ty'). 1321 | Proof. 1322 | move => /= H H0; apply subst_reducibility => //=. 1323 | by rewrite /insert take0 sub0n !drop0 /=; apply H0, reducibility_isrc. 1324 | Qed. 1325 | 1326 | Lemma reduce_lemma ctx ctx' preds t ty : 1327 | [seq Some c.2 | c <- ctx] \|- t \: ty -> 1328 | Forall (fun p => RC ctx' p.1 p.2) preds -> 1329 | Forall (fun c => #ctx' \|- c.1 \: subst_typ 0 (unzip1 preds) c.2 /\ 1330 | reducible ctx' c.2 preds c.1) ctx -> 1331 | reducible ctx' ty preds 1332 | (subst_term 0 0 (unzip1 ctx) (typemap (subst_typ^~ (unzip1 preds)) 0 t)). 1333 | Proof. 1334 | have Hty: forall t ty ctx ctx' preds, 1335 | [seq Some c.2 | c <- ctx] \|- t \: ty -> 1336 | Forall (fun p => RC ctx' p.1 p.2) preds -> 1337 | Forall (fun c => #ctx' \|- c.1 \: subst_typ 0 (unzip1 preds) c.2 /\ 1338 | reducible ctx' c.2 preds c.1) ctx -> 1339 | #ctx' \|- 1340 | subst_term 0 0 (unzip1 ctx) (typemap (subst_typ^~ (unzip1 preds)) 0 t) \: 1341 | subst_typ 0 (unzip1 preds) ty. 1342 | clear => t ty ctx ctx' preds H H0 H1. 1343 | move: (fun t ty => @subject_subst0 t ty (#ctx') 1344 | [seq (c.1, subst_typ 0 (unzip1 preds) c.2) | c <- ctx]). 1345 | rewrite /unzip1 -!map_comp !/comp -/unzip1; apply => /=. 1346 | - by elim: ctx H1 {H} => //= [[]] /= t' ty' ctx IH [] [-> _] /=. 1347 | - by move: (substtyp_preserves_typing 0 (unzip1 preds) H); 1348 | rewrite -map_comp; apply ctxleq_preserves_typing, ctxleq_appr. 1349 | elim: t ty ctx ctx' preds => /=. 1350 | - move => v ty ctx ctx' preds H H0 H1. 1351 | rewrite shifttyp_zero shift_zero subn0 size_map. 1352 | elim: ctx v H H1 => //=; first by case. 1353 | case => t ty' ctx IH [] //=. 1354 | + by case/eqP => <- [] []. 1355 | + by move => v H [H1 H2]; rewrite subSS; apply IH. 1356 | - move => tl IHtl tr IHtr ty ctx ctx' preds /typing_appP [tyl H H0] H1 H2. 1357 | by move: (IHtl (tyl ->t ty) ctx ctx' preds H H1 H2) => /=; 1358 | apply; [apply Hty | apply IHtr]. 1359 | - move => tyl t IHt ty ctx ctx' preds /typing_absP [tyr -> H] H0 H1. 1360 | apply abs_reducibility => //; 1361 | first by apply (Hty (abs tyl t)) => //; rewrite typing_absE. 1362 | by move => t' H2 H3; rewrite subst_app //=; 1363 | apply (IHt tyr ((t', tyl) :: ctx)). 1364 | - move => t IHt tyr ty ctx ctx' preds /typing_uappP [tyl -> H] H0 H1. 1365 | by apply uapp_reducibility => //; apply IHt. 1366 | - move => t IHt ty ctx ctx' preds /typing_uabsP [ty' -> H] H0 H1. 1367 | apply uabs_reducibility => //; 1368 | first by apply (Hty (uabs t)) => //; rewrite typing_uabsE. 1369 | move => v P H2. 1370 | rewrite -(subst_substtyp_distr 0 [:: v]) // typemap_compose /=. 1371 | have /(typemap_eq 0 t) -> : forall i ty, 1372 | subst_typ (i + 0) [:: v] (subst_typ (i + 1) (unzip1 preds) ty) = 1373 | subst_typ i (unzip1 ((v, P) :: preds)) ty by 1374 | move => i ty''; rewrite addn0 addn1 subst_app_ty. 1375 | move: (IHt ty' 1376 | [seq (c.1, shift_typ 1 0 c.2) | c <- ctx] ctx' ((v, P) :: preds)). 1377 | rewrite /unzip1 -!map_comp /comp /=; apply => //=. 1378 | + by move: H; rewrite -map_comp /comp /=. 1379 | + apply Forall_map; move: H1; apply Forall_impl => [[t' ty'']] /= []. 1380 | rewrite subst_shift_cancel_ty3 //= drop0 !add0n subSS sub0n shift_zero_ty. 1381 | case (shift_reducibility ty'' [:: (v, P)] ctx' t' (leq0n (size preds))). 1382 | rewrite /insert take0 drop0 sub0n; auto. 1383 | Qed. 1384 | 1385 | Lemma identical_substitution ctx' ctx t ty n : 1386 | ctx' ++ ctx \|- t \: ty -> 1387 | subst_term (size ctx') n [seq var i | i <- iota 0 (size ctx)] t = t. 1388 | Proof. 1389 | elim: t ctx' ctx ty n => //=. 1390 | - move => v ctx' ctx ty n; case: ifP => //; elimleq. 1391 | rewrite /typing /= size_map size_iota -nth_map' /= 1392 | nth_cat ltnNge leq_addr //= addKn addnC => H; congr addn. 1393 | have H0: v < size ctx by elim: v ctx H => [| v IH] []. 1394 | by move: (ltnW H0); rewrite -subn_eq0 => /eqP ->; rewrite nth_iota. 1395 | - by move => tl IHtl tr IHtr ctx' ctx tyr n /typing_appP [tyl H H0]; 1396 | rewrite (IHtl _ _ (tyl ->t tyr)) // (IHtr _ _ tyl). 1397 | - move => tyl t IHt ctx' ctx ty n /typing_absP [tyr _ H]. 1398 | rewrite (IHt (Some tyl :: ctx') _ tyr) //. 1399 | - move => t IHt tyr ctx' ctx ty n /typing_uappP [tyl _ H]. 1400 | by rewrite (IHt ctx' ctx (tyabs tyl)). 1401 | - move => t IHt ctx' ctx ty n /typing_uabsP [ty' _]; 1402 | rewrite map_cat => /IHt => /(_ n.+1). 1403 | by rewrite !size_map => ->. 1404 | Qed. 1405 | 1406 | Theorem typed_term_is_snorm ctx t ty : ctx \|- t \: ty -> SN t. 1407 | Proof. 1408 | move => H. 1409 | have {H}: map some (map (odflt (tyvar 0)) ctx) \|- t \: ty by 1410 | move: H; apply ctxleq_preserves_typing; 1411 | elim: ctx {t ty} => // [[]] // ty ctx H; rewrite ctxleqE eqxx. 1412 | move: {ctx} (map _ ctx) => ctx H; move: (H). 1413 | have ->: map some ctx = 1414 | [seq Some c.2 | c <- zip (map var (iota 0 (size ctx))) ctx] by 1415 | elim: {H} ctx 0 => //= {t ty} ty ctx IH n; rewrite -IH. 1416 | move/reduce_lemma => /(_ (map some ctx) [::] I). 1417 | rewrite /= unzip1_zip ?size_map ?size_iota ?leqnn; last by []. 1418 | have H0: Forall 1419 | (fun c => # [seq Some i | i <- ctx] \|- c.1 \: subst_typ 0 [::] c.2 /\ 1420 | reducible (map some ctx) c.2 [::] c.1) 1421 | (zip (map var (iota 0 (size ctx))) ctx). 1422 | apply Forall_nth; case => {H t ty} t ty n; rewrite 1423 | size_zip size_map size_iota minnn nth_map' (nth_map' (@fst _ _)) /= 1424 | -/unzip1 -/unzip2 unzip1_zip ?unzip2_zip ?size_map ?size_iota // => H; 1425 | rewrite (nth_map 0 t var) ?size_iota // nth_iota // add0n. 1426 | suff: # [seq Some i | i <- ctx] \|- n \: subst_typ 0 [::] (nth ty ctx n) by 1427 | intuition apply (CR4' (@reducibility_isrc _ (nth ty ctx n) [::] I)). 1428 | by rewrite /typing /= subst_nil_ty nth_rcons size_map H (nth_map ty). 1429 | move/(_ H0) => {H0} /(rc_cr1 (reducibility_isrc _ _)) /= /(_ I). 1430 | have ->: typemap (subst_typ^~ [::]) 0 t = t by 1431 | elim: {H} t 0 => //=; move: subst_nil_ty; congruence'. 1432 | move: (@identical_substitution [::] (map some ctx) t ty 0 H). 1433 | rewrite subst_nil_ty /= size_map => ->. 1434 | move/(_ (ctxleq_preserves_typing _ H)). 1435 | by rewrite -cats1 => /(_ (ctxleq_appr _ _)). 1436 | Qed. 1437 | 1438 | End strong_normalization_proof_typed. 1439 | 1440 | Module strong_normalization_proof_kripke. 1441 | 1442 | Import subject_reduction_proof. 1443 | 1444 | Section CRs. 1445 | Variable (ty : typ) (P : context typ -> term -> Prop). 1446 | Definition CR_typed := forall ctx t, P ctx t -> ctx \|- t \: ty. 1447 | Definition CR0 := forall ctx ctx' t, ctx <=c ctx' -> P ctx t -> P ctx' t. 1448 | Definition CR1 := forall ctx t, P ctx t -> SN t. 1449 | Definition CR2 := forall ctx t t', t ->r1 t' -> P ctx t -> P ctx t'. 1450 | Definition CR3 := forall ctx t, 1451 | ctx \|- t \: ty -> neutral t -> (forall t', t ->r1 t' -> P ctx t') -> P ctx t. 1452 | Record RC : Prop := 1453 | reducibility_candidate { 1454 | rc_typed : CR_typed ; 1455 | rc_cr0 : CR0 ; 1456 | rc_cr1 : CR1 ; 1457 | rc_cr2 : CR2 ; 1458 | rc_cr3 : CR3 1459 | }. 1460 | End CRs. 1461 | 1462 | Lemma CR2' ty P ctx t t' : RC ty P -> t ->r t' -> P ctx t -> P ctx t'. 1463 | Proof. 1464 | move => H; elim/clos_refl_trans_1n_ind: t t' / => //= x y z H0 H1 H2 H3. 1465 | by apply H2, (rc_cr2 H H0). 1466 | Qed. 1467 | 1468 | Lemma CR4 ty P ctx t : RC ty P -> 1469 | ctx \|- t \: ty -> neutral t -> (forall t', ~ t ->r1 t') -> P ctx t. 1470 | Proof. by move/rc_cr3 => H H0 H1 H2; apply H => // t' /H2. Qed. 1471 | 1472 | Lemma CR4' ty P ctx (v : nat) : RC ty P -> ctxindex ctx v ty -> P ctx v. 1473 | Proof. move/CR4 => H H0; apply H => // t' H1; inversion H1. Qed. 1474 | 1475 | Notation SN' ty := (fun ctx t => ctx \|- t \: ty /\ SN t). 1476 | 1477 | Lemma snorm_isrc ty : RC ty (SN' ty). 1478 | Proof. 1479 | (constructor; move => /=) => 1480 | [ctx t [] // | ctx ctx' t H [H0 H1] | ctx t [] // | 1481 | ctx t t' H [H0 [H1]] | ctx t ]; intuition. 1482 | - by apply (ctxleq_preserves_typing H). 1483 | - by apply (subject_reduction H). 1484 | - by constructor => t' /H1 []. 1485 | Qed. 1486 | 1487 | Definition rcfun tyl tyr (P Q : context typ -> term -> Prop) ctx u : Prop := 1488 | ctx \|- u \: tyl ->t tyr /\ 1489 | forall ctx' v, ctx <=c ctx' -> P ctx' v -> Q ctx' (u @ v). 1490 | 1491 | Lemma rcfun_isrc tyl tyr P Q : 1492 | RC tyl P -> RC tyr Q -> RC (tyl ->t tyr) (rcfun tyl tyr P Q). 1493 | Proof. 1494 | move => HP HQ; rewrite /rcfun; constructor; move => /=; first tauto. 1495 | - move => ctx ctx' t H [H0 H1]; split; last move => ctx'' t' H2 H3. 1496 | + by apply (ctxleq_preserves_typing H). 1497 | + by apply (H1 ctx'' t') => //; apply ctxleq_trans with ctx'. 1498 | - move => ctx t [H] 1499 | /(_ _ _ (ctxleq_appr _ _) (CR4' HP (ctxindex_last ctx tyl))) /(rc_cr1 HQ). 1500 | by apply (acc_preservation (f := app^~_)); auto. 1501 | - move => ctx t t' H [H0 H1]; split; last move => ctx' t''. 1502 | + by apply (subject_reduction H). 1503 | + by move => /H1 {H1} H1 /H1 {H1}; apply (rc_cr2 HQ); constructor. 1504 | - move => ctx t H H0 H1; split => // ctx' t' H2 H3. 1505 | elim/Acc_ind: t' / {H3} (rc_cr1 HP H3) (H3) => t' _ IH H3. 1506 | apply (rc_cr3 HQ) => //=; first (apply/typing_appP; exists tyl). 1507 | + by apply (ctxleq_preserves_typing H2). 1508 | + by apply (rc_typed HP). 1509 | + move => t'' H4; move: H0; inversion H4; subst => //= _. 1510 | * by apply (proj2 (H1 _ H7) ctx'). 1511 | * by apply IH => //; apply (rc_cr2 HP H7). 1512 | Qed. 1513 | 1514 | Fixpoint reducible ty (preds : seq (typ * (context typ -> term -> Prop))) : 1515 | context typ -> term -> Prop := 1516 | match ty with 1517 | | tyvar v => nth (SN' (v - size preds)) (unzip2 preds) v 1518 | | tyfun tyl tyr => 1519 | let s := subst_typ 0 (unzip1 preds) in 1520 | rcfun (s tyl) (s tyr) (reducible tyl preds) (reducible tyr preds) 1521 | | tyabs ty => fun ctx t => 1522 | ctx \|- t \: tyabs (subst_typ 1 (unzip1 preds) ty) /\ 1523 | forall ty' P, RC ty' P -> reducible ty ((ty', P) :: preds) ctx (t @' ty') 1524 | end. 1525 | 1526 | Lemma reducibility_isrc ty preds : 1527 | Forall (fun p => RC p.1 p.2) preds -> 1528 | RC (subst_typ 0 (unzip1 preds) ty) (reducible ty preds). 1529 | Proof. 1530 | elim: ty preds => /= [n | tyl IHl tyr IHr | ty IH] preds. 1531 | - rewrite shift_zero_ty subn0. 1532 | elim: preds n => [| P preds IH []] /=. 1533 | + move => n _; rewrite !nth_nil /= !subn0; apply snorm_isrc. 1534 | + by case. 1535 | + by move => n [H] /(IH n) {IH}; rewrite !subSS. 1536 | - by move => H; apply rcfun_isrc; [apply IHl | apply IHr]. 1537 | - move => H; constructor. 1538 | + by move => /= ctx t []. 1539 | + move => /= ctx ctx' t H0 [H1 H2]; split; 1540 | first by move: H1; apply ctxleq_preserves_typing. 1541 | move => ty' P H3. 1542 | move: (rc_cr0 (IH ((_, _) :: _) (conj H3 H))) => /(_ ctx ctx' _ H0). 1543 | by apply; apply H2. 1544 | + move => /= ctx t [_] /(_ 0 _ (snorm_isrc _)) 1545 | /(rc_cr1 (IH ((_, _) :: _) (conj (snorm_isrc _) H))). 1546 | by apply (acc_preservation (f := uapp^~_)) => x y H0; constructor. 1547 | + move => /= ctx t t' H0 [H1 H2]; split; 1548 | first by move: H1; apply subject_reduction. 1549 | move => ty' P H3; move: (H2 _ _ H3). 1550 | by apply (rc_cr2 (IH ((_, _) :: _) (conj H3 H))); constructor. 1551 | + move => /= ctx t H0 H1 H2; split => // ty' P H3. 1552 | apply (rc_cr3 (IH ((_, _) :: _) (conj H3 H))) => //=. 1553 | + by apply/typing_uappP; 1554 | exists (subst_typ 1 (unzip1 preds) ty) => //; rewrite subst_app_ty. 1555 | * by move => t' H4; move: H0; inversion H4; subst => //= _; apply H2. 1556 | Qed. 1557 | 1558 | Lemma shift_reducibility c ty preds preds' ctx t : 1559 | c <= size preds -> 1560 | (reducible (shift_typ (size preds') c ty) 1561 | (insert preds' preds (tyvar 0, SN' 0) c) ctx t <-> 1562 | reducible ty preds ctx t). 1563 | Proof. 1564 | elim: ty c preds ctx t => [v | tyl IHl tyr IHr | ty IH] c preds ctx t H. 1565 | - rewrite /= /unzip2 map_insert nth_insert size_map size_insert; elimif_omega. 1566 | rewrite (_ : v - size preds = 0) //; ssromega. 1567 | - rewrite /= /rcfun /unzip1 map_insert -(size_map (@fst _ _)) /= 1568 | !subst_shift_cancel_ty4 ?size_map //. 1569 | by split; case => H0 H1; split => // ctx' t' H2 1570 | /(IHl c _ _ _ H) /(H1 _ _ H2) /(IHr c _ _ _ H). 1571 | - rewrite /= /unzip1 map_insert -(size_map (@fst _ _)) -add1n 1572 | subst_shift_cancel_ty4 ?size_map //. 1573 | by split; case => H0 H1; split => // ty' P H2; 1574 | apply (IH c.+1 ((ty', P) :: preds)) => //; apply H1. 1575 | Qed. 1576 | 1577 | Lemma subst_reducibility ty preds n tys ctx t : 1578 | n <= size preds -> 1579 | (reducible (subst_typ n tys ty) preds ctx t <-> 1580 | reducible ty 1581 | (insert [seq (subst_typ 0 (drop n (unzip1 preds)) ty, 1582 | reducible ty (drop n preds)) | ty <- tys] 1583 | preds (tyvar 0, SN' 0) n) 1584 | ctx t). 1585 | Proof. 1586 | elim: ty preds n tys ctx t => 1587 | [v | tyl IHl tyr IHr | ty IH] preds n tys ctx t. 1588 | - rewrite /= size_insert size_map /unzip2 map_insert /= -map_comp /comp /= 1589 | nth_insert size_map -/unzip2; elimif_omega. 1590 | + by move/maxn_idPr => ->; rewrite nth_default ?leq_addr //= addnC. 1591 | + move => H0. 1592 | rewrite (nth_map (tyvar (v - size tys))) //=. 1593 | move: (shift_reducibility (nth (tyvar (v - size tys)) tys v) 1594 | (take n preds) ctx t (leq0n (size (drop n preds)))). 1595 | by rewrite /insert take0 drop0 /= cat_take_drop size_takel. 1596 | + by move => H; rewrite (_ : v - size preds = 0) //; ssromega. 1597 | - move => H /=; rewrite /rcfun /= /unzip1 map_insert -!map_comp /comp /= 1598 | -!subst_subst_compose_ty ?size_map // -/unzip1 add0n. 1599 | by split; case => H1 H2; split => 1600 | // ctx' t' H3 /IHl => /(_ H) /(H2 _ _ H3) /IHr => /(_ H). 1601 | - move => H /=; rewrite /unzip1 map_insert -!map_comp /comp /= 1602 | -subst_subst_compose_ty ?size_map // add1n -/unzip1. 1603 | split; case => H0 H1; split => // ty' P /H1 {H0 H1}. 1604 | + by move/IH => /= /(_ H); rewrite /insert /= subSS. 1605 | + by move => H0; apply IH. 1606 | Qed. 1607 | 1608 | Lemma abs_reducibility tyl tyr preds ctx t : 1609 | ctx \|- abs (subst_typ 0 (unzip1 preds) tyl) t 1610 | \: subst_typ 0 (unzip1 preds) (tyl ->t tyr) -> 1611 | Forall (fun p => RC p.1 p.2) preds -> 1612 | (forall ctx' t', ctx <=c ctx' -> 1613 | reducible tyl preds ctx' t' -> 1614 | reducible tyr preds ctx' (subst_term 0 0 [:: t'] t)) -> 1615 | reducible (tyl ->t tyr) preds ctx (abs (subst_typ 0 (unzip1 preds) tyl) t). 1616 | Proof. 1617 | move => /= H H0. 1618 | move: (reducible tyl preds) (reducible tyr preds) 1619 | (reducibility_isrc tyl H0) (reducibility_isrc tyr H0) => 1620 | {H0} P Q HP HQ H0; split => // ctx' t' H1 H2. 1621 | have H3: SN t by 1622 | move: (rc_cr1 HQ (H0 _ _ H1 H2)); 1623 | apply acc_preservation => x y; apply subst_reduction1. 1624 | move: t t' H3 {H0 H2} (rc_cr1 HP H2) H H1 (H2) (H0 _ _ H1 H2). 1625 | refine (Acc_ind2 _) => t t' IHt IHt' H H0 H1 H2; apply (rc_cr3 HQ) => //=; 1626 | first (apply/typing_appP; exists (subst_typ 0 (unzip1 preds) tyl)). 1627 | - by move: H; apply ctxleq_preserves_typing. 1628 | - by apply (rc_typed HP). 1629 | - move => t'' H3; inversion H3; subst => //. 1630 | + inversion H7; subst; apply IHt; auto. 1631 | * by move: H; apply subject_reduction. 1632 | * by apply (rc_cr2 HQ (subst_reduction1 0 0 [:: t'] H8)). 1633 | + apply IHt' => //; first by apply (rc_cr2 HP H7). 1634 | move: H2; apply (CR2' HQ). 1635 | by apply (@subst_reduction t 0 0 [:: (t', t2')]) => /=; 1636 | intuition apply rtc_step. 1637 | Qed. 1638 | 1639 | Lemma uabs_reducibility ty preds ctx t : 1640 | ctx \|- uabs t \: subst_typ 0 (unzip1 preds) (tyabs ty) -> 1641 | Forall (fun p => RC p.1 p.2) preds -> 1642 | (forall v P, RC v P -> 1643 | reducible ty ((v, P) :: preds) ctx (typemap (subst_typ^~ [:: v]) 0 t)) -> 1644 | reducible (tyabs ty) preds ctx (uabs t). 1645 | Proof. 1646 | move => /= H H0 H1; split => // ty' P H2. 1647 | move: (reducible _ _) (@reducibility_isrc ty ((ty', P) :: preds)) 1648 | (H1 ty' P H2) => P' /= /(_ (conj H2 H0)) {H0 H1 H2} H0 H1. 1649 | have H2: SN t by 1650 | move: (rc_cr1 H0 H1); 1651 | apply acc_preservation => x y; apply substtyp_reduction1. 1652 | elim/Acc_ind': t / H2 H H1 => t H1 H H2; apply (rc_cr3 H0) => //=. 1653 | - by apply/typing_uappP; 1654 | exists (subst_typ 1 (unzip1 preds) ty) => //; rewrite subst_app_ty. 1655 | - move => t' H3; inversion H3; subst => //; inversion H7; subst; apply H1 => //. 1656 | + by apply (subject_reduction H7). 1657 | + apply (rc_cr2 H0 (substtyp_reduction1 _ _ H5) H2). 1658 | Qed. 1659 | 1660 | Lemma uapp_reducibility ty ty' preds ctx t : 1661 | Forall (fun p => RC p.1 p.2) preds -> 1662 | reducible (tyabs ty) preds ctx t -> 1663 | reducible (subst_typ 0 [:: ty'] ty) preds ctx 1664 | (t @' subst_typ 0 (unzip1 preds) ty'). 1665 | Proof. 1666 | move => /= H H0. 1667 | apply subst_reducibility => //=. 1668 | rewrite /insert take0 sub0n !drop0 /=. 1669 | by apply H0, reducibility_isrc. 1670 | Qed. 1671 | 1672 | Lemma reduce_lemma ctx ctx' preds t ty : 1673 | [seq Some c.2 | c <- ctx] \|- t \: ty -> 1674 | Forall (fun p => RC p.1 p.2) preds -> 1675 | Forall (fun c => reducible c.2 preds ctx' c.1) ctx -> 1676 | reducible ty preds ctx' 1677 | (subst_term 0 0 (unzip1 ctx) (typemap (subst_typ^~ (unzip1 preds)) 0 t)). 1678 | Proof. 1679 | have Hty: forall t ty ctx ctx' preds, 1680 | [seq Some c.2 | c <- ctx] \|- t \: ty -> 1681 | Forall (fun p => RC p.1 p.2) preds -> 1682 | Forall (fun c => reducible c.2 preds ctx' c.1) ctx -> 1683 | ctx' \|- 1684 | subst_term 0 0 (unzip1 ctx) (typemap (subst_typ^~ (unzip1 preds)) 0 t) \: 1685 | subst_typ 0 (unzip1 preds) ty. 1686 | clear => t ty ctx ctx' preds H H0 H1. 1687 | move: (fun t ty => @subject_subst0 t ty ctx' 1688 | [seq (c.1, subst_typ 0 (unzip1 preds) c.2) | c <- ctx]). 1689 | rewrite /unzip1 -!map_comp !/comp -/unzip1; apply => /=. 1690 | - by elim: ctx H1 {H} => //= [[]] /= t' ty' ctx IH [] 1691 | /(rc_typed (reducibility_isrc ty' H0)) ->. 1692 | - by move: (substtyp_preserves_typing 0 (unzip1 preds) H); 1693 | rewrite -map_comp; apply ctxleq_preserves_typing, ctxleq_appr. 1694 | elim: t ty ctx ctx' preds => /=. 1695 | - move => v ty ctx ctx' preds H H0 H1. 1696 | rewrite shifttyp_zero shift_zero subn0 size_map. 1697 | elim: ctx v H H1 => //=; first by case. 1698 | case => t ty' ctx IH [] //=. 1699 | + by case/eqP => <- []. 1700 | + by move => v H [H1 H2]; rewrite subSS; apply IH. 1701 | - move => tl IHtl tr IHtr ty ctx ctx' preds /typing_appP [tyl H H0] H1 H2. 1702 | by move: (IHtl (tyl ->t ty) ctx ctx' preds H H1 H2) => /= [_]; 1703 | apply => //; apply IHtr. 1704 | - move => tyl t IHt ty ctx ctx' preds /typing_absP [tyr -> H] H0 H1. 1705 | apply abs_reducibility => //; 1706 | first by apply (Hty (abs tyl t)) => //; rewrite typing_absE. 1707 | move => ctx'' t' H2 H3; rewrite subst_app //=. 1708 | apply (IHt tyr ((t', tyl) :: ctx)) => //=; split => //; move: H1. 1709 | by apply Forall_impl => {t' ty H H3} [[t' ty]] /=; 1710 | apply (rc_cr0 (reducibility_isrc ty H0) H2). 1711 | - move => t IHt tyr ty ctx ctx' preds /typing_uappP [tyl -> H] H0 H1. 1712 | by apply uapp_reducibility => //; apply IHt. 1713 | - move => t IHt ty ctx ctx' preds /typing_uabsP [ty' -> H] H0 H1. 1714 | apply uabs_reducibility => //; 1715 | first by apply (Hty (uabs t)) => //; rewrite typing_uabsE. 1716 | move => v P H2. 1717 | rewrite -(subst_substtyp_distr 0 [:: v]) // typemap_compose /=. 1718 | have /(typemap_eq 0 t) -> : forall i ty, 1719 | subst_typ (i + 0) [:: v] (subst_typ (i + 1) (unzip1 preds) ty) = 1720 | subst_typ i (unzip1 ((v, P) :: preds)) ty by 1721 | move => i ty''; rewrite addn0 addn1 subst_app_ty. 1722 | move: (IHt ty' 1723 | [seq (c.1, shift_typ 1 0 c.2) | c <- ctx] ctx' ((v, P) :: preds)). 1724 | rewrite /unzip1 -!map_comp /comp /=; apply => //=. 1725 | + by move: H; rewrite -map_comp /comp /=. 1726 | + apply Forall_map; move: H1; apply Forall_impl => [[t' ty'']] /=. 1727 | case (shift_reducibility ty'' [:: (v, P)] ctx' t' (leq0n (size preds))). 1728 | rewrite /insert take0 drop0 sub0n => _ /=; apply. 1729 | Qed. 1730 | 1731 | Theorem typed_term_is_snorm ctx t ty : ctx \|- t \: ty -> SN t. 1732 | Proof. 1733 | move => H. 1734 | have {H}: map some (map (odflt (tyvar 0)) ctx) \|- t \: ty by 1735 | move: H; apply ctxleq_preserves_typing; 1736 | elim: ctx {t ty} => // [[]] // ty ctx H; rewrite ctxleqE eqxx. 1737 | move: {ctx} (map _ ctx) => ctx. 1738 | have ->: map some ctx = 1739 | [seq Some c.2 | c <- zip (map var (iota 0 (size ctx))) ctx] by 1740 | elim: ctx 0 => //= {t ty} ty ctx IH n; rewrite -IH. 1741 | move => /reduce_lemma; move => /(_ (map some ctx) [::] I). 1742 | rewrite unzip1_zip ?size_map ?size_iota ?leqnn; last by []. 1743 | have H: Forall 1744 | (fun c => reducible c.2 [::] (map some ctx) c.1) 1745 | (zip (map var (iota 0 (size ctx))) ctx) by 1746 | apply Forall_nth; case => {t ty} t ty n; rewrite 1747 | size_zip size_map size_iota minnn nth_map' (nth_map' (@fst _ _)) /= 1748 | -/unzip1 -/unzip2 unzip1_zip ?unzip2_zip ?size_map ?size_iota // => H; 1749 | rewrite (nth_map 0 t var) ?size_iota // nth_iota // add0n; 1750 | apply (CR4' (@reducibility_isrc (nth ty ctx n) [::] I)) => //=; 1751 | rewrite subst_nil_ty (nth_map ty). 1752 | move/(_ H) => {H} /(rc_cr1 (reducibility_isrc _ _)) /= /(_ I). 1753 | set f := subst_term _ _ _; set g := typemap _ _. 1754 | apply (acc_preservation (f := f \o g)) => x y H. 1755 | by apply subst_reduction1, substtyp_reduction1. 1756 | Qed. 1757 | 1758 | End strong_normalization_proof_kripke. 1759 | -------------------------------------------------------------------------------- /coq/deBruijn/STLC.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. 2 | From LCAC Require Import Relations_ext seq_ext_base ssrnat_ext seq_ext. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | Inductive typ := tyvar of nat | tyfun of typ & typ. 9 | Inductive term := var of nat | app of term & term | abs of typ & term. 10 | 11 | Coercion tyvar : nat >-> typ. 12 | Coercion var : nat >-> term. 13 | 14 | Notation "ty ->t ty'" := (tyfun ty ty') (at level 55, right associativity). 15 | Notation "t @ t'" := (app t t') (at level 65, left associativity). 16 | 17 | Fixpoint eqtyp t1 t2 := 18 | match t1, t2 with 19 | | tyvar n, tyvar m => eqn n m 20 | | t1l ->t t1r, t2l ->t t2r => eqtyp t1l t2l && eqtyp t1r t2r 21 | | _, _ => false 22 | end. 23 | 24 | Lemma eqnP' : Equality.axiom eqn. 25 | Proof. 26 | elim => [| n IH] [| m] //=; try by constructor. 27 | by case: (IH m) => H; constructor; [f_equal | case]. 28 | Defined. 29 | 30 | Lemma eqtypP : Equality.axiom eqtyp. 31 | Proof. 32 | elim => [n | t1l IHl t1r IHr] [m | t2l t2r] /=; try by constructor. 33 | - by case: eqnP'; constructor; [f_equal | case]. 34 | - move: (IHl t2l) (IHr t2r) => [] Hl [] Hr; 35 | constructor; by [f_equal | case]. 36 | Defined. 37 | 38 | Canonical typ_eqMixin := EqMixin eqtypP. 39 | Canonical typ_eqType := Eval hnf in EqType typ typ_eqMixin. 40 | 41 | Fixpoint eqterm t1 t2 := 42 | match t1, t2 with 43 | | var n, var m => n == m 44 | | t1l @ t1r, t2l @ t2r => eqterm t1l t2l && eqterm t1r t2r 45 | | abs ty1 t1, abs ty2 t2 => (ty1 == ty2) && eqterm t1 t2 46 | | _, _ => false 47 | end. 48 | 49 | Lemma eqtermP : Equality.axiom eqterm. 50 | Proof. 51 | move => t1 t2; apply: (iffP idP) => [| <-]; 52 | last by elim: t1 => //= [t1l -> | ty1 t1 ->] //; rewrite andbT. 53 | by elim: t1 t2 => [n | t1l IHl t1r IHr | ty1 t1 IH] 54 | [// m /eqP -> | //= t2l t2r /andP [] /IHl -> /IHr -> | 55 | //= ty2 t2 /andP [] /eqP -> /IH ->]. 56 | Defined. 57 | 58 | Canonical term_eqMixin := EqMixin eqtermP. 59 | Canonical term_eqType := Eval hnf in EqType term term_eqMixin. 60 | 61 | Fixpoint shift d c t : term := 62 | match t with 63 | | var n => var (if c <= n then n + d else n) 64 | | t1 @ t2 => shift d c t1 @ shift d c t2 65 | | abs ty t1 => abs ty (shift d c.+1 t1) 66 | end. 67 | 68 | Notation substitutev ts m n := 69 | (shift n 0 (nth (var (m - n - size ts)) ts (m - n))) (only parsing). 70 | 71 | Fixpoint substitute n ts t : term := 72 | match t with 73 | | var m => if n <= m then substitutev ts m n else m 74 | | t1 @ t2 => substitute n ts t1 @ substitute n ts t2 75 | | abs ty t' => abs ty (substitute n.+1 ts t') 76 | end. 77 | 78 | Reserved Notation "t ~>b1 t'" (at level 70, no associativity). 79 | 80 | Inductive betared1 : relation term := 81 | | betared1beta ty t1 t2 : abs ty t1 @ t2 ~>b1 substitute 0 [:: t2] t1 82 | | betared1appl t1 t1' t2 : t1 ~>b1 t1' -> t1 @ t2 ~>b1 t1' @ t2 83 | | betared1appr t1 t2 t2' : t2 ~>b1 t2' -> t1 @ t2 ~>b1 t1 @ t2' 84 | | betared1abs ty t t' : t ~>b1 t' -> abs ty t ~>b1 abs ty t' 85 | where "t ~>b1 t'" := (betared1 t t'). 86 | 87 | Notation betared := [* betared1]. 88 | Notation betaeq := [= betared1]. 89 | Infix "~>b" := betared (at level 70, no associativity). 90 | Infix "==b" := betaeq (at level 70, no associativity). 91 | 92 | Hint Constructors betared1. 93 | 94 | Fixpoint typing_rec (ctx : context typ) (t : term) : option typ := 95 | match t with 96 | | var n => nth None ctx n 97 | | tl @ tr => 98 | if typing_rec ctx tl is Some (tyl ->t tyr) 99 | then (if typing_rec ctx tr == Some tyl then Some tyr else None) 100 | else None 101 | | abs ty t => omap (tyfun ty) (typing_rec (Some ty :: ctx) t) 102 | end. 103 | 104 | Definition typing := nosimpl typing_rec. 105 | 106 | Notation "ctx \|- t \: ty" := (Some ty == typing ctx t) 107 | (at level 69, no associativity). 108 | Notation "ctx \|- t \: ty" := (Some (ty : typ) == typing ctx t) 109 | (at level 69, no associativity, only parsing). 110 | 111 | Lemma typing_varE ctx (n : nat) ty : ctx \|- n \: ty = ctxindex ctx n ty. 112 | Proof. by rewrite /typing /=. Qed. 113 | 114 | Lemma typing_appP ctx t1 t2 ty : 115 | reflect (exists2 tyl, ctx \|- t1 \: tyl ->t ty & ctx \|- t2 \: tyl) 116 | (ctx \|- t1 @ t2 \: ty). 117 | Proof. 118 | apply: (iffP idP); rewrite /typing /=. 119 | - by move: (typing_rec ctx t1) => [] // [] // tyl tyr; 120 | case: ifP => // /eqP -> /eqP [] ->; exists tyl. 121 | - by case => tyl /eqP <- /eqP <-; rewrite eqxx. 122 | Qed. 123 | 124 | Lemma typing_absP ctx t tyl ty : 125 | reflect (exists2 tyr, ty = tyl ->t tyr & Some tyl :: ctx \|- t \: tyr) 126 | (ctx \|- abs tyl t \: ty). 127 | Proof. 128 | apply: (iffP idP); rewrite /typing /=. 129 | - by case: typing_rec => //= tyr /eqP [] ->; exists tyr. 130 | - by case => tyr ->; case: typing_rec => // tyr' /eqP [] <-. 131 | Qed. 132 | 133 | Lemma typing_absE ctx t tyl tyr : 134 | ctx \|- abs tyl t \: tyl ->t tyr = Some tyl :: ctx \|- t \: tyr. 135 | Proof. 136 | by rewrite /typing /=; case: typing_rec => //= tyr'; 137 | rewrite /eq_op /= /eq_op /= -/eq_op eqxx. 138 | Qed. 139 | 140 | Notation SN := (Acc (fun x y => betared1 y x)). 141 | 142 | Definition neutral t := if t is abs _ _ then false else true. 143 | 144 | Lemma shiftzero n t : shift 0 n t = t. 145 | Proof. by elim: t n; congruence' => m n; rewrite addn0 if_same. Qed. 146 | 147 | Lemma shift_add d d' c c' t : 148 | c <= c' <= c + d -> shift d' c' (shift d c t) = shift (d' + d) c t. 149 | Proof. case/andP; do 2 elimleq; elim: t c; congruence' => *; elimif_omega. Qed. 150 | 151 | Lemma shift_shift_distr d c d' c' t : 152 | c' <= c -> shift d' c' (shift d c t) = shift d (d' + c) (shift d' c' t). 153 | Proof. elimleq; elim: t c'; congruence' => *; elimif_omega. Qed. 154 | 155 | Lemma shift_subst_distr n d c ts t : 156 | c <= n -> shift d c (substitute n ts t) = substitute (d + n) ts (shift d c t). 157 | Proof. 158 | elimleq; elim: t c => /=; congruence' => v c; 159 | elimif_omega; rewrite shift_add; elimif_omega. 160 | Qed. 161 | 162 | Lemma subst_shift_distr n d c ts t : 163 | n <= c -> 164 | shift d c (substitute n ts t) = 165 | substitute n (map (shift d (c - n)) ts) (shift d (size ts + c) t). 166 | Proof. 167 | elimleq; elim: t n; congruence' => v n; elimif_omega. 168 | - rewrite !nth_default ?size_map /=; elimif_omega. 169 | - rewrite -shift_shift_distr // nth_map' /=. 170 | congr shift; apply nth_equal; rewrite size_map; elimif_omega. 171 | Qed. 172 | 173 | Lemma subst_shift_cancel n d c ts t : 174 | c <= n -> size ts + n <= d + c -> 175 | substitute n ts (shift d c t) = shift (d - size ts) c t. 176 | Proof. 177 | do 2 elimleq; elim: t c; congruence' => v c; 178 | elimif_omega; rewrite nth_default /=; elimif_omega. 179 | Qed. 180 | 181 | Lemma subst_subst_distr n m xs ys t : 182 | m <= n -> 183 | substitute n xs (substitute m ys t) = 184 | substitute m (map (substitute (n - m) xs) ys) (substitute (size ys + n) xs t). 185 | Proof. 186 | elimleq; elim: t m; congruence' => v m /=; elimif_omega. 187 | - rewrite (@subst_shift_cancel m) // size_map 1?nth_default /=; elimif_omega. 188 | - rewrite size_map -shift_subst_distr // nth_map' /=; 189 | congr shift; apply nth_equal; rewrite size_map; elimif_omega. 190 | Qed. 191 | 192 | Lemma subst_app n xs ys t : 193 | substitute n xs (substitute (size xs + n) ys t) = substitute n (xs ++ ys) t. 194 | Proof. 195 | elim: t n; congruence' => v n; rewrite nth_cat size_cat; 196 | elimif_omega; rewrite subst_shift_cancel; elimif_omega. 197 | Qed. 198 | 199 | Lemma subst_nil n t : substitute n [::] t = t. 200 | Proof. elim: t n; congruence' => v n; rewrite nth_nil /=; elimif_omega. Qed. 201 | 202 | Lemma subst_betared1 n ts t t' : 203 | t ~>b1 t' -> substitute n ts t ~>b1 substitute n ts t'. 204 | Proof. 205 | move => H; elim/betared1_ind: t t' / H n => /=; auto => t t' ty n. 206 | by rewrite subst_subst_distr //= add1n subn0. 207 | Qed. 208 | 209 | Module subject_reduction_proof. 210 | 211 | Lemma ctxleq_preserves_typing ctx1 ctx2 t ty : 212 | ctx1 <=c ctx2 -> ctx1 \|- t \: ty -> ctx2 \|- t \: ty. 213 | Proof. 214 | elim: t ty ctx1 ctx2 => /= [n | tl IHl tr IHr | tyl t IH] ty ctx1 ctx2. 215 | - move/ctxleqP; apply. 216 | - by move => H /typing_appP [tyl H0 H1]; apply/typing_appP; 217 | exists tyl; [apply (IHl _ ctx1) | apply (IHr _ ctx1)]. 218 | - by move => H /typing_absP [tyr -> H0]; rewrite typing_absE; 219 | move: H0; apply IH; rewrite ctxleqE eqxx. 220 | Qed. 221 | 222 | Lemma typing_shift t c ctx1 ctx2 : 223 | typing (ctxinsert ctx2 ctx1 c) (shift (size ctx2) c t) = typing ctx1 t. 224 | Proof. 225 | rewrite /typing. 226 | elim: t c ctx1 => /= [n | tl IHl tr IHr | tyl t IH] c ctx1. 227 | - by rewrite nth_insert; elimif_omega. 228 | - by rewrite IHl IHr. 229 | - by rewrite -(IH c.+1). 230 | Qed. 231 | 232 | Lemma subject_subst t ty n ctx ctx' : 233 | all (fun p => drop n ctx \|- p.1 \: p.2) ctx' -> 234 | ctxinsert [seq Some p.2 | p <- ctx'] ctx n \|- t \: ty -> 235 | ctx \|- substitute n (unzip1 ctx') t \: ty. 236 | Proof. 237 | elim: t ty n ctx => /= [m | tl IHl tr IHr | tyl t IH] ty n ctx H. 238 | - rewrite /typing /= nth_insert !size_map => /eqP ->; elimif. 239 | + apply/eqP/esym; rewrite nth_default ?size_map /=; elimif_omega. 240 | + rewrite -/typing !(nth_map (var 0, tyvar 0)) //. 241 | move: {H H0} (all_nthP (var 0, tyvar 0) H m H0). 242 | rewrite -(typing_shift _ 0 _ (ctxinsert [::] (take n ctx) n)) 243 | size_insert /= add0n size_take minnC minKn /insert take0 244 | sub0n take_minn minnn size_take minnE subKn ?leq_subr //= 245 | drop_take_nil drop0 cats0 -catA -{4}(cat_take_drop n ctx). 246 | apply ctxleq_preserves_typing; rewrite ctxleq_appl. 247 | by case: (leqP' n (size ctx)) => 248 | //= /ltnW /drop_oversize ->; rewrite cats0; 249 | apply/ctxleqP => /= n' ty' /eqP; rewrite nth_nseq if_same. 250 | - by case/typing_appP => tyl H0 H1; apply/typing_appP; exists tyl; auto. 251 | - by case/typing_absP => tyr -> H0; rewrite typing_absE; apply IH. 252 | Qed. 253 | 254 | Lemma subject_subst0 t ty ctx ctx' : 255 | all (fun p => ctx \|- p.1 \: p.2) ctx' -> 256 | [seq Some p.2 | p <- ctx'] ++ ctx \|- t \: ty -> 257 | ctx \|- substitute 0 (unzip1 ctx') t \: ty. 258 | Proof. 259 | by move: (@subject_subst t ty 0 ctx ctx'); rewrite /insert take0 sub0n drop0. 260 | Qed. 261 | 262 | Arguments subject_subst [t ty n ctx] _ _ _. 263 | Arguments subject_subst0 [t ty ctx] _ _ _. 264 | 265 | Theorem subject_reduction ctx t1 t2 ty : 266 | t1 ~>b1 t2 -> ctx \|- t1 \: ty -> ctx \|- t2 \: ty. 267 | Proof. 268 | move => H; elim/betared1_ind: t1 t2 / H ctx ty => /=. 269 | - move => ty t1 t2 ctx ty2 /typing_appP [tyl] /typing_absP [tyr [-> ->]] H H0. 270 | by apply (subject_subst0 [:: (t2, ty)]) => //=; rewrite H0. 271 | - by move => t1 t1' t2 H H0 ctx ty /typing_appP [tyl H1 H2]; 272 | apply/typing_appP; exists tyl; [apply H0 | apply H2]. 273 | - by move => t1 t2 t2' H H0 ctx ty /typing_appP [tyl H1 H2]; 274 | apply/typing_appP; exists tyl; [apply H1 | apply H0]. 275 | - by move => tyl t t' H H0 ctx ty /typing_absP [tyr -> H1]; 276 | rewrite typing_absE; apply H0. 277 | Qed. 278 | 279 | Hint Resolve ctxleq_preserves_typing subject_subst subject_reduction. 280 | 281 | End subject_reduction_proof. 282 | 283 | (******************************************************************************* 284 | Strong normalization proof with the type-free version of reducibility 285 | *******************************************************************************) 286 | 287 | Module strong_normalization_proof_typefree. 288 | 289 | Fixpoint reducible (ty : typ) (t : term) : Prop := 290 | match ty with 291 | | tyvar n => SN t 292 | | tyl ->t tyr => forall t', reducible tyl t' -> reducible tyr (t @ t') 293 | end. 294 | 295 | Lemma CR2 t t' ty : t ~>b1 t' -> reducible ty t -> reducible ty t'. 296 | Proof. 297 | elim: ty t t' => /= [_ | tyl IHl tyr IHr] t t' H. 298 | - by case; apply. 299 | - by move => H0 t'' H1; apply IHr with (t @ t''); auto. 300 | Qed. 301 | 302 | Hint Resolve CR2. 303 | 304 | Lemma CR1_and_CR3 ty : 305 | (forall t, reducible ty t -> SN t) /\ 306 | (forall t, neutral t -> 307 | (forall t', t ~>b1 t' -> reducible ty t') -> reducible ty t). 308 | Proof. 309 | elim: ty; first by []. 310 | move => /= tyl [IHl1 IHl2] tyr [IHr1 IHr2]; 311 | split => [t H | tl H H0 tr H1]. 312 | - suff: SN ([fun t => t @ 0] t) by apply acc_preservation; constructor. 313 | apply IHr1, IHr2 => // t' H0. 314 | apply (CR2 H0), H, IHl2 => // t'' H1; inversion H1. 315 | - elim/Acc_ind': tr / {H1} (IHl1 tr H1) (H1) => tr IH H1. 316 | apply IHr2 => // t' H2; move: H; inversion H2; subst => // _; eauto. 317 | Qed. 318 | 319 | Definition CR1 t ty := proj1 (CR1_and_CR3 ty) t. 320 | Definition CR3 t ty := proj2 (CR1_and_CR3 ty) t. 321 | 322 | Lemma abs_reducibility t tyl tyr : 323 | (forall t', reducible tyl t' -> reducible tyr (substitute 0 [:: t'] t)) -> 324 | reducible (tyl ->t tyr) (abs tyl t). 325 | Proof. 326 | move => /= H t' H0. 327 | have H1: SN t by 328 | move: (CR1 (H t' H0)); apply acc_preservation => x y; apply subst_betared1. 329 | move: t t' H1 {H0} (CR1 H0) H (H0); refine (Acc_ind2 _) => t t' IH IH' H H0. 330 | apply CR3 => // t'' H1; inversion H1; subst; eauto; inversion H5; subst. 331 | by apply IH; auto => t'' H2; apply (CR2 (subst_betared1 0 [:: t''] H6)), H. 332 | Qed. 333 | 334 | Lemma reduce_lemma ctx (ctx' : seq (term * typ)) t ty : 335 | [seq Some p.2 | p <- ctx'] ++ ctx \|- t \: ty -> 336 | Forall (fun p => reducible p.2 p.1) ctx' -> 337 | reducible ty (substitute 0 (unzip1 ctx') t). 338 | Proof. 339 | elim: t ty ctx ctx'. 340 | - move => /= n ty ctx ctx'. 341 | rewrite subn0 size_map shiftzero. 342 | elim: ctx' n => [| c' ctx' IH []] /=. 343 | + move => n _ _; rewrite nth_nil subn0. 344 | apply CR3 => //; auto => t' H0; inversion H0. 345 | + case/eqP => ->; tauto. 346 | + by move => n H [_ H0]; apply IH. 347 | - move => tl IHl tr IHr tyr ctx ctx' /typing_appP [tyl H H0] H1. 348 | by move: (IHl (tyl ->t tyr) ctx ctx') => /=; 349 | apply => //; apply IHr with ctx. 350 | - move => tyl t IH ty ctx ctx' /typing_absP [tyr -> H] H0. 351 | simpl substitute; apply abs_reducibility => t' H1. 352 | by rewrite subst_app /=; apply (IH tyr ctx ((t', tyl) :: ctx')). 353 | Qed. 354 | 355 | Theorem typed_term_is_sn ctx t ty : ctx \|- t \: ty -> SN t. 356 | Proof. 357 | by move/(@reduce_lemma ctx [::]) => /= /(_ I); rewrite subst_nil; apply CR1. 358 | Qed. 359 | 360 | End strong_normalization_proof_typefree. 361 | 362 | (******************************************************************************* 363 | Strong normalization proof with the Kripke-style reducibility 364 | *******************************************************************************) 365 | 366 | Module strong_normalization_proof_kripke. 367 | 368 | Import subject_reduction_proof. 369 | 370 | Fixpoint reducible (ctx : context typ) (ty : typ) (t : term) : Prop := 371 | match ty with 372 | | tyvar n => SN t 373 | | tyl ->t tyr => forall t' ctx', 374 | ctx <=c ctx' -> ctx' \|- t' \: tyl -> reducible ctx' tyl t' -> 375 | reducible ctx' tyr (t @ t') 376 | end. 377 | 378 | Lemma ctxleq_preserves_reducibility ctx ctx' t ty : 379 | ctx <=c ctx' -> reducible ctx ty t -> reducible ctx' ty t. 380 | Proof. 381 | case: ty => //= tyl tyr H H0 t1 ctx'' H1. 382 | apply H0; move: H H1; apply ctxleq_trans. 383 | Qed. 384 | 385 | Lemma CR2 ctx t t' ty : 386 | t ~>b1 t' -> reducible ctx ty t -> reducible ctx ty t'. 387 | Proof. 388 | elim: ty ctx t t'. 389 | - by move => /= _ _ t t' H []; apply. 390 | - move => /= tyl IHl tyr IHr ctx t1 t2 H H0 t3 ctx' H1 H2 H3. 391 | by apply IHr with (t1 @ t3); auto. 392 | Qed. 393 | 394 | Hint Resolve ctxleq_preserves_reducibility CR2. 395 | 396 | Lemma CR1_and_CR3 ty : 397 | (forall ctx t, ctx \|- t \: ty -> reducible ctx ty t -> SN t) /\ 398 | (forall ctx t, ctx \|- t \: ty -> neutral t -> 399 | (forall t', t ~>b1 t' -> reducible ctx ty t') -> reducible ctx ty t). 400 | Proof. 401 | elim: ty; first by []. 402 | move => /= tyl [IHl1 IHl2] tyr [IHr1 IHr2]; 403 | split => [ctx t H H0 | ctx tl H H0 H1 tr ctx' H2 H3 H4]. 404 | - set H1 := ctxindex_last ctx tyl. 405 | have H2: ctx ++ [:: Some tyl] \|- t @ size ctx \: tyr by 406 | apply/typing_appP; exists tyl; eauto. 407 | suff: SN ([fun t => t @ size ctx] t) by 408 | apply acc_preservation; constructor. 409 | apply (IHr1 _ _ H2), IHr2 => // t' H3. 410 | apply (CR2 H3), H0 => //=. 411 | apply IHl2 => // x H4; inversion H4. 412 | - have H5: SN tr by apply IHl1 with ctx'. 413 | elim/Acc_ind': tr / H5 H2 H3 H4 => tr IH H2 H3 H4. 414 | have H5: ctx' \|- tl @ tr \: tyr by apply/typing_appP; exists tyl; eauto. 415 | apply IHr2 => // t' H6; move: H0; inversion H6; subst => // _; eauto. 416 | Qed. 417 | 418 | Definition CR1 t ty := proj1 (CR1_and_CR3 ty) t. 419 | Definition CR3 t ty := proj2 (CR1_and_CR3 ty) t. 420 | 421 | Lemma abs_reducibility ctx t tyl tyr : 422 | ctx \|- abs tyl t \: tyl ->t tyr -> 423 | (forall t' ctx', ctx <=c ctx' -> ctx' \|- t' \: tyl -> 424 | reducible ctx' tyl t' -> reducible ctx' tyr (substitute 0 [:: t'] t)) -> 425 | reducible ctx (tyl ->t tyr) (abs tyl t). 426 | Proof. 427 | move => /= H H0 t' ctx' H1 H2 H3. 428 | have H4: ctx' \|- substitute 0 [:: t'] t \: tyr by 429 | apply (subject_subst0 [:: (t', tyl)]); rewrite /= ?H2 // -typing_absE; 430 | move: H; apply ctxleq_preserves_typing. 431 | have {H4} H4: SN t by 432 | move: (CR1 H4 (H0 t' ctx' H1 H2 H3)); 433 | apply acc_preservation => x y; apply subst_betared1. 434 | move: t t' {H2 H3} H4 (CR1 H2 H3) H H0 H1 (H2) (H3); 435 | refine (Acc_ind2 _) => t t' IH IH' H H0 H1 H2 H3; apply CR3 => //. 436 | - apply/typing_appP; exists tyl; last apply H2. 437 | by move: H; apply ctxleq_preserves_typing. 438 | - move => t'' H4. 439 | inversion H4; subst => {H4}; eauto; inversion H8; subst => {H8}. 440 | apply IH => //; eauto => t'' ctx'' H4 H5 H6. 441 | by apply (CR2 (subst_betared1 0 [:: t''] H7)), H0. 442 | Qed. 443 | 444 | Lemma reduce_lemma ctx (ctx' : seq (term * typ)) t ty : 445 | [seq Some p.2 | p <- ctx'] ++ ctx \|- t \: ty -> 446 | all (fun p => ctx \|- p.1 \: p.2) ctx' -> 447 | Forall (fun p => reducible ctx p.2 p.1) ctx' -> 448 | reducible ctx ty (substitute 0 (unzip1 ctx') t). 449 | Proof. 450 | elim: t ty ctx ctx'. 451 | - move => /= n ty ctx ctx'. 452 | rewrite subn0 size_map shiftzero. 453 | elim: ctx' n => [| c' ctx' IH []] /=. 454 | + move => n H _ _; rewrite nth_nil subn0. 455 | apply CR3 => //; auto => t' H0; inversion H0. 456 | + case/eqP => ->; tauto. 457 | + by move => n H /andP [_ H0] [_]; apply IH. 458 | - move => tl IHl tr IHr tyr ctx ctx' /typing_appP [tyl H H0] H1 H2. 459 | move: (IHl (tyl ->t tyr) ctx ctx') => /=; apply; auto. 460 | by apply subject_subst0. 461 | - move => tyl t IH ty ctx ctx' /typing_absP [tyr -> H] H0 H1. 462 | simpl substitute; apply abs_reducibility; 463 | first by apply (@subject_subst0 (abs tyl t)) => //; rewrite typing_absE. 464 | move => /= t2 ctx2 H2 H3 H4. 465 | rewrite subst_app /=. 466 | apply (IH tyr ctx2 ((t2, tyl) :: ctx')) => /=. 467 | + by move: H; rewrite -!typing_absE; 468 | apply ctxleq_preserves_typing; rewrite ctxleq_appl. 469 | + rewrite H3 /=; move: H0; apply sub_all => p; eauto. 470 | + intuition; move: H1; apply Forall_impl; eauto. 471 | Qed. 472 | 473 | Theorem typed_term_is_sn ctx t ty : ctx \|- t \: ty -> SN t. 474 | Proof. 475 | move => H; move: (@reduce_lemma ctx [::] _ _ H) => /= /(_ erefl I). 476 | by rewrite subst_nil; apply CR1. 477 | Qed. 478 | 479 | End strong_normalization_proof_kripke. 480 | 481 | (******************************************************************************* 482 | Strong normalization proof with the typed version of reducibility 483 | *******************************************************************************) 484 | 485 | Module strong_normalization_proof_typed. 486 | 487 | Import subject_reduction_proof. 488 | 489 | Fixpoint list_hyp ty : seq typ := 490 | if ty is tyl ->t tyr then tyl :: list_hyp tyl ++ list_hyp tyr else [::]. 491 | 492 | Fixpoint list_hyp' (ctx : context typ) t : seq typ := 493 | oapp list_hyp [::] (typing ctx t) ++ 494 | match t with 495 | | var n => [::] 496 | | app tl tr => list_hyp' ctx tl ++ list_hyp' ctx tr 497 | | abs ty t => list_hyp' (Some ty :: ctx) t 498 | end. 499 | 500 | Lemma list_hyp'E ctx t : 501 | list_hyp' ctx t = 502 | oapp list_hyp [::] (typing ctx t) ++ 503 | match t with 504 | | var n => [::] 505 | | app tl tr => list_hyp' ctx tl ++ list_hyp' ctx tr 506 | | abs ty t => list_hyp' (Some ty :: ctx) t 507 | end. 508 | Proof. by case: t. Qed. 509 | 510 | Lemma ctxleq_list_hyp' ctx ctx' t ty : 511 | ctx <=c ctx' -> ctx \|- t \: ty -> list_hyp' ctx' t = list_hyp' ctx t. 512 | Proof. 513 | elim: t ty ctx ctx' => //=. 514 | - move => n ty ctx ctx' H H0. 515 | by rewrite -(eqP (ctxleq_preserves_typing H H0)) -(eqP H0). 516 | - move => tl IHl tr IHr tyr ctx ctx' H H0. 517 | rewrite -(eqP (ctxleq_preserves_typing H H0)) -(eqP H0) /=; congr cat. 518 | case/typing_appP: H0 => tyl H0 H1. 519 | by rewrite (IHl (tyl ->t tyr) ctx ctx') // (IHr tyl ctx ctx'). 520 | - move => tyl t IH ty ctx ctx' H H0. 521 | rewrite -(eqP (ctxleq_preserves_typing H H0)) -(eqP H0) /=; congr cat. 522 | case/typing_absP: H0 => tyr _ H0. 523 | by rewrite (IH tyr (Some tyl :: ctx) (Some tyl :: ctx')) // ctxleqE eqxx. 524 | Qed. 525 | 526 | Fixpoint reducible (ctx : context typ) (ty : typ) (t : term) : Prop := 527 | match ty with 528 | | tyvar n => SN t 529 | | tyl ->t tyr => forall t', 530 | ctx \|- t' \: tyl -> reducible ctx tyl t' -> reducible ctx tyr (t @ t') 531 | end. 532 | 533 | Lemma CR2 ctx t t' ty : 534 | t ~>b1 t' -> reducible ctx ty t -> reducible ctx ty t'. 535 | Proof. 536 | elim: ty ctx t t'. 537 | - by move => /= _ _ t t' H []; apply. 538 | - move => /= tyl IHl tyr IHr ctx t1 t2 H H0 t3 H1 H2. 539 | by apply IHr with (t1 @ t3); auto. 540 | Qed. 541 | 542 | Hint Resolve CR2. 543 | 544 | Lemma CR1_and_CR3 ty : 545 | (forall (ctx : context typ) t, 546 | all (fun ty => Some ty \in ctx) (list_hyp ty) -> 547 | ctx \|- t \: ty -> reducible ctx ty t -> SN t) /\ 548 | (forall ctx t, 549 | all (fun ty => Some ty \in ctx) (list_hyp ty) -> 550 | ctx \|- t \: ty -> neutral t -> 551 | (forall t', t ~>b1 t' -> reducible ctx ty t') -> reducible ctx ty t). 552 | Proof. 553 | elim: ty; first by []. 554 | move => /= tyl [IHl1 IHl2] tyr [IHr1 IHr2]; split => ctx tl. 555 | - case/andP => /(nthP None) [x _] /esym /eqP H; 556 | rewrite all_cat; case/andP => H0 H1 H2 H3. 557 | have H4: ctx \|- tl @ x \: tyr by apply/typing_appP; exists tyl. 558 | suff: SN ([fun t => t @ x] tl) by apply acc_preservation; constructor. 559 | by apply (IHr1 ctx) => //; apply IHr2 => // t' H5; 560 | apply (CR2 H5); apply H3 => //; apply IHl2 => // y H6; inversion H6. 561 | - case/andP => _; rewrite all_cat; case/andP => H H0 H1 H2 H3 tr H4 H5. 562 | have H6: SN tr by apply (IHl1 ctx). 563 | elim/Acc_ind': tr / H6 H4 H5 => tr IH H4 H5. 564 | have H6: ctx \|- tl @ tr \: tyr by apply/typing_appP; exists tyl. 565 | apply IHr2 => // t' H7; move: H2; inversion H7; subst => // _; eauto. 566 | Qed. 567 | 568 | Definition CR1 t ty := proj1 (CR1_and_CR3 ty) t. 569 | Definition CR3 t ty := proj2 (CR1_and_CR3 ty) t. 570 | 571 | Lemma abs_reducibility (ctx : context typ) t tyl tyr : 572 | all (fun ty => Some ty \in ctx) (list_hyp (tyl ->t tyr)) -> 573 | ctx \|- abs tyl t \: tyl ->t tyr -> 574 | (forall t', ctx \|- t' \: tyl -> 575 | reducible ctx tyl t' -> reducible ctx tyr (substitute 0 [:: t'] t)) -> 576 | reducible ctx (tyl ->t tyr) (abs tyl t). 577 | Proof. 578 | move => /= /andP [] H; rewrite all_cat; case/andP => H0 H1 H2 H3 t' H4 H5. 579 | have H6: ctx \|- substitute 0 [:: t'] t \: tyr by 580 | apply (subject_subst0 [:: (t', tyl)]); rewrite /= ?H4 // -typing_absE. 581 | have {H6} H6: SN t by 582 | move: (CR1 H1 H6 (H3 t' H4 H5)); 583 | apply acc_preservation => x y; apply subst_betared1. 584 | move: t t' {H4 H5} H6 (CR1 H0 H4 H5) H2 H3 (H4) (H5); 585 | refine (Acc_ind2 _) => t t' IH IH' H2 H3 H4 H5; apply CR3 => //. 586 | - by apply/typing_appP; exists tyl. 587 | - move => t'' H6. 588 | inversion H6; subst => {H6}; eauto; inversion H10; subst => {H10}. 589 | apply IH => //; eauto => t'' ctx'' H6. 590 | by apply (CR2 (subst_betared1 0 [:: t''] H9)), H3. 591 | Qed. 592 | 593 | Lemma reduce_lemma ctx (ctx' : seq (term * typ)) t ty : 594 | [seq Some p.2 | p <- ctx'] ++ ctx \|- t \: ty -> 595 | all (fun ty => Some ty \in ctx) 596 | (list_hyp' ([seq Some p.2 | p <- ctx'] ++ ctx) t) -> 597 | all (fun p => ctx \|- p.1 \: p.2) ctx' -> 598 | Forall (fun p => reducible ctx p.2 p.1) ctx' -> 599 | reducible ctx ty (substitute 0 (unzip1 ctx') t). 600 | Proof. 601 | elim: t ty ctx ctx'. 602 | - move => /= n ty ctx ctx' H. 603 | rewrite -(eqP H) cats0 subn0 size_map shiftzero /=. 604 | elim: ctx' n H => [| c' ctx' IH []] /=. 605 | + move => n H H0 _ _; rewrite nth_nil subn0. 606 | apply CR3 => // t' H1; inversion H1. 607 | + case/eqP => ->; tauto. 608 | + by move => n H H0 /andP [_ H1] [_]; apply IH. 609 | - move => /= tl IHl tr IHr tyr ctx ctx' H. 610 | rewrite -(eqP H) /=; case/typing_appP: H => tyl H H0. 611 | rewrite !all_cat; case/and3P => H1 H2 H3 H4 H5. 612 | move: (IHl (tyl ->t tyr) ctx ctx') => /=; apply; auto. 613 | by apply subject_subst0. 614 | - move => /= tyl t IH ty ctx ctx' H. 615 | rewrite -(eqP H) /=; case/typing_absP: H => tyr -> H. 616 | rewrite all_cat; case/andP => H0 H1 H2 H3. 617 | simpl substitute; apply abs_reducibility => //; 618 | first by apply (@subject_subst0 (abs tyl t)) => //; rewrite typing_absE. 619 | move => /= t2 H4 H5; rewrite subst_app /=. 620 | by apply (IH tyr ctx ((t2, tyl) :: ctx')) => //=; rewrite H4. 621 | Qed. 622 | 623 | Theorem typed_term_is_sn ctx t ty : ctx \|- t \: ty -> SN t. 624 | Proof. 625 | move => H. 626 | have H0: ctx ++ map some (list_hyp' ([::] ++ ctx) t) \|- t \: ty by 627 | move: H; apply ctxleq_preserves_typing, ctxleq_appr. 628 | move: (@reduce_lemma _ [::] _ _ H0) => /=. 629 | rewrite (ctxleq_list_hyp' (ctxleq_appr ctx _) H). 630 | have ->: forall xs ys, all (fun x => Some x \in xs ++ map some ys) ys by 631 | move => A xs ys; apply/allP => x H1; 632 | rewrite mem_cat; apply/orP/or_intror/map_f. 633 | move/(_ erefl erefl I); rewrite subst_nil; apply CR1 => //. 634 | by apply/allP => /= ty' H1; rewrite list_hyp'E mem_cat; 635 | apply/orP/or_intror/map_f; rewrite mem_cat -(eqP H) H1. 636 | Qed. 637 | 638 | End strong_normalization_proof_typed. 639 | 640 | (******************************************************************************* 641 | Normal forms: beta normal form and long beta/eta normal form 642 | *******************************************************************************) 643 | 644 | Fixpoint beta_normal (nf : bool) (t : term) : bool := 645 | match t with 646 | | var m => true 647 | | tl @ tr => beta_normal false tl && beta_normal true tr 648 | | abs _ t => nf && beta_normal true t 649 | end. 650 | 651 | Lemma beta_normal_imply (nf nf' : bool) (t : term) : 652 | nf ==> nf' -> beta_normal nf t -> beta_normal nf' t. 653 | Proof. by case: t nf nf' => //= _ t [] []. Qed. 654 | 655 | Lemma beta_normal_shift (d c : nat) (nf : bool) (t : term) : 656 | beta_normal nf (shift d c t) = beta_normal nf t. 657 | Proof. 658 | by elim: t nf c => [| tl IHl tr IHr _ c | _ t IH []] //=; rewrite IHl IHr. 659 | Qed. 660 | 661 | Lemma beta_normalPn (nf : bool) (t : term) : 662 | reflect 663 | ((if t is abs _ _ then ~~ nf else false) \/ exists t', t ~>b1 t') 664 | (~~ beta_normal nf t). 665 | Proof. 666 | apply/(iffP idP); elim: t nf => [n | tl IHl tr IHr | ty t IH] nf //=. 667 | - rewrite negb_and => /orP [/IHl | /IHr] {IHl IHr} [] H; right. 668 | - by case: tl H => //= ty tl _; eexists; constructor. 669 | - by case: H => tl' H; exists (tl' @ tr); constructor. 670 | - by case: tr H. 671 | - by case: H => tr' H; exists (tl @ tr'); constructor. 672 | - case: nf => //=; last by left. 673 | case/IH; first by case t. 674 | by case => t' H; right; exists (abs ty t'); constructor. 675 | - by case => // -[] t H; inversion H. 676 | - case => // -[] t H; rewrite negb_and; inversion H; subst => //=. 677 | + by rewrite IHl //=; right; exists t1'. 678 | + by rewrite IHr ?orbT //=; right; exists t2'. 679 | - by case: nf => //= -[] // [t' H]; 680 | inversion H; subst; apply IH; right; exists t'0. 681 | Qed. 682 | 683 | Lemma beta_normalP (nf : bool) (t : term) : 684 | reflect 685 | ((if t is abs _ _ then nf else true) /\ forall t', ~ t ~>b1 t') 686 | (beta_normal nf t). 687 | Proof. 688 | apply/(iffP idP); 689 | last by rewrite -(negbK (beta_normal _ _)); 690 | case => H H0; apply/beta_normalPn => -[]; 691 | [case: t nf H H0 => //= k t [] | case => t' /H0]. 692 | rewrite -(negbK (beta_normal _ _)) => /beta_normalPn H; split. 693 | - by case: t nf H => //= k t [] //= []; left. 694 | - by move => t' H0; apply H; right; exists t'. 695 | Qed. 696 | 697 | Lemma beta_normal_equiv_rtc (t t' : term) : 698 | beta_normal true t' -> t ==b t' -> t ~>b t'. 699 | Proof. 700 | Abort. 701 | 702 | Lemma beta_normal_uniqueness (t t' : term) (ty : typ) : 703 | t ==b t' -> beta_normal true t -> beta_normal true t' -> t = t'. 704 | Proof. 705 | Abort. 706 | 707 | Lemma beta_normal_existence (t : term) : 708 | SN t -> exists2 t', t ~>b t' & beta_normal true t'. 709 | Proof. 710 | elim/Acc_ind': t / => t IH. 711 | case: (boolP (beta_normal true t)); first by move => H; exists t. 712 | case/beta_normalPn; first by case t. 713 | by case => t' H; case: (IH _ H) => t'' H0 H1; 714 | exists t'' => //;apply rt1n_trans with t'. 715 | Qed. 716 | -------------------------------------------------------------------------------- /coq/deBruijn/Untyped.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. 2 | From LCAC Require Import Relations_ext seq_ext_base ssrnat_ext seq_ext. 3 | Require FunInd. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Printing Implicit Defensive. 8 | 9 | Inductive term : Set := var of nat | app of term & term | abs of term. 10 | 11 | Coercion var : nat >-> term. 12 | 13 | Fixpoint eqterm t1 t2 := 14 | match t1, t2 with 15 | | var n, var m => n == m 16 | | app t1l t1r, app t2l t2r => eqterm t1l t2l && eqterm t1r t2r 17 | | abs t1, abs t2 => eqterm t1 t2 18 | | _, _ => false 19 | end. 20 | 21 | Lemma eqtermP : Equality.axiom eqterm. 22 | Proof. 23 | move => t1 t2; apply: (iffP idP) => [| <-]; last by elim: t1 => //= t1l ->. 24 | by elim: t1 t2 => [n | t1l IH t1r IH' | t1 IH] 25 | [// m /eqP -> | //= t2l t2r /andP [] /IH -> /IH' -> | // t2 /IH ->]. 26 | Defined. 27 | 28 | Canonical term_eqMixin := EqMixin eqtermP. 29 | Canonical term_eqType := Eval hnf in EqType term term_eqMixin. 30 | 31 | Fixpoint shift d c t : term := 32 | match t with 33 | | var n => var (if c <= n then n + d else n) 34 | | app t1 t2 => app (shift d c t1) (shift d c t2) 35 | | abs t1 => abs (shift d c.+1 t1) 36 | end. 37 | 38 | Notation substitutev ts m n := 39 | (shift n 0 (nth (var (m - n - size ts)) ts (m - n))) (only parsing). 40 | 41 | Fixpoint substitute n ts t : term := 42 | match t with 43 | | var m => if n <= m then substitutev ts m n else m 44 | | app t1 t2 => app (substitute n ts t1) (substitute n ts t2) 45 | | abs t' => abs (substitute n.+1 ts t') 46 | end. 47 | 48 | Reserved Notation "t ->b1 t'" (at level 70, no associativity). 49 | 50 | Inductive betared1 : relation term := 51 | | betared1beta t1 t2 : app (abs t1) t2 ->b1 substitute 0 [:: t2] t1 52 | | betared1appl t1 t1' t2 : t1 ->b1 t1' -> app t1 t2 ->b1 app t1' t2 53 | | betared1appr t1 t2 t2' : t2 ->b1 t2' -> app t1 t2 ->b1 app t1 t2' 54 | | betared1abs t t' : t ->b1 t' -> abs t ->b1 abs t' 55 | where "t ->b1 t'" := (betared1 t t'). 56 | 57 | Notation betared := [* betared1]. 58 | Infix "->b" := betared (at level 70, no associativity). 59 | 60 | Hint Constructors betared1. 61 | 62 | Lemma shiftzero n t : shift 0 n t = t. 63 | Proof. by elim: t n; congruence' => v n; rewrite addn0 if_same. Qed. 64 | 65 | Lemma shift_add d d' c c' t : 66 | c <= c' <= c + d -> shift d' c' (shift d c t) = shift (d' + d) c t. 67 | Proof. case/andP; do 2 elimleq; elim: t c; congruence' => *; elimif_omega. Qed. 68 | 69 | Lemma shift_shift_distr d c d' c' t : 70 | c' <= c -> shift d' c' (shift d c t) = shift d (d' + c) (shift d' c' t). 71 | Proof. elimleq; elim: t c'; congruence' => *; elimif_omega. Qed. 72 | 73 | Lemma shift_subst_distr n d c ts t : 74 | c <= n -> shift d c (substitute n ts t) = substitute (d + n) ts (shift d c t). 75 | Proof. 76 | by elimleq; elim: t c; congruence' => v c; elimif; 77 | rewrite shift_add //= add0n leq_addr. 78 | Qed. 79 | 80 | Lemma subst_shift_distr n d c ts t : 81 | n <= c -> 82 | shift d c (substitute n ts t) = 83 | substitute n (map (shift d (c - n)) ts) (shift d (size ts + c) t). 84 | Proof. 85 | elimleq; elim: t n; congruence' => v n; elimif. 86 | - rewrite !nth_default ?size_map /=; elimif_omega. 87 | - rewrite -shift_shift_distr // nth_map' /=; 88 | congr shift; apply nth_equal; rewrite size_map; elimif_omega. 89 | Qed. 90 | 91 | (* 92 | Lemma subst_shift_distr' n d c ts t : 93 | shift d (n + c) (substitute n ts t) = 94 | substitute n [seq shift d c i | i <- ts] (shift d (size ts + (n + c)) t). 95 | Proof. by rewrite subst_shift_distr ?addKn // leq_addr. Qed. 96 | *) 97 | 98 | Lemma subst_shift_cancel n d c ts t : 99 | c <= n -> size ts + n <= d + c -> 100 | substitute n ts (shift d c t) = shift (d - size ts) c t. 101 | Proof. 102 | do 2 elimleq; elim: t c; congruence' => v c; 103 | elimif; rewrite nth_default /=; elimif_omega. 104 | Qed. 105 | 106 | Lemma subst_subst_distr n m xs ys t : 107 | m <= n -> 108 | substitute n xs (substitute m ys t) = 109 | substitute m (map (substitute (n - m) xs) ys) (substitute (size ys + n) xs t). 110 | Proof. 111 | elimleq; elim: t m; congruence' => v m; elimif. 112 | - rewrite nth_default ?(@subst_shift_cancel m) // ?size_map /=; elimif_omega. 113 | - rewrite -shift_subst_distr // nth_map' /=; 114 | congr shift; apply nth_equal; rewrite size_map; elimif_omega. 115 | Qed. 116 | 117 | Lemma subst_app n xs ys t : 118 | substitute n xs (substitute (size xs + n) ys t) = substitute n (xs ++ ys) t. 119 | Proof. 120 | elim: t n; congruence' => v n; rewrite nth_cat size_cat; 121 | elimif_omega; rewrite subst_shift_cancel; elimif_omega. 122 | Qed. 123 | 124 | Lemma subst_nil n t : substitute n [::] t = t. 125 | Proof. elim: t n; congruence' => m n; rewrite nth_nil /=; elimif_omega. Qed. 126 | 127 | Lemma subst_betared1 n ts t t' : 128 | t ->b1 t' -> substitute n ts t ->b1 substitute n ts t'. 129 | Proof. 130 | move => H; elim/betared1_ind: t t' / H n => /=; auto => t t' n. 131 | by rewrite subst_subst_distr //= add1n subn0. 132 | Qed. 133 | 134 | (* small example for PPL2015 paper *) 135 | Lemma shift_betared t t' d c : t ->b1 t' -> shift d c t ->b1 shift d c t'. 136 | Proof. 137 | move => H; elim/betared1_ind: t t' / H d c => /=; auto => t1 t2 d c. 138 | rewrite subst_shift_distr //= add1n subn0; auto. 139 | Qed. 140 | 141 | Module confluence_proof. 142 | 143 | Reserved Notation "t ->bp t'" (at level 70, no associativity). 144 | 145 | Inductive parred : relation term := 146 | | parredvar n : var n ->bp var n 147 | | parredapp t1 t1' t2 t2' : 148 | t1 ->bp t1' -> t2 ->bp t2' -> app t1 t2 ->bp app t1' t2' 149 | | parredabs t t' : t ->bp t' -> abs t ->bp abs t' 150 | | parredbeta t1 t1' t2 t2' : 151 | t1 ->bp t1' -> t2 ->bp t2' -> app (abs t1) t2 ->bp substitute 0 [:: t2'] t1' 152 | where "t ->bp t'" := (parred t t'). 153 | 154 | Hint Constructors parred. 155 | 156 | Function reduce_all_redex t : term := 157 | match t with 158 | | var _ => t 159 | | app (abs t1) t2 => 160 | substitute 0 [:: reduce_all_redex t2] (reduce_all_redex t1) 161 | | app t1 t2 => app (reduce_all_redex t1) (reduce_all_redex t2) 162 | | abs t' => abs (reduce_all_redex t') 163 | end. 164 | 165 | Lemma parred_refl t : parred t t. 166 | Proof. elim: t; auto. Qed. 167 | 168 | Lemma betaredappl t1 t1' t2 : t1 ->b t1' -> app t1 t2 ->b app t1' t2. 169 | Proof. apply (rtc_map' (fun x y => @betared1appl x y t2)). Qed. 170 | 171 | Lemma betaredappr t1 t2 t2' : t2 ->b t2' -> app t1 t2 ->b app t1 t2'. 172 | Proof. apply (rtc_map' (@betared1appr t1)). Qed. 173 | 174 | Lemma betaredabs t t' : t ->b t' -> abs t ->b abs t'. 175 | Proof. apply (rtc_map' betared1abs). Qed. 176 | 177 | Hint Resolve parred_refl betaredappl betaredappr betaredabs. 178 | 179 | Lemma betared1_in_parred : inclusion betared1 parred. 180 | Proof. apply betared1_ind; auto. Qed. 181 | 182 | Lemma parred_in_betared : inclusion parred betared. 183 | Proof. 184 | apply parred_ind; auto => t1 t1' t2 t2' H H0 H1 H2. 185 | - apply rtc_trans with (app t1' t2); auto. 186 | - apply rtc_trans with (app (abs t1') t2); auto. 187 | apply rtc_trans with (app (abs t1') t2'); auto. 188 | by apply rtc_step. 189 | Qed. 190 | 191 | Lemma shift_parred t t' d c : t ->bp t' -> shift d c t ->bp shift d c t'. 192 | Proof. 193 | move => H; elim/parred_ind: t t' / H d c => //=; 194 | auto => t1 t1' t2 t2' H H0 H1 H2 d c. 195 | rewrite subst_shift_distr //= add1n subn0; auto. 196 | Qed. 197 | 198 | Lemma subst_parred n ps t t' : 199 | Forall (prod_curry parred) ps -> t ->bp t' -> 200 | substitute n [seq fst p | p <- ps] t ->bp 201 | substitute n [seq snd p | p <- ps] t'. 202 | Proof. 203 | move => H H0; elim/parred_ind: t t' / H0 n => /=; auto. 204 | - move => v n; elimif; rewrite !size_map; apply shift_parred. 205 | elim: ps v H => //= [[t t']] ps IH [| v] [] //= H H0. 206 | by rewrite subSS; apply IH. 207 | - move => t1 t1' t2 t2' H0 H1 H2 H3 n. 208 | by rewrite subst_subst_distr //= add1n subn0; auto. 209 | Qed. 210 | 211 | Lemma parred_all_lemma t t' : t ->bp t' -> t' ->bp reduce_all_redex t. 212 | Proof with auto. 213 | elim/reduce_all_redex_ind: {t}_ t'. 214 | - by move => t n H t' H0; inversion H0; subst. 215 | - move => _ t1 t2 _ H H0 t' H1; inversion H1; subst. 216 | + inversion H4; subst... 217 | + apply (@subst_parred 0 [:: (t2', reduce_all_redex t2)]) => /=... 218 | - move => _ t1 t2 _ H H0 H1 t' H2; inversion H2; subst => //... 219 | - move => _ t1 _ H t2 H0; inversion H0; subst... 220 | Qed. 221 | 222 | Lemma parred_confluent : confluent parred. 223 | Proof. 224 | by move => t1 t2 t3 H H0; exists (reduce_all_redex t1); apply parred_all_lemma. 225 | Qed. 226 | 227 | Theorem betared_confluent : confluent betared. 228 | Proof. 229 | apply (rtc_confluent' betared1_in_parred parred_in_betared parred_confluent). 230 | Qed. 231 | 232 | End confluence_proof. 233 | -------------------------------------------------------------------------------- /coq/lib/Relations_ext.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat. 2 | From Coq Require Import Relations Relation_Operators. 3 | Export Relations Relation_Operators. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Printing Implicit Defensive. 8 | 9 | Notation "[* R ]" := 10 | (clos_refl_trans_1n _ R) (at level 5, no associativity) : type_scope. 11 | Notation "[= R ]" := [* fun x y => R x y \/ R y x ] 12 | (at level 5, no associativity) : type_scope. 13 | 14 | Notation inclusion R R' := (forall t1 t2, R t1 t2 -> R' t1 t2). 15 | Notation same_relation R R' := (forall t1 t2, R t1 t2 <-> R' t1 t2). 16 | Notation wconfluent R := 17 | (forall t1 t2 t3, R t1 t2 -> R t1 t3 -> 18 | exists2 t4, [*R%type] t2 t4 & [*R%type] t3 t4). 19 | Notation confluent R := 20 | (forall t1 t2 t3, R t1 t2 -> R t1 t3 -> 21 | exists2 t4, R%type t2 t4 & R%type t3 t4). 22 | Notation church_rosser R := 23 | (forall t1 t2, [= R] t1 t2 -> 24 | exists2 t3, [*R%type] t1 t3 & [*R%type] t2 t3). 25 | 26 | Hint Resolve rt1n_refl. 27 | 28 | Lemma rtc_trans A R : transitive A [*R]. 29 | Proof. by move => t1 t2 t3; rewrite -!clos_rt_rt1n_iff; apply rt_trans. Qed. 30 | 31 | Lemma rtc_step A (R : relation A) : inclusion R [*R]. 32 | Proof. by move => x y H; apply rt1n_trans with y. Qed. 33 | 34 | Lemma rtc_map A (R R' : relation A) : inclusion R R' -> inclusion [*R] [*R']. 35 | Proof. 36 | move => H t t'; elim/clos_refl_trans_1n_ind: t t' / => // t1 t2 T3 H0 H1 H2. 37 | apply rt1n_trans with t2; auto. 38 | Qed. 39 | 40 | Lemma rtc_map' A B (R : relation A) (R' : relation B) (f : A -> B) : 41 | inclusion R (fun x y => R' (f x) (f y)) -> 42 | inclusion [*R] (fun x y => [*R'] (f x) (f y)). 43 | Proof. 44 | move => H t t'; elim/clos_refl_trans_1n_ind: t t' / => //= x y z H0 H1 H2. 45 | apply rt1n_trans with (f y); auto. 46 | Qed. 47 | 48 | Lemma rtc_nest_elim A (R : relation A) : same_relation [*[*R]] [*R]. 49 | Proof. 50 | move => t1 t2; split. 51 | - by elim/clos_refl_trans_1n_ind: t1 t2 / => // t1 t2 t3 H _ H0; 52 | apply rtc_trans with t2. 53 | - apply clos_rt1n_step. 54 | Qed. 55 | 56 | Lemma rtc_semi_confluent' A (R R' : relation A) : 57 | (forall t1 t2 t3, R t1 t2 -> R' t1 t3 -> exists2 t4, R' t2 t4 & [*R] t3 t4) -> 58 | (forall t1 t2 t3, [*R] t1 t2 -> R' t1 t3 -> 59 | exists2 t4, R' t2 t4 & [*R] t3 t4). 60 | Proof. 61 | move => H t1 t2 t3 H0; elim/clos_refl_trans_1n_ind: t1 t2 / H0 t3. 62 | - by move => t1 t3 H0; exists t3. 63 | - move => t1 t1' t2 H0 H1 IH t3 H2. 64 | case: (H t1 t1' t3 H0 H2) => {H H0 H2} t3' H H0. 65 | case: (IH t3' H) => {IH H} t4 H H2. 66 | by exists t4 => //; apply rtc_trans with t3'. 67 | Qed. 68 | 69 | Lemma rtc_semi_confluent A (R R' : relation A) : 70 | (forall t1 t2 t3, R t1 t2 -> R' t1 t3 -> exists2 t4, R' t2 t4 & R t3 t4) -> 71 | (forall t1 t2 t3, [*R] t1 t2 -> R' t1 t3 -> 72 | exists2 t4, R' t2 t4 & [*R] t3 t4). 73 | Proof. 74 | move => H; apply rtc_semi_confluent' => t1 t2 t3 H0 H1. 75 | by case: (H _ _ _ H0 H1) => t4 H2 H3; exists t4 => //; apply rtc_step. 76 | Qed. 77 | 78 | Lemma rtc_confluent A (R : relation A) : confluent R -> confluent [*R]. 79 | Proof. 80 | move => H; apply rtc_semi_confluent => t1 t2 t3 H0 H1. 81 | by case: (rtc_semi_confluent H H1 H0) => t4 H2 H3; exists t4. 82 | Qed. 83 | 84 | Lemma rtc_confluent' A (R R' : relation A) : 85 | inclusion R R' -> inclusion R' [*R] -> confluent R' -> confluent [*R]. 86 | Proof. 87 | move => H H0 H1 t1 t2 t3 H2 H3. 88 | case: (rtc_confluent H1 (rtc_map H H2) (rtc_map H H3)) => t4 H4 H5. 89 | by exists t4; apply rtc_nest_elim; apply (rtc_map H0). 90 | Qed. 91 | 92 | Lemma equiv_refl A (R : relation A) : reflexive A [= R]. 93 | Proof. by []. Qed. 94 | 95 | Lemma equiv_trans A (R : relation A) : transitive A [= R]. 96 | Proof. apply rtc_trans. Qed. 97 | 98 | Lemma equiv_sym A (R : relation A) : symmetric A [= R]. 99 | Proof. 100 | move => x y; elim/clos_refl_trans_1n_ind: x y / => //= x y z H _ H0. 101 | by apply equiv_trans with y => //; apply rtc_step; case: H; auto. 102 | Qed. 103 | 104 | Lemma equiv_includes_rtc A (R : relation A) : inclusion [* R] [= R]. 105 | Proof. 106 | move => x y; elim/clos_refl_trans_1n_ind: x y / => //= x y z H _ H0. 107 | by apply rt1n_trans with y => //; left. 108 | Qed. 109 | 110 | Lemma confluent_CR A (R : relation A) : confluent [* R] <-> church_rosser R. 111 | Proof. 112 | split; last by move => H x y z Hy Hz; apply H, equiv_trans with x; 113 | [apply equiv_sym, equiv_includes_rtc | apply equiv_includes_rtc]. 114 | move => H x y; elim/clos_refl_trans_1n_ind: x y /; first by move => x; exists x. 115 | move => x y z [] H0 H1 [c H2 H3]. 116 | - by exists c => //; apply rt1n_trans with y. 117 | - by case: (H _ _ _ (rtc_step H0) H2) => c' H4 H5; 118 | exists c' => //; apply rtc_trans with c. 119 | Qed. 120 | 121 | Section Z. 122 | 123 | Variable (A : Type) (R : relation A) (comp : A -> A). 124 | Variable (Z : forall x y, R x y -> [*R] y (comp x) /\ [*R] (comp x) (comp y)). 125 | 126 | Lemma Z_confluent : confluent [*R]. 127 | Proof. 128 | have Z1 x y : R x y -> [*R] y (comp x) by firstorder. 129 | have Z2 x y : R x y -> [*R] (comp x) (comp y) by firstorder. 130 | apply rtc_semi_confluent' => t1 t2 t3 H H0; exists (comp t3). 131 | - apply rtc_trans with (comp t1); first apply Z1 => //. 132 | apply rtc_nest_elim; move: t1 t3 H0 {H}; apply rtc_map', Z2. 133 | - case: H0. 134 | + by apply rt1n_trans with t2 => //; apply Z1. 135 | + move => {t3} t1' t3 H0 H1; 136 | elim/clos_refl_trans_1n_ind: t1' t3 / H1 t1 H0 {H}; eauto => t1' t1 H. 137 | by apply rtc_trans with (comp t1); [apply Z1 | apply Z2]. 138 | Qed. 139 | 140 | End Z. 141 | 142 | Lemma rtc_preservation A (P : A -> Prop) (R : relation A) : 143 | (forall x y, R x y -> P x -> P y) -> forall x y, [*R] x y -> P x -> P y. 144 | Proof. 145 | move => H x y; elim/clos_refl_trans_1n_ind: x y / => //= x y z H0 _ H1 H2. 146 | exact (H1 (H x y H0 H2)). 147 | Qed. 148 | 149 | Lemma Acc_ind' (A : Type) (R : relation A) (P : A -> Prop) : 150 | (forall x, (forall y, R y x -> P y) -> P x) -> forall x, Acc R x -> P x. 151 | Proof. move => H x H0; elim/Acc_ind: x / H0 H => x _ IH H; apply H; auto. Qed. 152 | 153 | Lemma acc_preservation A B (RA : relation A) (RB : relation B) (f : A -> B) a : 154 | (forall x y, RA x y -> RB (f x) (f y)) -> Acc RB (f a) -> Acc RA a. 155 | Proof. 156 | move => H H0; move: {1 2}(f a) H0 (erefl (f a)) => b H0. 157 | elim/Acc_ind': b / H0 a => b H0 a H1; subst b. 158 | by constructor => a' H1; apply (fun x => H0 _ x _ erefl), H. 159 | Qed. 160 | 161 | Lemma Acc_ind2 162 | (A B : Type) (RA : relation A) (RB : relation B) (P : A -> B -> Prop) : 163 | (forall x y, (forall x', RA x' x -> P x' y) -> 164 | (forall y', RB y' y -> P x y') -> P x y) -> 165 | forall x y, Acc RA x -> Acc RB y -> P x y. 166 | Proof. 167 | move => H x y Hx Hy; 168 | elim/Acc_ind: x / Hx y Hy => x Hx IHx y Hy; elim/Acc_ind: y / Hy => y Hy IHy. 169 | by apply H => [x' | y'] H2; [apply IHx | apply IHy]. 170 | Qed. 171 | 172 | Lemma Acc_wconfluent A (R : relation A) : 173 | (forall x, Acc (fun a => R^~ a) x) -> wconfluent R -> confluent [*R]. 174 | Proof. 175 | move => Hacc Hwconfluent x; elim/Acc_ind': x / (Hacc x) => x IH y z Hy Hz. 176 | inversion Hy; subst; first by exists z. 177 | inversion Hz; subst; first by exists y. 178 | move: {Hy Hz} y0 y1 H H0 H1 H2 => y' z' Hy Hy0 Hz Hz0. 179 | case: (Hwconfluent _ _ _ Hy Hz) => x' Hy1 Hz1. 180 | case: (IH _ Hy _ _ Hy0 Hy1) => y'' Hy2 Hy3. 181 | case: (IH _ Hz _ _ Hz0 Hz1) => z'' Hz2 Hz3. 182 | case: (IH _ Hy _ _ (rtc_trans Hy0 Hy2) (rtc_trans Hy1 Hz3)) => x'' Hy4 Hz4. 183 | by exists x''; [apply (rtc_trans Hy2 Hy4) | apply (rtc_trans Hz2 Hz4)]. 184 | Qed. 185 | -------------------------------------------------------------------------------- /coq/lib/seq_ext.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. 2 | From LCAC Require Import seq_ext_base ssrnat_ext. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | (* insert *) 9 | 10 | Definition insert A xs ys d n : seq A := 11 | take n ys ++ nseq (n - size ys) d ++ xs ++ drop n ys. 12 | 13 | Section Insert. 14 | 15 | Variable (A B : Type). 16 | 17 | Lemma size_insert (xs ys : seq A) d n : 18 | size (insert xs ys d n) = size xs + maxn n (size ys). 19 | Proof. rewrite /insert !size_cat size_nseq size_take size_drop; ssromega. Qed. 20 | 21 | Lemma map_insert (f : A -> B) xs ys d n : 22 | map f (insert xs ys d n) = insert (map f xs) (map f ys) (f d) n. 23 | Proof. by rewrite /insert !map_cat map_take map_nseq size_map map_drop. Qed. 24 | 25 | Lemma nth_insert (xs ys : seq A) d d' n m : 26 | nth d (insert xs ys d' n) m = 27 | if m < n then nth d' ys m else 28 | if m < n + size xs then nth d' xs (m - n) else nth d ys (m - size xs). 29 | Proof. 30 | rewrite /insert !nth_cat size_take size_nseq -subnDA nth_drop nth_take'. 31 | have ->: minn n (size ys) + (n - size ys) = n by ssromega. 32 | elimif_omega; rewrite nth_nseq H0 nth_default //; ssromega. 33 | Qed. 34 | 35 | Lemma cons_insert n y (xs ys : seq A) d : 36 | y :: (insert xs ys d n) = insert xs (y :: ys) d n.+1. 37 | Proof. by rewrite /insert /= subSS. Qed. 38 | 39 | Lemma take_insert n (xs ys : seq A) d : 40 | take n (insert xs ys d n) = take n ys ++ nseq (n - size ys) d. 41 | Proof. 42 | by rewrite /insert take_cat size_take ltnNge geq_minl /= minnE subKn 43 | ?leq_subr // take_cat size_nseq ltnn subnn take0 cats0. 44 | Qed. 45 | 46 | Lemma drop_insert n (xs ys : seq A) d : 47 | drop (n + size xs) (insert xs ys d n) = drop n ys. 48 | Proof. 49 | rewrite /insert !catA drop_cat !size_cat size_take size_nseq drop_addn. 50 | elimif_omega. 51 | Qed. 52 | 53 | End Insert. 54 | 55 | (* context *) 56 | 57 | Definition context A := (seq (option A)). 58 | 59 | Notation ctxnth := (nth None). 60 | Notation ctxindex xs n x := (Some x == ctxnth xs n). 61 | Notation ctxmap f xs := (map (omap f) xs). 62 | Notation ctxinsert xs ys n := (insert xs ys None n). 63 | 64 | Section Context. 65 | 66 | Variable (A : eqType). 67 | 68 | Fixpoint ctxleq_rec (xs ys : context A) : bool := 69 | match xs, ys with 70 | | [::], _ => true 71 | | (None :: xs), [::] => ctxleq_rec xs [::] 72 | | (None :: xs), (_ :: ys) => ctxleq_rec xs ys 73 | | (Some _ :: _), [::] => false 74 | | (Some x :: xs), (Some y :: ys) => (x == y) && ctxleq_rec xs ys 75 | | (Some _ :: _), (None :: _) => false 76 | end. 77 | 78 | Definition ctxleq := nosimpl ctxleq_rec. 79 | 80 | Infix "<=c" := ctxleq (at level 70, no associativity). 81 | 82 | Lemma ctxleqE xs ys : 83 | (xs <=c ys) = 84 | ((head None xs == head None ys) || (head None xs == None)) && 85 | (behead xs <=c behead ys). 86 | Proof. 87 | move: xs ys => [| [x |] xs] [| [y |] ys] //=. 88 | by case/boolP: (Some x == None) => // _; rewrite orbF. 89 | Qed. 90 | 91 | Lemma ctxleqP (xs ys : context A) : 92 | reflect (forall n a, ctxindex xs n a -> ctxindex ys n a) (xs <=c ys). 93 | Proof. 94 | apply: (iffP idP); elim: xs ys => [| x xs IH]. 95 | - by move => ys _ n a; rewrite nth_nil. 96 | - by case => [| y ys]; rewrite ctxleqE /= => 97 | /andP [] /orP [] /eqP -> H [] //= n a /(IH _ H) //; rewrite nth_nil. 98 | - by move => ys; rewrite ctxleqE eqxx orbT. 99 | - by case => [| y ys]; rewrite ctxleqE /= => H; apply/andP; 100 | (split; 101 | [ case: x H; rewrite ?eqxx ?orbT // => x H; rewrite (H 0 x) | 102 | apply IH => n a /(H n.+1 a) ]). 103 | Qed. 104 | 105 | Lemma ctxleqxx (xs : context A) : xs <=c xs. 106 | Proof. by apply/ctxleqP. Qed. 107 | 108 | Lemma ctxleq_trans (xs ys zs : context A) : 109 | xs <=c ys -> ys <=c zs -> xs <=c zs. 110 | Proof. do 2 move/ctxleqP => ?; apply/ctxleqP; auto. Qed. 111 | 112 | Lemma ctxleq_app (xs xs' ys ys' : context A) : 113 | size xs = size xs' -> 114 | (xs ++ ys) <=c (xs' ++ ys') = (xs <=c xs') && (ys <=c ys'). 115 | Proof. 116 | elim: xs xs' => [| x xs IH] [] //= x' xs' []. 117 | by rewrite ctxleqE /=; move/IH => ->; apply esym; rewrite ctxleqE /= andbA. 118 | Qed. 119 | 120 | Lemma ctxleq_appl (xs ys zs : context A) : 121 | (xs ++ ys <=c xs ++ zs) = (ys <=c zs). 122 | Proof. by rewrite ctxleq_app // ctxleqxx. Qed. 123 | 124 | Lemma ctxleq_appr (xs ys : context A) : xs <=c (xs ++ ys). 125 | Proof. by rewrite -{1}(cats0 xs) ctxleq_appl. Qed. 126 | 127 | Lemma ctxindex_last ctx (x : A) : ctxindex (ctx ++ [:: Some x]) (size ctx) x. 128 | Proof. by rewrite nth_cat ltnn subnn. Qed. 129 | 130 | Lemma ctxinsert_leq n (xs ys zs : context A) : 131 | ctxinsert xs ys n <=c zs -> 132 | exists xs' ys', [&& zs <=c ctxinsert xs' ys' n, ctxinsert xs' ys' n <=c zs, 133 | xs <=c xs' & ys <=c ys']. 134 | Proof. 135 | Abort. 136 | 137 | End Context. 138 | 139 | Infix "<=c" := ctxleq (at level 70, no associativity). 140 | 141 | Lemma ctxnth_map A B (f : A -> B) xs n : 142 | ctxnth (ctxmap f xs) n = omap f (ctxnth xs n). 143 | Proof. by elim: xs n => [| x xs IH] []. Defined. 144 | 145 | Lemma ctxindex_map (A B : eqType) (f : A -> B) xs n x : 146 | ctxindex xs n x -> ctxindex (ctxmap f xs) n (f x). 147 | Proof. by rewrite ctxnth_map => /eqP <-. Qed. 148 | 149 | Lemma ctxleq_map (A B : eqType) (f : A -> B) xs ys : 150 | xs <=c ys -> ctxmap f xs <=c ctxmap f ys. 151 | Proof. 152 | move/ctxleqP => H; apply/ctxleqP => n a. 153 | move: (H n); rewrite -!(nth_map' (omap f) None). 154 | by case: (ctxnth xs n) => //= a' /(_ a' (eqxx _)) /eqP => <-. 155 | Qed. 156 | 157 | Hint Resolve ctxindex_map ctxleqxx ctxleq_trans ctxleq_app 158 | ctxleq_appl ctxleq_appr ctxleq_map. 159 | 160 | (* Forall *) 161 | 162 | Fixpoint Forall A (P : A -> Prop) xs := 163 | if xs is x :: xs then P x /\ Forall P xs else True. 164 | 165 | Lemma Forall_impl : 166 | forall (A : Type) (P Q : A -> Prop) xs, 167 | (forall a, P a -> Q a) -> Forall P xs -> Forall Q xs. 168 | Proof. move => A P Q xs H; elim: xs; firstorder. Qed. 169 | 170 | Lemma Forall_map (A B : Type) (f : A -> B) P xs : 171 | Forall P (map f xs) <-> Forall (P \o f) xs. 172 | Proof. by elim: xs => //= x xs ->. Qed. 173 | 174 | Lemma Forall_nth (A : Type) (P : A -> Prop) xs : 175 | Forall P xs <-> (forall x m, m < size xs -> P (nth x xs m)). 176 | Proof. 177 | elim: xs => //= x' xs [] IH IH'; split. 178 | - by case => H H0 x [] //= n; rewrite ltnS; apply IH. 179 | - move => H; split. 180 | + by apply (H x' 0). 181 | + by apply IH' => x m; apply (H x m.+1). 182 | Qed. 183 | 184 | Lemma allP' (A : eqType) (P : pred A) xs : 185 | reflect (Forall P xs) (all P xs). 186 | Proof. 187 | apply (iffP idP); elim: xs => //= x xs IH. 188 | - by case/andP => -> /IH. 189 | - by case => ->. 190 | Qed. 191 | 192 | (* zip *) 193 | 194 | Fixpoint zipwith (A B C : Type) (f : A -> B -> C) xd yd 195 | (xs : seq A) (ys : seq B) : seq C := 196 | match xs, ys with 197 | | [::], ys => map (f xd) ys 198 | | xs, [::] => map (f^~ yd) xs 199 | | x :: xs', y :: ys' => f x y :: zipwith f xd yd xs' ys' 200 | end. 201 | -------------------------------------------------------------------------------- /coq/lib/seq_ext_base.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Printing Implicit Defensive. 6 | 7 | Section Seq. 8 | 9 | Variable (A B : Type). 10 | 11 | (* drop and take *) 12 | 13 | Lemma drop_addn n m (xs : seq A) : drop n (drop m xs) = drop (n + m) xs. 14 | Proof. 15 | elim: m xs => [| m IH] [] // x xs. 16 | - by rewrite addn0. 17 | - by rewrite addnS /= IH. 18 | Qed. 19 | 20 | Lemma take_minn n m (xs : seq A) : take n (take m xs) = take (minn n m) xs. 21 | Proof. by elim: n m xs => [| n IH] [| m] [] // x xs; rewrite minnSS /= IH. Qed. 22 | 23 | Lemma drop_take_distr n m (xs : seq A) : 24 | drop n (take m xs) = take (m - n) (drop n xs). 25 | Proof. 26 | elim: n m xs => [m xs | n IH [xs | m [] //= _ xs]]. 27 | - by rewrite !drop0 subn0. 28 | - by rewrite sub0n !take0. 29 | - by rewrite subSS IH. 30 | Qed. 31 | 32 | Lemma take_drop_distr n m (xs : seq A) : 33 | take n (drop m xs) = drop m (take (n + m) xs). 34 | Proof. by rewrite drop_take_distr addnK. Qed. 35 | 36 | Lemma drop_take_nil n (xs : seq A) : drop n (take n xs) = [::]. 37 | Proof. by rewrite drop_take_distr subnn take0. Qed. 38 | 39 | Lemma size_take n (xs : seq A) : size (take n xs) = minn n (size xs). 40 | Proof. 41 | elim: n xs => [xs | n IH [] //= _ xs]. 42 | - by rewrite take0 min0n. 43 | - by rewrite minnSS IH. 44 | Qed. 45 | 46 | Lemma nth_take' x n m (xs : seq A) : 47 | nth x (take n xs) m = if n <= m then x else nth x xs m. 48 | Proof. 49 | case: (leqP n m) => H. 50 | - by rewrite nth_default // size_take geq_min H. 51 | - by rewrite nth_take. 52 | Qed. 53 | 54 | (* nth *) 55 | 56 | Lemma nth_map' (f : A -> B) x xs n : f (nth x xs n) = nth (f x) (map f xs) n. 57 | Proof. by elim: xs n => [n | x' xs IH []] //=; rewrite !nth_nil. Qed. 58 | 59 | Lemma nth_equal (a b : A) xs n : 60 | (size xs <= n -> a = b) -> nth a xs n = nth b xs n. 61 | Proof. by elim: xs n => [n /= -> | x xs IH []]. Qed. 62 | 63 | End Seq. 64 | -------------------------------------------------------------------------------- /coq/lib/ssrnat_ext.v: -------------------------------------------------------------------------------- 1 | Require Import Omega Psatz. 2 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. 3 | From LCAC Require Import seq_ext_base. 4 | 5 | Set Implicit Arguments. 6 | Unset Strict Implicit. 7 | Unset Printing Implicit Defensive. 8 | 9 | Definition natE := 10 | (addSn, addnS, add0n, addn0, sub0n, subn0, subSS, 11 | min0n, minn0, max0n, maxn0, leq0n). 12 | 13 | (* Extended comparison predicates. *) 14 | 15 | Variant leq_xor_gtn' m n : 16 | bool -> bool -> bool -> bool -> 17 | nat -> nat -> nat -> nat -> nat -> nat -> Set := 18 | | LeqNotGtn' of m <= n : 19 | leq_xor_gtn' m n (m < n) false true (n <= m) n n m m 0 (n - m) 20 | | GtnNotLeq' of n < m : 21 | leq_xor_gtn' m n false true false true m m n n (m - n) 0. 22 | 23 | Lemma leqP' m n : leq_xor_gtn' m n 24 | (m < n) (n < m) (m <= n) (n <= m) 25 | (maxn m n) (maxn n m) (minn m n) (minn n m) 26 | (m - n) (n - m). 27 | Proof. 28 | case: (leqP m n) => H. 29 | - by move: (H); rewrite -subn_eq0 => /eqP ->; constructor. 30 | - rewrite (ltnW H) leqNgt ltnS ltnW //. 31 | by move: (ltnW H); rewrite -subn_eq0 => /eqP ->; constructor. 32 | Qed. 33 | 34 | Variant compare_nat' m n : 35 | bool -> bool -> bool -> bool -> bool -> 36 | nat -> nat -> nat -> nat -> nat -> nat -> Set := 37 | | CompareNatLt' of m < n : 38 | compare_nat' m n true false false true false n n m m 0 (n - m) 39 | | CompareNatGt' of m > n : 40 | compare_nat' m n false true false false true m m n n (m - n) 0 41 | | CompareNatEq' of m = n : 42 | compare_nat' m n false false true true true m m m m 0 0. 43 | 44 | Lemma ltngtP' m n : compare_nat' m n 45 | (m < n) (n < m) (m == n) (m <= n) (n <= m) 46 | (maxn m n) (maxn n m) (minn m n) (minn n m) 47 | (m - n) (n - m). 48 | Proof. 49 | case: (ltngtP m n) => [H|H|->]; last by rewrite subnn; constructor. 50 | - by move: (ltnW H); rewrite -subn_eq0 => /eqP ->; constructor. 51 | - by move: (ltnW H); rewrite -subn_eq0 => /eqP ->; constructor. 52 | Qed. 53 | 54 | (* simpl_natarith *) 55 | 56 | Module simpl_natarith. 57 | Lemma lem1_1 ml mr n r : ml = r + n -> ml + mr = r + mr + n. 58 | Proof. by move => ->; rewrite addnAC. Qed. 59 | Lemma lem1_2 ml mr n r : mr = r + n -> ml + mr = ml + r + n. 60 | Proof. by move => ->; rewrite addnA. Qed. 61 | Lemma lem1_3 m' n r : m' = r + n -> m'.+1 = r.+1 + n. 62 | Proof. by move => ->; rewrite addSn. Qed. 63 | Lemma lem2_1 ml mr n r : ml - n = r -> ml - mr - n = r - mr. 64 | Proof. by move => <-; rewrite subnAC. Qed. 65 | Lemma lem2_2 m' n r : m' - n = r -> m'.-1 - n = r.-1. 66 | Proof. by move => <-; rewrite -subnS -add1n subnDA subn1. Qed. 67 | Lemma lem2_3 m n r : m = r + n -> m - n = r. 68 | Proof. by move => ->; rewrite addnK. Qed. 69 | Lemma lem3_1 m m' m'' nl nl' nl'' nr nr' : 70 | m - nl = m' - nl' -> m' - nl' - nr = m'' - nl'' - nr' -> 71 | m - (nl + nr) = m'' - (nl'' + nr'). 72 | Proof. by rewrite !subnDA => -> ->. Qed. 73 | Lemma lem3_2 m n r : m - (n + 1) = r -> m - n.+1 = r. 74 | Proof. by rewrite addn1. Qed. 75 | Lemma lem3_3 m n r : m - n = r -> m - n = r - 0. 76 | Proof. by rewrite subn0. Qed. 77 | Lemma lem4_1 m n m' n' : m - n = m' - n' -> (m <= n) = (m' <= n'). 78 | Proof. by rewrite -!subn_eq0 => ->. Qed. 79 | End simpl_natarith. 80 | 81 | Ltac simpl_natarith1 m n := 82 | match m with 83 | | n => constr: (esym (add0n n)) 84 | | ?ml + ?mr => let H := simpl_natarith1 ml n in 85 | constr: (simpl_natarith.lem1_1 mr H) 86 | | ?ml + ?mr => let H := simpl_natarith1 mr n in 87 | constr: (simpl_natarith.lem1_2 ml H) 88 | | ?m'.+1 => let H := simpl_natarith1 m' n in 89 | constr: (simpl_natarith.lem1_3 H) 90 | | ?m'.+1 => match n with 1 => constr: (esym (addn1 m')) end 91 | end. 92 | 93 | Ltac simpl_natarith2 m n := 94 | match m with 95 | | ?ml - ?mr => let H := simpl_natarith2 ml n in 96 | constr: (simpl_natarith.lem2_1 mr H) 97 | | ?m'.-1 => let H := simpl_natarith2 m' n in 98 | constr: (simpl_natarith.lem2_2 H) 99 | | _ => let H := simpl_natarith1 m n in constr: (simpl_natarith.lem2_3 H) 100 | end. 101 | 102 | Ltac simpl_natarith3 m n := 103 | lazymatch n with 104 | | ?nl + ?nr => 105 | simpl_natarith3 m nl; 106 | match goal with |- _ = ?m1 -> _ => 107 | let H := fresh "H" in 108 | move => H; simpl_natarith3 m1 nr; move/(simpl_natarith.lem3_1 H) => {H} 109 | end 110 | | _ => 111 | match n with 112 | | ?n'.+1 => 113 | lazymatch n' with 114 | | 0 => fail 115 | | _ => simpl_natarith3 m (n' + 1); move/simpl_natarith.lem3_2 116 | end 117 | | _ => let H := simpl_natarith2 m n in move: (simpl_natarith.lem3_3 H) 118 | | _ => move: (erefl (m - n)) 119 | end 120 | end. 121 | 122 | Ltac simpl_natarith := 123 | let tac x := 124 | lazymatch goal with 125 | | |- ?x = ?x -> _ => move => _; rewrite !natE 126 | | _ => move => ->; rewrite ?natE 127 | end in 128 | repeat 129 | (match goal with 130 | H : context [?m - ?n] |- _ => move: H; simpl_natarith3 m n; tac 0 => H 131 | end || 132 | match goal with 133 | |- context [?m - ?n] => simpl_natarith3 m n; tac 0 134 | end || 135 | match goal with 136 | H : context [?m <= ?n] |- _ => 137 | move: H; simpl_natarith3 m n; move/simpl_natarith.lem4_1; tac 0 => H 138 | end || 139 | match goal with 140 | |- context [?m <= ?n] => 141 | simpl_natarith3 m n; move/simpl_natarith.lem4_1; tac 0 142 | end); 143 | try done; 144 | repeat match goal with 145 | | H : is_true true |- _ => clear H 146 | (* "move => {H}" may unfold the "is_true" *) 147 | end. 148 | 149 | (* elimleq *) 150 | 151 | Tactic Notation "elimleq" := 152 | match goal with |- is_true (?n <= ?m) -> _ => 153 | is_var m; 154 | (let H := fresh "H" in move/subnKC => H; rewrite <- H in *; clear H); 155 | let rec tac := 156 | lazymatch reverse goal with 157 | | H : context [m] |- _ => move: H; tac => H 158 | | _ => move: {m} (m - n) => m; rewrite ?(addKn, addnK) 159 | end in tac; simpl_natarith 160 | end. 161 | 162 | Tactic Notation "elimleq" constr(H) := move: H; elimleq. 163 | 164 | (* ssromega *) 165 | 166 | Tactic Notation "find_minneq_hyp" constr(m) constr(n) := 167 | match goal with 168 | | H : is_true (m <= n) |- _ => rewrite (minn_idPl H) 169 | | H : is_true (n <= m) |- _ => rewrite (minn_idPr H) 170 | | H : is_true (m < n) |- _ => rewrite (minn_idPl (ltnW H)) 171 | | H : is_true (n < m) |- _ => rewrite (minn_idPr (ltnW H)) 172 | | |- _ => case (leqP' m n) 173 | end; rewrite ?natE. 174 | 175 | Tactic Notation "find_maxneq_hyp" constr(m) constr(n) := 176 | match goal with 177 | | H : is_true (m <= n) |- _ => rewrite (maxn_idPr H) 178 | | H : is_true (n <= m) |- _ => rewrite (maxn_idPl H) 179 | | H : is_true (m < n) |- _ => rewrite (maxn_idPr (ltnW H)) 180 | | H : is_true (n < m) |- _ => rewrite (maxn_idPl (ltnW H)) 181 | | |- _ => case (leqP' m n) 182 | end; rewrite ?natE. 183 | 184 | Ltac replace_minn_maxn := 185 | try (rewrite <- minnE in * || rewrite <- maxnE in * ); 186 | match goal with 187 | | H : context [minn ?m ?n] |- _ => move: H; find_minneq_hyp n m => H 188 | | H : context [maxn ?m ?n] |- _ => move: H; find_maxneq_hyp n m => H 189 | | |- context [minn ?m ?n] => find_minneq_hyp m n 190 | | |- context [maxn ?m ?n] => find_maxneq_hyp m n 191 | end; 192 | try (let x := fresh "x" in move => x). 193 | 194 | Ltac arith_hypo_ssrnat2coqnat := 195 | match goal with 196 | | H : is_true false |- _ => by move: H 197 | | H : is_true (_ && _) |- _ => let H0 := fresh "H" in case/andP: H => H H0 198 | | H : is_true (_ || _) |- _ => case/orP in H 199 | | H : is_true (_ < _) |- _ => move/ltP in H 200 | | H : is_true (_ <= _) |- _ => move/leP in H 201 | | H : is_true (_ == _) |- _ => move/eqP in H 202 | | H : is_true (_ != _) |- _ => move/eqP in H 203 | end. 204 | 205 | Ltac arith_goal_ssrnat2coqnat := 206 | match goal with 207 | | |- is_true (_ && _) => apply/andP; split 208 | | |- is_true (_ || _) => apply/orP 209 | | |- is_true (_ < _) => apply/ltP 210 | | |- is_true (_ <= _) => apply/leP 211 | | |- is_true (_ == _) => apply/eqP 212 | | |- is_true (_ != _) => apply/eqP 213 | end. 214 | 215 | Ltac ssromega := 216 | repeat (let x := fresh "x" in move => /= x); 217 | do ?replace_minn_maxn; 218 | try done; 219 | repeat match goal with H : is_true (?m <= ?n) |- _ => elimleq H end; 220 | repeat (rewrite <- leqNgt in * || rewrite <- ltnNge in * ); 221 | do ?unfold addn, subn, muln, addn_rec, subn_rec, muln_rec in *; 222 | do ?arith_hypo_ssrnat2coqnat; 223 | do ?arith_goal_ssrnat2coqnat; 224 | simpl Equality.sort in *; 225 | lia. 226 | (* 227 | omega. 228 | *) 229 | 230 | Ltac elimif' := 231 | (match goal with 232 | | |- context [if ?m < ?n then _ else _] => case (leqP' n m) 233 | | |- context [if ?m <= ?n then _ else _] => case (leqP' m n) 234 | | |- context [if ?b then _ else _] => case (ifP b) 235 | end; 236 | move => //=; elimif'; let hyp := fresh "H" in move => hyp) || 237 | idtac. 238 | 239 | Ltac elimif := 240 | elimif'; simpl_natarith; 241 | repeat match goal with H : is_true (?m <= ?n) |- _ => elimleq H end. 242 | 243 | Ltac elimif_omega := 244 | elimif; 245 | try (repeat match goal with 246 | | |- @eq nat _ _ => idtac 247 | | |- nth _ _ _ = nth _ _ _ => apply nth_equal 248 | | |- _ => f_equal 249 | end; ssromega). 250 | 251 | (* congruence' *) 252 | 253 | Ltac congruence' := simpl; try (move: addSn addnS; congruence). 254 | 255 | (* test code for ssromega *) 256 | 257 | Module ssromega_test. 258 | Goal forall m n, minn (maxn m n) m = m. ssromega. Qed. 259 | Goal forall m n, minn n (maxn m n) = n. ssromega. Qed. 260 | Goal forall m n, maxn (minn m n) m = m. ssromega. Qed. 261 | Goal forall m n, maxn n (minn m n) = n. ssromega. Qed. 262 | Goal forall m n, maxn m n = m + (n - m). ssromega. Qed. 263 | Goal forall m n, minn m n = m - (m - n). ssromega. Qed. 264 | Goal forall m n, minn m n = m <-> m <= n. split; ssromega. Qed. 265 | Goal forall m n, maxn m n = n <-> m <= n. split; ssromega. Qed. 266 | Goal forall m n, maxn m n - minn m n = (m - n) + (n - m). ssromega. Qed. 267 | Goal forall m n, minn m n - maxn m n = 0. ssromega. Qed. 268 | Goal forall m n, minn m n + maxn m n = m + n. ssromega. Qed. 269 | Goal forall m n, minn m n + (m - n) = m. ssromega. Qed. 270 | Goal forall m n, maxn m n - (n - m) = m. ssromega. Qed. 271 | End ssromega_test. 272 | --------------------------------------------------------------------------------