├── Equivalence.v ├── Makefile ├── README ├── Semantics.v ├── Substitution.v ├── Syntax.v └── Termination.v /Equivalence.v: -------------------------------------------------------------------------------- 1 | Require Import Termination. 2 | Require Import Setoid. 3 | 4 | Open Scope t_scope. 5 | Open Scope list_scope. 6 | 7 | Reserved Notation " M ≃ N " (at level 70). 8 | Inductive Eq2 M N : Prop := 9 | | equivT : forall 10 | (HRedM : M ↦* TT) 11 | (HRedN : N ↦* TT), 12 | M ≃ N 13 | | equivF : forall 14 | (HRedM : M ↦* FF) 15 | (HRedN : N ↦* FF), 16 | M ≃ N 17 | where " M ≃ N " := (Eq2 M N). 18 | 19 | Lemma step_det : forall M N₁ N₂ 20 | (Hred₁ : M ↦ N₁) 21 | (Hred₂ : M ↦ N₂), 22 | N₁ = N₂. 23 | Proof. 24 | intros; generalize dependent N₂; induction Hred₁; intros; 25 | try (inversion Hred₂; subst; [reflexivity | inversion HR]); 26 | inversion Hred₂; subst; (inversion Hred₁; fail) || f_equal; intuition. 27 | Qed. 28 | 29 | Lemma value_not_red : forall V M 30 | (Hval : value V) 31 | (Hred : V ↦* M), 32 | M = V. 33 | Proof. 34 | induction 2; [reflexivity |]. 35 | inversion H; subst; inversion Hval. 36 | Qed. 37 | 38 | Lemma steps_val_det : forall M V₁ V₂ 39 | (Hval₁ : value V₁) 40 | (Hval₂ : value V₂) 41 | (Hred₁ : M ↦* V₁) 42 | (Hred₂ : M ↦* V₂), 43 | V₁ = V₂. 44 | Proof. 45 | intros; generalize dependent V₂; induction Hred₁; intros. 46 | symmetry; apply value_not_red; assumption. 47 | inversion Hred₂; subst. 48 | inversion H; subst; inversion Hval₂. 49 | assert (y = y0) by (eauto using step_det); subst. 50 | apply IHHred₁; assumption. 51 | Qed. 52 | 53 | Instance Eq2_sym : Symmetric Eq2. 54 | Proof. 55 | intros x. 56 | unfold Symmetric; intros. 57 | inversion H; auto using Eq2. 58 | Qed. 59 | 60 | Instance Eq2_trans : Transitive Eq2. 61 | Proof. 62 | unfold Transitive; intros. 63 | inversion H0; subst; inversion H; subst; auto using Eq2; 64 | assert (TT = FF) by (eauto using steps_val_det, value); discriminate. 65 | Qed. 66 | 67 | Lemma Eq2_type_refl : forall M (HT : nil ⊢ M ::: 2), M ≃ M. 68 | Proof. 69 | intros; assert (HR := types_red _ _ _ nil HT I I). 70 | simpl in *; inversion HR; subst; auto using Eq2. 71 | Qed. 72 | 73 | Inductive Eqω : te -> te -> Prop := 74 | | Eqωz : forall M N (HMz : M ↦* z) (HNz : N ↦* z), Eqω M N 75 | | Eqωs : forall M N M' N' (HMs : M ↦* s M') (HNs : N ↦* s N') (HEq : Eqω M' N'), Eqω M N. 76 | 77 | Instance Eqω_sym : Symmetric Eqω. 78 | Proof. 79 | unfold Symmetric; intros; induction H; eauto using Eqω. 80 | Qed. 81 | 82 | Instance Eqω_trans : Transitive Eqω. 83 | Proof. 84 | unfold Transitive; intros. 85 | revert z H0; induction H; intros u Hu; inversion Hu; subst; eauto using Eqω. 86 | assert (z = s M') by (eauto using steps_val_det, value); discriminate. 87 | assert (z = s N') by (eauto using steps_val_det, value); discriminate. 88 | assert (HT : s N' = s M'0) by (eauto using steps_val_det, value); inversion HT; subst; 89 | eauto using Eqω. 90 | Qed. 91 | 92 | CoInductive Eqs (P : te -> te -> Prop) : te -> te -> Prop := 93 | | EqS : forall M N (HEqHD : P (hd M) (hd N)) (HEqTL : Eqs P (tl M) (tl N)), Eqs P M N. 94 | 95 | Instance Eqs_sym P (HSym : Symmetric P) : Symmetric (Eqs P). 96 | Proof. 97 | unfold Symmetric; cofix; intros x y H. 98 | inversion H; subst; constructor. 99 | symmetry; assumption. 100 | apply Eqs_sym; assumption. 101 | Qed. 102 | 103 | Instance Eqs_trans P (HTran : Transitive P) : Transitive (Eqs P). 104 | Proof. 105 | cofix; unfold Transitive; intros u v w Huv Hvw. 106 | inversion Huv; subst; inversion Hvw; subst; constructor. 107 | eapply HTran; eassumption. 108 | eapply Eqs_trans; eassumption. 109 | Qed. 110 | 111 | Inductive ctx : Type := 112 | | hole : ctx 113 | | ct_lam : ty -> ctx -> ctx 114 | | ct_appL: te -> ctx -> ctx 115 | | ct_appR: te -> ctx -> ctx 116 | | ct_s : ctx -> ctx 117 | | ct_rec0: te -> te -> ctx -> ctx 118 | | ct_rec1: te -> te -> ctx -> ctx 119 | | ct_rec2: te -> te -> ctx -> ctx 120 | | ct_hd : ctx -> ctx 121 | | ct_tl : ctx -> ctx 122 | | ct_sd0 : te -> te -> ctx -> ctx 123 | | ct_sd1 : te -> te -> ctx -> ctx 124 | | ct_sd2 : te -> te -> ctx -> ctx. 125 | 126 | Fixpoint plug (E : ctx) (M : te) : te := 127 | match E with 128 | | hole => M 129 | | ct_lam A E => λ A, plug E M 130 | | ct_appR N E => N @ (plug E M) 131 | | ct_appL N E => (plug E M) @ N 132 | | ct_s E => s (plug E M) 133 | | ct_rec0 M₀ M₁ E => rec (plug E M) M₀ M₁ 134 | | ct_rec1 M₀ M₁ E => rec M₀ (plug E M) M₁ 135 | | ct_rec2 M₀ M₁ E => rec M₀ M₁ (plug E M) 136 | | ct_hd E => hd (plug E M) 137 | | ct_tl E => tl (plug E M) 138 | | ct_sd0 M₀ M₁ E => seed (plug E M) M₀ M₁ 139 | | ct_sd1 M₀ M₁ E => seed M₀ (plug E M) M₁ 140 | | ct_sd2 M₀ M₁ E => seed M₀ M₁ (plug E M) 141 | end. 142 | 143 | Reserved Notation " E · F " (at level 40). 144 | 145 | Fixpoint compose (E E' : ctx) : ctx := 146 | match E with 147 | | hole => E' 148 | | ct_lam A E => ct_lam A (E · E') 149 | | ct_appL M E => ct_appL M (E · E') 150 | | ct_appR N E => ct_appR N (E · E') 151 | | ct_s E => ct_s (E · E') 152 | | ct_rec0 M N E => ct_rec0 M N (E · E') 153 | | ct_rec1 M N E => ct_rec1 M N (E · E') 154 | | ct_rec2 M N E => ct_rec2 M N (E · E') 155 | | ct_hd E => ct_hd (E · E') 156 | | ct_tl E => ct_tl (E · E') 157 | | ct_sd0 M N E => ct_sd0 M N (E · E') 158 | | ct_sd1 M N E => ct_sd1 M N (E · E') 159 | | ct_sd2 M N E => ct_sd2 M N (E · E') 160 | end where " E · E' " := (compose E E'). 161 | 162 | Lemma plug_compose : forall M E E', 163 | plug E (plug E' M) = plug (E · E') M. 164 | Proof. 165 | induction E; intros; simpl; try rewrite IHE; reflexivity. 166 | Qed. 167 | 168 | (* TODO Define context typing inductively and prove the property below holds *) 169 | Reserved Notation " E ::: [ Γ , A ] → [ Δ , B ] " (at level 70). 170 | 171 | Inductive tyctx : ctx -> list ty -> ty -> list ty -> ty -> Prop := 172 | | cty_hole : forall Γ A, hole ::: [Γ, A] → [Γ, A] 173 | | cty_lam : forall Γ Δ A B C E 174 | (HCT : E ::: [Γ, B] → [A :: Δ, C]), 175 | ct_lam A E ::: [Γ, B] → [Δ, A → C] 176 | | cty_appL : forall Γ Δ A B C E M 177 | (HCT : E ::: [Γ, A] → [Δ, B → C]) 178 | (HTR : Δ ⊢ M ::: B), 179 | ct_appL M E ::: [Γ, A] → [Δ, C] 180 | | cty_appR : forall Γ Δ A B C E M 181 | (HCT : E ::: [Γ, A] → [Δ, B]) 182 | (HTR : Δ ⊢ M ::: B → C), 183 | ct_appR M E ::: [Γ, A] → [Δ, C] 184 | | cty_s : forall Γ Δ A E 185 | (HCT : E ::: [Γ, A] → [Δ, ω]), 186 | ct_s E ::: [Γ, A] → [Δ, ω] 187 | | cty_rec0 : forall Γ Δ A B M₀ M₁ E 188 | (HCT : E ::: [Γ, A] → [Δ, ω]) 189 | (HT0 : Δ ⊢ M₀ ::: B) 190 | (HT1 : ω :: B :: Δ ⊢ M₁ ::: B), 191 | ct_rec0 M₀ M₁ E ::: [Γ, A] → [Δ, B] 192 | | cty_rec1 : forall Γ Δ A B M M₁ E 193 | (HCT : E ::: [Γ, A] → [Δ, B]) 194 | (HTM : Δ ⊢ M ::: ω) 195 | (HT1 : ω :: B :: Δ ⊢ M₁ ::: B), 196 | ct_rec1 M M₁ E ::: [Γ, A] → [Δ, B] 197 | | cty_rec2 : forall Γ Δ A B M M₀ E 198 | (HCT : E ::: [Γ, A] → [ω :: B :: Δ, B]) 199 | (HTM : Δ ⊢ M ::: ω) 200 | (HT1 : Δ ⊢ M₀ ::: B), 201 | ct_rec2 M M₀ E ::: [Γ, A] → [Δ, B] 202 | | cty_hd : forall Γ Δ A B E 203 | (HCT : E ::: [Γ, A] → [Δ, stream B]), 204 | ct_hd E ::: [Γ, A] → [Δ, B] 205 | | cty_tl : forall Γ Δ A B E 206 | (HCT : E ::: [Γ, A] → [Δ, stream B]), 207 | ct_tl E ::: [Γ, A] → [Δ, stream B] 208 | | cty_seed0 : forall Γ Δ A B M₀ M₁ E 209 | (HCT : E ::: [Γ, A] → [Δ, B]) 210 | (HT0 : B :: Δ ⊢ M₀ ::: B) 211 | (HT1 : B :: Δ ⊢ M₁ ::: B), 212 | ct_sd0 M₀ M₁ E ::: [Γ, A] → [Δ, stream B] 213 | | cty_seed1 : forall Γ Δ A B M M₁ E 214 | (HCT : E ::: [Γ, A] → [B :: Δ, B]) 215 | (HTM : Δ ⊢ M ::: B) 216 | (HT1 : B :: Δ ⊢ M₁ ::: B), 217 | ct_sd1 M M₁ E ::: [Γ, A] → [Δ, stream B] 218 | | cty_seed2 : forall Γ Δ A B M M₀ E 219 | (HCT : E ::: [Γ, A] → [B :: Δ, B]) 220 | (HTM : Δ ⊢ M ::: B) 221 | (HT1 : B :: Δ ⊢ M₀ ::: B), 222 | ct_sd2 M M₀ E ::: [Γ, A] → [Δ, stream B] 223 | 224 | where " E ::: [ Γ , A ] → [ Δ , B ] " := (tyctx E Γ A Δ B). 225 | 226 | Lemma context_type : forall {E Γ Δ A B M} 227 | (HCT : E ::: [Γ, A] → [Δ, B]) 228 | (HMT : Γ ⊢ M ::: A), 229 | Δ ⊢ plug E M ::: B. 230 | Proof. 231 | induction 1; intros; simpl in *; eauto using types. 232 | Qed. 233 | 234 | Lemma context_comp : forall {E F Γ Δ Σ A B C} 235 | (HC1 : F ::: [Δ, B] → [Σ, C]) 236 | (HC2 : E ::: [Γ, A] → [Δ, B]), 237 | F · E ::: [Γ, A] → [Σ, C]. 238 | Proof. 239 | intros; generalize dependent Γ; revert A E; induction HC1; 240 | intros; simpl in *; eauto using tyctx. 241 | Qed. 242 | 243 | Definition obs_eq Γ A (M N : {K : te | Γ ⊢ K ::: A}) := 244 | forall E (HTE : E ::: [ Γ, A ] → [ nil, 2 ]), 245 | plug E (proj1_sig M) ≃ plug E (proj1_sig N). 246 | Notation " Γ ⊢ M == N ::: A " := (obs_eq Γ A M N) (at level 70). 247 | 248 | Section Defs. 249 | Variable P : forall Γ A, relation {K : te | Γ ⊢ K ::: A}. 250 | 251 | Definition consistent := 252 | forall M N, P nil 2 M N -> proj1_sig M ≃ proj1_sig N. 253 | 254 | Definition closed := 255 | forall {E Γ Δ A B M N} 256 | (HTE : E ::: [Γ, A] → [Δ, B]) 257 | (HPMN : P Γ A M N), 258 | (P Δ B (exist _ (plug E (proj1_sig M)) (context_type HTE (proj2_sig M))) 259 | (exist _ (plug E (proj1_sig N)) (context_type HTE (proj2_sig N)))). 260 | 261 | Record ConsCong := 262 | { Cong_Equiv : forall Γ A, Equivalence (P Γ A); 263 | Cong_Cons : consistent; 264 | Cong_Closed : closed}. 265 | 266 | End Defs. 267 | 268 | Lemma obs_eq_cong : ConsCong obs_eq. 269 | Proof. 270 | split; intros. 271 | split; unfold Reflexive, Symmetric, Transitive; intros. 272 | (* refl *) 273 | unfold obs_eq; intros. 274 | destruct x. 275 | eauto using Eq2_type_refl, @context_type. 276 | (* symm *) 277 | unfold obs_eq; intros; symmetry; apply H; assumption. 278 | (* tran *) 279 | unfold obs_eq; intros; etransitivity; eauto; fail. 280 | (* cons *) 281 | unfold consistent; intros. 282 | specialize (H hole); simpl in H; apply H. 283 | constructor. 284 | (* cong *) 285 | unfold closed, obs_eq; intros; simpl in *. 286 | simpl; rewrite !plug_compose; apply HPMN. 287 | eapply context_comp; eauto. 288 | Qed. 289 | 290 | Lemma obs_eq_coarsest : 291 | forall (P : forall Γ A, relation {K : te | Γ ⊢ K ::: A}) 292 | (HCC : ConsCong P) 293 | Γ A M N (HIn : P Γ A M N), 294 | Γ ⊢ M == N ::: A. 295 | Proof. 296 | unfold obs_eq; intros. 297 | destruct HCC as [Equiv Cons Cong]. 298 | unfold consistent in Cons. 299 | specialize (Cons (exist _ (plug E (proj1_sig M)) (context_type HTE (proj2_sig M))) (exist _ (plug E (proj1_sig N)) (context_type HTE (proj2_sig N)))); simpl in Cons; apply Cons; clear Cons. 300 | apply Cong; assumption. 301 | Qed. 302 | 303 | (* Logical Equivalence *) 304 | 305 | Reserved Notation " M ≡ N ::: A " (at level 70). 306 | 307 | Fixpoint log_equiv (A : ty) (M N : te) : Prop := 308 | nil ⊢ M ::: A /\ nil ⊢ N ::: A /\ 309 | match A with 310 | | 2 => M ≃ N 311 | | ω => Eqω M N 312 | | A₁ → A₂ => forall M' N' (HEq : M' ≡ N' ::: A₁), M @ M' ≡ N @ N' ::: A₂ 313 | | stream A => Eqs (log_equiv A) M N 314 | end where " M ≡ N ::: A " := (log_equiv A M N) : t_scope. 315 | 316 | Instance log_equiv_symm A : Symmetric (log_equiv A). 317 | Proof. 318 | induction A; unfold Symmetric; simpl; intuition. 319 | Qed. 320 | 321 | Instance log_equiv_trans A : Transitive (log_equiv A). 322 | Proof. 323 | induction A; unfold Transitive; simpl; intuition; etransitivity; eauto. 324 | eapply H5; etransitivity; [symmetry |]; eauto. 325 | Qed. 326 | 327 | Lemma log_equiv_typesL : forall M N A (HEq : M ≡ N ::: A), 328 | nil ⊢ M ::: A. 329 | Proof. 330 | destruct A; simpl in *; intuition. 331 | Qed. 332 | Lemma log_equiv_typesR : forall M N A (HEq : M ≡ N ::: A), 333 | nil ⊢ N ::: A. 334 | Proof. 335 | destruct A; simpl in *; intuition. 336 | Qed. 337 | 338 | Reserved Notation " γ₁ ∼ γ₂ ::: Γ " (at level 70). 339 | Fixpoint ctx_log_equiv Γ γ₁ γ₂ := 340 | match Γ, γ₁, γ₂ with 341 | | nil, nil, nil => True 342 | | A :: Γ, M :: γ₁, N :: γ₂ => (M ≡ N ::: A) /\ (γ₁ ∼ γ₂ ::: Γ) 343 | | _, _, _ => False 344 | end where " γ₁ ∼ γ₂ ::: Γ " := (ctx_log_equiv Γ γ₁ γ₂) : t_scope. 345 | 346 | Instance ctx_log_equiv_sym Γ : Symmetric (ctx_log_equiv Γ). 347 | Proof. 348 | induction Γ; unfold Symmetric; intros; destruct x; destruct y; 349 | try contradiction; [tauto |]. 350 | simpl in *; destruct H; split; symmetry; assumption. 351 | Qed. 352 | 353 | Instance ctx_log_equiv_trans Γ : Transitive (ctx_log_equiv Γ). 354 | Proof. 355 | induction Γ; unfold Transitive; intros; 356 | destruct x; destruct y; destruct z; try contradiction; [tauto |]. 357 | simpl in *; destruct H; destruct H0; split; etransitivity; eassumption. 358 | Qed. 359 | 360 | Lemma ctx_lequiv_typesL : forall Γ γ₁ γ₂ (HEq : γ₁ ∼ γ₂ ::: Γ), γ₁ :–: Γ. 361 | Proof. 362 | induction Γ; destruct γ₁; destruct γ₂; simpl in *; trivial || contradiction || 363 | intuition eauto using log_equiv_typesL. 364 | Qed. 365 | Lemma ctx_lequiv_typesR : forall Γ γ₁ γ₂ (HEq : γ₁ ∼ γ₂ ::: Γ), γ₂ :–: Γ. 366 | Proof. 367 | intros; symmetry in HEq; eauto using ctx_lequiv_typesL. 368 | Qed. 369 | 370 | Definition open_simeq Γ A (M N : {K : te | Γ ⊢ K ::: A}) := forall γ₁ γ₂ 371 | (HEΓ : γ₁ ∼ γ₂ ::: Γ), 372 | [γ₁ ! 0](proj1_sig M) ≡ [γ₂ ! 0](proj1_sig N) ::: A. 373 | Notation " Γ ⊢ M ∼ N ::: A " := (open_simeq Γ A M N) (at level 70) : t_scope. 374 | 375 | Instance open_simeq_symm Γ A : Symmetric (open_simeq Γ A). 376 | Proof. 377 | unfold Symmetric, open_simeq; intros. 378 | symmetry; apply H; try symmetry; assumption. 379 | Qed. 380 | 381 | Lemma open_simeq_trans : forall Γ A, transitive _ (open_simeq Γ A). 382 | Proof. 383 | unfold transitive, open_simeq; intros. 384 | etransitivity; [apply H | apply H0]; try eassumption. 385 | etransitivity; [symmetry |]; eassumption. 386 | Qed. 387 | 388 | Lemma open_simeq_cons : consistent open_simeq. 389 | Proof. 390 | unfold consistent; intros. 391 | apply (H nil nil I). 392 | Qed. 393 | 394 | Scheme types_Ind := Induction for types Sort Prop. 395 | 396 | Hint Rewrite sub_lam sub_app sub_z sub_s sub_rec sub_hd sub_tl sub_seed 397 | sub_TT sub_FF : sub. 398 | 399 | Ltac msimp := 400 | autorewrite with sub. 401 | 402 | Lemma head_expansion : forall A M M' N 403 | (HT : nil ⊢ M ::: A) 404 | (HR : M' ≡ N ::: A) 405 | (HS : M ↦ M'), 406 | M ≡ N ::: A. 407 | Proof. 408 | induction A; intros; simpl in *; intuition. 409 | (* bool *) 410 | inversion H2; [ apply equivT | apply equivF ]; auto; econstructor; eauto. 411 | (* nat *) 412 | inversion H2; subst; clear H2. 413 | apply Eqωz; auto; econstructor; eassumption. 414 | eapply Eqωs; eauto; econstructor; eassumption. 415 | (* arr *) 416 | intros. 417 | eapply IHA2; eauto using log_equiv_typesL, types, step. 418 | (* stream *) 419 | clear H1 H; generalize dependent M'; revert N M HT. 420 | cofix; intros. 421 | inversion H2; subst; clear H2; apply EqS; [eapply IHA | eapply head_expansion]; 422 | eauto using types, step. 423 | Qed. 424 | 425 | (* lifting of head expansion to reflexive transitive closure *) 426 | Lemma head_exp_star : forall A M M' N 427 | (HT : nil ⊢ M ::: A) 428 | (HR : M' ≡ N ::: A) 429 | (HS : M ↦* M'), 430 | M ≡ N ::: A. 431 | Proof. 432 | intros; induction HS; [assumption |]. 433 | eapply head_expansion; [| | eassumption]; eauto using preservation. 434 | Qed. 435 | 436 | Lemma preservation_star : forall A M N 437 | (HT : nil ⊢ M ::: A) 438 | (HR : M ↦* N), 439 | nil ⊢ N ::: A. 440 | Proof. 441 | induction 2; eauto using preservation. 442 | Qed. 443 | 444 | Lemma seed_cong : forall M M₀ M₁ N N₀ N₁ K L A 445 | (HMN : M ≡ N ::: A) 446 | (HMN₀ : forall K L, K ≡ L ::: A -> [K ↑ 0]M₀ ≡ [L ↑ 0]N₀ ::: A) 447 | (HMN₁ : forall K L, K ≡ L ::: A -> [K ↑ 0]M₁ ≡ [L ↑ 0]N₁ ::: A) 448 | (HTK : nil ⊢ K ::: stream A) 449 | (HTL : nil ⊢ L ::: stream A) 450 | (HRK : K ↦* seed M M₀ M₁) 451 | (HRL : L ↦* seed N N₀ N₁), 452 | K ≡ L ::: stream A. 453 | Proof. 454 | intros; simpl; intuition. 455 | generalize dependent M; generalize dependent N; revert K L HTK HTL. 456 | cofix; intros; simpl; apply EqS. 457 | apply head_exp_star with (M' := [M ↑ 0]M₀); eauto using types. 458 | symmetry; apply head_exp_star with (M' := [N ↑ 0]N₀); eauto using types. 459 | symmetry; eapply HMN₀; assumption. 460 | apply clos_rt_rt1n; eapply rt_trans; 461 | [apply clos_rt1n_rt, hd_cong_star; eassumption |]; 462 | eauto using rt_step, step. 463 | apply clos_rt_rt1n; eapply rt_trans; 464 | [apply clos_rt1n_rt, hd_cong_star; eassumption |]; 465 | eauto using rt_step, step. 466 | eapply seed_cong; try (apply tc_tl; auto). 467 | apply clos_rt_rt1n; eapply rt_trans; 468 | [apply clos_rt1n_rt, tl_cong_star; eassumption |]; 469 | eauto using rt_step, step. 470 | eauto. 471 | apply clos_rt_rt1n; eapply rt_trans; 472 | [apply clos_rt1n_rt, tl_cong_star; eassumption |]; 473 | eauto using rt_step, step. 474 | Qed. 475 | 476 | Lemma rec_cong : forall A M M₀ M₁ N N₀ N₁ 477 | (HTM₁ : ω :: A :: nil ⊢ M₁ ::: A) 478 | (HTN₁ : ω :: A :: nil ⊢ N₁ ::: A) 479 | (HMN : M ≡ N ::: ω) 480 | (HMN₀ : M₀ ≡ N₀ ::: A) 481 | (HMN₁ : forall K L (HKL : K ≡ L ::: ω) (HRKL : rec K M₀ M₁ ≡ rec L N₀ N₁ ::: A), 482 | [K :: rec K M₀ M₁ :: nil ! 0]M₁ ≡ [L :: rec L N₀ N₁ :: nil ! 0]N₁ ::: A), 483 | rec M M₀ M₁ ≡ rec N N₀ N₁ ::: A. 484 | Proof. 485 | intros; simpl. 486 | destruct HMN as [HTM [HTN HMN]]; induction HMN. 487 | apply head_exp_star with M₀; eauto using types, log_equiv_typesL. 488 | symmetry; apply head_exp_star with N₀; eauto using types, log_equiv_typesR. 489 | symmetry; eauto. 490 | apply clos_rt_rt1n; eapply rt_trans; 491 | [apply clos_rt1n_rt, rec_cong_star; eassumption |]; 492 | eauto using rt_step, step. 493 | apply clos_rt_rt1n; eapply rt_trans; 494 | [apply clos_rt1n_rt, rec_cong_star; eassumption |]; 495 | eauto using rt_step, step. 496 | apply head_exp_star with ([M' :: rec M' M₀ M₁ :: nil ! 0]M₁); 497 | eauto using types, log_equiv_typesL. 498 | symmetry; apply head_exp_star with ([N' :: rec N' N₀ N₁ :: nil ! 0]N₁); 499 | eauto using types, log_equiv_typesR. 500 | assert (HTM' : nil ⊢ M' ::: ω) by 501 | (eapply preservation_star in HMs; eauto; inversion HMs; subst; auto). 502 | assert (HTN' : nil ⊢ N' ::: ω) by 503 | (eapply preservation_star in HNs; eauto; inversion HNs; subst; auto). 504 | symmetry; eapply HMN₁; simpl; eauto. 505 | apply clos_rt_rt1n; eapply rt_trans; 506 | [apply clos_rt1n_rt, rec_cong_star; eassumption |]; 507 | eauto using rt_step, step. 508 | apply clos_rt_rt1n; eapply rt_trans; 509 | [apply clos_rt1n_rt, rec_cong_star; eassumption |]; 510 | eauto using rt_step, step. 511 | Qed. 512 | 513 | Lemma refl : forall Γ A, Reflexive (open_simeq Γ A). 514 | Proof. 515 | unfold Reflexive; intros. 516 | destruct x as [M HMT]. 517 | induction HMT using types_Ind; 518 | unfold open_simeq; simpl; intros; msimp; 519 | eauto using Eq2_type_refl, types. 520 | (* var *) 521 | replace n with (0 + n) by reflexivity; generalize 0; generalize dependent n; 522 | generalize dependent γ₂; revert γ₁; induction Γ; intros; 523 | destruct γ₁; destruct γ₂; try contradiction. 524 | destruct n; inversion HFind. 525 | destruct n; simpl in *; [inversion HFind; subst; clear HFind |]. 526 | rewrite plus_comm; simpl; rewrite <- !subst_gt; simpl; auto with arith. 527 | destruct (eq_nat_dec n0 n0); tauto. 528 | replace (n0 + S n) with (S n0 + n) by omega. 529 | eapply IHΓ with (n0 := S n0) in HFind; [| intuition eassumption]. 530 | replace n0 with (length (@nil te) + n0) by reflexivity; 531 | erewrite <- !closed_sub with (Γ := nil); simpl; 532 | eauto using log_equiv_typesL, log_equiv_typesR. 533 | (* lam *) 534 | repeat split; try (eapply tc_lam, subst_types; simpl; 535 | eauto using ctx_lequiv_typesL, ctx_lequiv_typesR). 536 | intros; eapply head_expansion; eauto using step. 537 | eapply tc_app; [| eauto using log_equiv_typesL]; eapply tc_lam, subst_types; simpl; 538 | eauto using ctx_lequiv_typesL. 539 | symmetry; eapply head_expansion; eauto using step. 540 | eapply tc_app; [| eauto using log_equiv_typesR]; eapply tc_lam, subst_types; simpl; 541 | eauto using ctx_lequiv_typesR. 542 | symmetry. 543 | rewrite !subcomp; eapply IHHMT; simpl; auto. 544 | (* app *) 545 | specialize (IHHMT1 _ _ HEΓ); simpl in IHHMT1; apply IHHMT1. 546 | specialize (IHHMT2 _ _ HEΓ); simpl in IHHMT2; apply IHHMT2. 547 | (* z *) 548 | repeat constructor. 549 | (* s *) 550 | repeat split; try (eapply tc_s, subst_types; simpl; 551 | eauto using ctx_lequiv_typesL, ctx_lequiv_typesR). 552 | apply IHHMT in HEΓ; simpl in *. 553 | destruct HEΓ as [HTL [HTR HEΓ]]. 554 | eapply Eqωs; eauto; constructor. 555 | (* rec *) 556 | eapply rec_cong. 557 | eapply subst_types; simpl; eauto using ctx_lequiv_typesL. 558 | eapply subst_types; simpl; eauto using ctx_lequiv_typesR. 559 | apply IHHMT1; auto. 560 | apply IHHMT2; auto. 561 | intros; rewrite !subcomp; apply IHHMT3; simpl; auto. 562 | (* hd *) 563 | apply IHHMT in HEΓ; simpl in *; destruct HEΓ as [HTL [HTR HEΓ]]; 564 | inversion HEΓ; subst; clear HEΓ; auto. 565 | (* tl *) 566 | apply IHHMT in HEΓ; simpl in *; destruct HEΓ as [HTL [HTR HEΓ]]; 567 | inversion HEΓ; subst; clear HEΓ; eauto using types. 568 | (* seed *) 569 | eapply seed_cong; try (constructor; fail); eauto. 570 | intros; change ([K :: γ₁ ! 0]M₀ ≡ [L :: γ₂ ! 0]M₀ ::: A); apply IHHMT2; simpl; auto. 571 | intros; change ([K :: γ₁ ! 0]M₁ ≡ [L :: γ₂ ! 0]M₁ ::: A); apply IHHMT3; simpl; auto. 572 | eapply tc_seed; repeat (eapply subst_types; simpl; eauto using ctx_lequiv_typesL). 573 | eapply tc_seed; repeat (eapply subst_types; simpl; eauto using ctx_lequiv_typesR). 574 | Qed. 575 | 576 | Scheme tyctx_Ind := Induction for tyctx Sort Prop. 577 | 578 | Lemma cong_log_equiv : closed open_simeq. 579 | Proof. 580 | unfold closed; intros. 581 | revert M N HPMN; induction HTE using tyctx_Ind; intros; simpl; 582 | unfold open_simeq; simpl; intros; msimp. 583 | (* hole *) 584 | apply HPMN; assumption. 585 | (* lam *) 586 | repeat split. 587 | eapply tc_lam, subst_types; eauto using ctx_lequiv_typesL; simpl. 588 | eapply context_type; eauto; apply (proj2_sig M). 589 | eapply tc_lam, subst_types; eauto using ctx_lequiv_typesR; simpl. 590 | eapply context_type; eauto; apply (proj2_sig N). 591 | intros. 592 | eapply head_expansion; eauto using step. 593 | eapply tc_app; [| eapply log_equiv_typesL; eauto]. 594 | eapply tc_lam, subst_types; eauto using ctx_lequiv_typesL; simpl. 595 | eapply context_type; eauto; apply (proj2_sig M). 596 | symmetry. 597 | eapply head_expansion; eauto using step. 598 | eapply tc_app; [| eapply log_equiv_typesR; eauto]. 599 | eapply tc_lam, subst_types; eauto using ctx_lequiv_typesR; simpl. 600 | eapply context_type; eauto; apply (proj2_sig N). 601 | symmetry. 602 | rewrite !subcomp. 603 | eapply IHHTE; simpl; eauto. 604 | (* appL *) 605 | eapply IHHTE; eauto. 606 | fold ([γ₁ ! 0]M ≡ [γ₂ ! 0]M ::: B). 607 | assert (HT := refl Δ B (exist _ M HTR)); eapply HT; eauto. 608 | (* appR *) 609 | assert (HT := refl _ _ (exist _ M HTR)); eapply HT; eauto. 610 | fold log_equiv. 611 | eapply IHHTE; eauto. 612 | (* s *) 613 | repeat split; try (apply IHHTE in HPMN; apply HPMN in HEΓ; simpl in HEΓ; intuition eauto using types; fail). 614 | eapply Eqωs; [constructor | constructor |]. 615 | eapply IHHTE; eauto. 616 | (* rec arg *) 617 | eapply rec_cong. 618 | eapply subst_types; simpl; eauto using ctx_lequiv_typesL. 619 | eapply subst_types; simpl; eauto using ctx_lequiv_typesR. 620 | apply IHHTE; auto. 621 | eapply (refl _ _ (exist _ M₀ HT0)); assumption. 622 | intros; rewrite !subcomp; 623 | eapply (refl _ _ (exist _ M₁ HT1)); simpl; auto. 624 | (* rec base *) 625 | eapply rec_cong. 626 | eapply subst_types; simpl; eauto using ctx_lequiv_typesL. 627 | eapply subst_types; simpl; eauto using ctx_lequiv_typesR. 628 | eapply (refl _ _ (exist _ M HTM)); simpl; auto. 629 | apply IHHTE; auto. 630 | intros; rewrite !subcomp; 631 | eapply (refl _ _ (exist _ M₁ HT1)); simpl; auto. 632 | (* rec step *) 633 | eapply rec_cong. 634 | eapply subst_types; simpl; eauto using ctx_lequiv_typesL. 635 | eapply context_type; eauto; apply (proj2_sig M0). 636 | eapply subst_types; simpl; eauto using ctx_lequiv_typesR. 637 | eapply context_type; eauto; apply (proj2_sig N). 638 | eapply (refl _ _ (exist _ M HTM)); simpl; auto. 639 | eapply (refl _ _ (exist _ M₀ HT1)); auto. 640 | intros; rewrite !subcomp; apply IHHTE; simpl; auto. 641 | (* hd *) 642 | eapply IHHTE in HEΓ; eauto; simpl in *. 643 | destruct HEΓ as [HTM [HTN HEΓ]]; inversion HEΓ; subst; clear HEΓ; auto. 644 | (* tl *) 645 | eapply IHHTE in HEΓ; eauto; simpl in *. 646 | destruct HEΓ as [HTM [HTN HEΓ]]; inversion HEΓ; auto; subst; clear HEΓ; eauto using types. 647 | (* seed arg *) 648 | eapply seed_cong; try (constructor; fail). 649 | apply IHHTE; auto. 650 | intros; change ([K :: γ₁ ! 0]M₀ ≡ [L :: γ₂ ! 0]M₀ ::: B). 651 | apply (refl _ _ (exist _ M₀ HT0)); simpl; auto. 652 | intros; change ([K :: γ₁ ! 0]M₁ ≡ [L :: γ₂ ! 0]M₁ ::: B). 653 | apply (refl _ _ (exist _ M₁ HT1)); simpl; auto. 654 | apply tc_seed; eapply subst_types; eauto using ctx_lequiv_typesL; simpl. 655 | eapply context_type; eauto; apply (proj2_sig M). 656 | apply tc_seed; eapply subst_types; eauto using ctx_lequiv_typesR; simpl. 657 | eapply context_type; eauto; apply (proj2_sig N). 658 | (* seed hd *) 659 | eapply seed_cong; try (constructor; fail). 660 | apply (refl _ _ (exist _ M HTM)); simpl; auto. 661 | intros; change ([K :: γ₁ ! 0](plug E (proj1_sig M0)) ≡ [L :: γ₂ ! 0](plug E (proj1_sig N)) ::: B). 662 | apply IHHTE; simpl; auto. 663 | intros; change ([K :: γ₁ ! 0]M₁ ≡ [L :: γ₂ ! 0]M₁ ::: B). 664 | apply (refl _ _ (exist _ M₁ HT1)); simpl; auto. 665 | apply tc_seed; eapply subst_types; eauto using ctx_lequiv_typesL; simpl. 666 | eapply context_type; eauto; apply (proj2_sig M0). 667 | apply tc_seed; eapply subst_types; eauto using ctx_lequiv_typesR; simpl. 668 | eapply context_type; eauto; apply (proj2_sig N). 669 | (* seed tl *) 670 | eapply seed_cong; try (constructor; fail). 671 | apply (refl _ _ (exist _ M HTM)); simpl; auto. 672 | intros; change ([K :: γ₁ ! 0]M₀ ≡ [L :: γ₂ ! 0]M₀ ::: B). 673 | apply (refl _ _ (exist _ M₀ HT1)); simpl; auto. 674 | intros; change ([K :: γ₁ ! 0](plug E (proj1_sig M0)) ≡ [L :: γ₂ ! 0](plug E (proj1_sig N)) ::: B). 675 | apply IHHTE; simpl; auto. 676 | apply tc_seed; eapply subst_types; eauto using ctx_lequiv_typesL; simpl. 677 | eapply context_type; eauto; apply (proj2_sig M0). 678 | apply tc_seed; eapply subst_types; eauto using ctx_lequiv_typesR; simpl. 679 | eapply context_type; eauto; apply (proj2_sig N). 680 | Qed. 681 | 682 | Program Definition plus : { K : te | nil ⊢ K ::: ω → ω → ω } := 683 | exist _ (λ ω, λ ω, rec (#1) (#0) (s (#1))) _. 684 | Next Obligation. 685 | do 2 apply tc_lam. 686 | apply tc_rec. 687 | apply tc_var; reflexivity. 688 | apply tc_var; reflexivity. 689 | apply tc_s, tc_var; reflexivity. 690 | Qed. 691 | 692 | Program Definition xp0 : { K : te | ω :: nil ⊢ K ::: ω} := 693 | exist _ (proj1_sig plus @ #0 @ z) _. 694 | Next Obligation. 695 | eapply tc_app; [| eauto using types]. 696 | eapply tc_app; [| apply tc_var; reflexivity]. 697 | change (nil ++ ω :: nil ⊢ proj1_sig plus ::: ω → ω → ω); apply weaken, (proj2_sig plus). 698 | Qed. 699 | 700 | Program Definition x : { K : te | ω :: nil ⊢ K ::: ω} := 701 | exist _ (#0) _. 702 | Next Obligation. 703 | apply tc_var; reflexivity. 704 | Qed. 705 | 706 | Lemma equiv : open_simeq (ω :: nil) ω x xp0. 707 | Proof. 708 | unfold open_simeq; intros. 709 | destruct γ₁; destruct γ₂; try contradiction; simpl in HEΓ. 710 | destruct γ₁; destruct γ₂; try tauto; destruct HEΓ as [[HTL [HTR HT]] _]. 711 | symmetry; apply head_exp_star with (M' := rec t0 z (s (#1))); simpl. 712 | eapply tc_app; [| eauto using types]. 713 | eapply tc_app; eauto; apply (proj2_sig plus). 714 | intuition. 715 | apply tc_rec; eauto using types. 716 | induction HT. 717 | apply Eqωz; auto. 718 | eapply clos_rt_rt1n, rt_trans. 719 | apply clos_rt1n_rt, rec_cong_star; eauto. 720 | apply rt_step, red_recz. 721 | eapply Eqωs. 722 | eapply clos_rt_rt1n, rt_trans; [| eapply rt_trans]. 723 | apply clos_rt1n_rt, rec_cong_star; eassumption. 724 | apply rt_step; eauto using step. 725 | msimp; simpl. 726 | apply rt_refl. 727 | eauto. 728 | eapply preservation_star in HNs; eauto. 729 | eapply preservation_star in HMs; eauto. 730 | inversion HMs; subst; inversion HNs; subst; clear HMs HNs. 731 | replace 0 with (length (@nil ty) + 0) by reflexivity; 732 | erewrite <- closed_sub with (Γ := nil); eauto. 733 | econstructor 2; eauto using step. 734 | msimp; simpl. 735 | econstructor 2; eauto using step. 736 | msimp; simpl. 737 | replace 0 with (length (@nil ty) + 0) by reflexivity; 738 | erewrite <- closed_sub with (Γ := nil); eauto. 739 | simpl; constructor. 740 | Qed. -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | MODULES := Syntax Substitution Semantics Termination 2 | 3 | VS := $(MODULES:%=%.v) 4 | 5 | .PHONY: coq clean 6 | 7 | coq: Makefile.coq 8 | $(MAKE) -j2 -f Makefile.coq 9 | 10 | Makefile.coq: Makefile $(VS) 11 | coq_makefile $(VS) -o Makefile.coq 12 | 13 | clean:: Makefile.coq 14 | $(MAKE) -f Makefile.coq clean 15 | rm -f Makefile.coq 16 | 17 | -------------------------------------------------------------------------------- /README: -------------------------------------------------------------------------------- 1 | This development formalises the notions from the first few lectures on 2 | Type Theory Foundations at OPLSS 2011. 3 | 4 | * Overview of the development: 5 | ** Syntax 6 | Defines the syntax of the language and some notations. Nothing fancy. 7 | ** Substitution 8 | Defines substitutions for de Bruijn indices. Pretty much the most boring 9 | thing possible. 10 | ** Semantics 11 | Defines static and dynamic semantics of the language, proves the standard 12 | properties: weakening, progress, preservation, etc. and lemmas needed by 13 | termination proof that are not related to logical relations. 14 | ** Termination 15 | Defines the reducibility relation and proceeds to prove head expansion 16 | and termination. 17 | ** Equivalence 18 | Defines contexts and context typings, observable equivalence and logical equivalence, 19 | and proves soundness of logical equivalence wrt observable equivalence. Provdies one 20 | example of a logical equivalence (x : ω ⊢ x ∼ x + 0 : ω). 21 | -------------------------------------------------------------------------------- /Semantics.v: -------------------------------------------------------------------------------- 1 | (** This file contains the definition of static and dynamic semantics, and basic 2 | properties, including preservation and progress lemmas *) 3 | 4 | Require Export Relations. 5 | Require Export List. 6 | Require Export Substitution. 7 | Require Import Omega. 8 | 9 | Open Scope t_scope. 10 | Open Scope list_scope. 11 | 12 | Global Reserved Notation " Γ '⊢' M ':::' A " (at level 70). 13 | Global Reserved Notation " γ ':–:' Γ " (at level 70). 14 | 15 | Section Statics. 16 | 17 | (* Typing relation for terms *) 18 | Inductive types : list ty -> te -> ty -> Prop := 19 | | tc_var : forall A Γ n 20 | (HFind : nth_error Γ n = Some A), 21 | Γ ⊢ #n ::: A 22 | | tc_lam : forall A B Γ M 23 | (HT : A :: Γ ⊢ M ::: B), 24 | Γ ⊢ λ A, M ::: A → B 25 | | tc_app : forall Γ A B M N 26 | (HTM : Γ ⊢ M ::: A → B) 27 | (HTN : Γ ⊢ N ::: A), 28 | Γ ⊢ M @ N ::: B 29 | | tc_z : forall Γ, 30 | Γ ⊢ z ::: ω 31 | | tc_s : forall Γ M 32 | (HT : Γ ⊢ M ::: ω), 33 | Γ ⊢ s M ::: ω 34 | | tc_rec : forall Γ A M M₀ M₁ 35 | (HTM : Γ ⊢ M ::: ω) 36 | (HT0 : Γ ⊢ M₀ ::: A) 37 | (HTS : ω :: A :: Γ ⊢ M₁ ::: A), 38 | Γ ⊢ rec M M₀ M₁ ::: A 39 | | tc_hd : forall Γ A M 40 | (HTM : Γ ⊢ M ::: stream A), 41 | Γ ⊢ hd M ::: A 42 | | tc_tl : forall Γ A M 43 | (HTM : Γ ⊢ M ::: stream A), 44 | Γ ⊢ tl M ::: stream A 45 | | tc_seed: forall Γ A M M₀ M₁ 46 | (HTM : Γ ⊢ M ::: A) 47 | (HTH : A :: Γ ⊢ M₀ ::: A) 48 | (HTT : A :: Γ ⊢ M₁ ::: A), 49 | Γ ⊢ seed M M₀ M₁ ::: stream A 50 | | tc_TT : forall Γ, Γ ⊢ TT ::: 2 51 | | tc_FF : forall Γ, Γ ⊢ FF ::: 2 52 | where " Γ ⊢ M ::: A " := (types Γ M A) : t_scope. 53 | 54 | (* Typing relation for substitutions (only for closed substitutions) *) 55 | Fixpoint tcmt γ Γ : Prop := 56 | match γ, Γ with 57 | | nil, nil => True 58 | | M :: γ, A :: Γ => nil ⊢ M ::: A /\ γ :–: Γ 59 | | _, _ => False 60 | end where " γ ':–:' Γ " := (tcmt γ Γ) : t_scope. 61 | 62 | End Statics. 63 | 64 | Global Reserved Notation " M ↦ N " (at level 70). 65 | Global Reserved Notation " M ↦* N " (at level 70). 66 | Section Dynamics. 67 | 68 | Inductive value : te -> Prop := 69 | | val_z : 70 | value z 71 | | val_s : forall M, 72 | value (s M) 73 | | val_lam : forall A M, 74 | value (λ A, M) 75 | | val_seed: forall M M₀ M₁, 76 | value (seed M M₀ M₁) 77 | | val_TT : value TT 78 | | val_FF : value FF. 79 | 80 | Inductive step : te -> te -> Prop := 81 | | red_β : forall A M N, 82 | (λ A, M) @ N ↦ [N :: nil ! 0]M 83 | | red_recz : forall M₀ M₁, 84 | rec z M₀ M₁ ↦ M₀ 85 | | red_recs : forall M M₀ M₁, 86 | rec (s M) M₀ M₁ ↦ [M :: rec M M₀ M₁ :: nil ! 0] M₁ 87 | | red_hds : forall M M₀ M₁, 88 | hd (seed M M₀ M₁) ↦ [M ↑ 0]M₀ 89 | | red_tls : forall M M₀ M₁, 90 | tl (seed M M₀ M₁) ↦ seed [M ↑ 0]M₁ M₀ M₁ 91 | | red_appC : forall M M' N 92 | (HR : M ↦ M'), 93 | M @ N ↦ M' @ N 94 | | red_recC : forall M M' M₀ M₁ 95 | (HR : M ↦ M'), 96 | rec M M₀ M₁ ↦ rec M' M₀ M₁ 97 | | red_hdC : forall M M' 98 | (HR : M ↦ M'), 99 | hd M ↦ hd M' 100 | | red_tlC : forall M M' 101 | (HR : M ↦ M'), 102 | tl M ↦ tl M' 103 | where " M ↦ N " := (step M N) : t_scope. 104 | Definition steps := clos_refl_trans_1n _ step. 105 | 106 | End Dynamics. 107 | 108 | Notation " Γ ⊢ M ::: A " := (types Γ M A) : t_scope. 109 | Notation " γ ':–:' Γ " := (tcmt γ Γ) : t_scope. 110 | Notation " M ↦* N " := (steps M N) : t_scope. 111 | Notation " M ↦ N " := (step M N) : t_scope. 112 | 113 | Section Properties. 114 | 115 | Lemma weaken : forall Γ Δ K A (HT : Γ ⊢ K ::: A), (Γ ++ Δ ⊢ K ::: A). 116 | Proof. 117 | induction 1; eauto using types; []. 118 | apply tc_var. 119 | generalize dependent n; induction Γ; intros; simpl; destruct n; simpl in *; 120 | discriminate || intuition. 121 | Qed. 122 | 123 | Lemma ssubst_type : forall M K A B Δ 124 | (HTK : nil ⊢ K ::: A) 125 | (HTM : Δ ++ A :: nil ⊢ M ::: B), 126 | Δ ⊢ [K ↑ length Δ]M ::: B. 127 | Proof. 128 | intros; remember (Δ ++ A :: nil) as Γ; generalize dependent Δ; 129 | induction HTM; simpl in *; intros; subst; simpl in *; eauto using types; [| | |]. 130 | destruct (eq_nat_dec (length Δ) n). 131 | subst; assert (HEq : A = A0); [| subst A0]. 132 | induction Δ; simpl in *; [inversion HFind; subst; tauto |]. 133 | apply IHΔ; assumption. 134 | replace Δ with (nil ++ Δ) by reflexivity; apply weaken; assumption. 135 | apply tc_var. 136 | generalize dependent n; induction Δ; simpl in *; intros. 137 | destruct n; [tauto | destruct n; discriminate]. 138 | destruct n; simpl in *; [assumption |]. 139 | apply IHΔ; [assumption |]. 140 | intro HEq; apply n0; f_equal; assumption. 141 | simpl; eapply tc_lam, IHHTM; tauto. 142 | simpl; apply tc_rec; [apply IHHTM1 | apply IHHTM2 | apply IHHTM3]; tauto. 143 | simpl; apply tc_seed; [apply IHHTM1 | apply IHHTM2 | apply IHHTM3]; tauto. 144 | Qed. 145 | 146 | Lemma subst_types : forall γ Γ Δ M A 147 | (HTC : γ :–: Γ) 148 | (HT : Δ ++ Γ ⊢ M ::: A), 149 | Δ ⊢ [γ!length Δ]M ::: A. 150 | Proof. 151 | induction γ; destruct Γ; simpl in *; intros; try contradiction; [|]. 152 | rewrite <- app_nil_end in HT; assumption. 153 | destruct HTC as [HTa HTC]. 154 | eapply ssubst_type; [eassumption |]. 155 | replace (S (length Δ)) with (length (Δ ++ t :: nil)) 156 | by (rewrite app_length, plus_comm; reflexivity). 157 | eapply IHγ; [eassumption |]. 158 | rewrite <- app_assoc; simpl; assumption. 159 | Qed. 160 | 161 | Lemma closed_sub : forall M A K n Γ 162 | (HT : Γ ⊢ M ::: A), 163 | M = [K ↑ length Γ + n]M. 164 | Proof. 165 | induction 1; try (simpl in *; f_equal; reflexivity || assumption). 166 | simpl. 167 | assert (n0 < length Γ). 168 | generalize dependent n0; induction Γ; intros; [destruct n0; discriminate |]. 169 | destruct n0; simpl; [auto with arith |]. 170 | simpl in HFind; apply IHΓ in HFind; auto with arith. 171 | destruct (eq_nat_dec (length Γ + n) n0); [| reflexivity]. 172 | contradict e; omega. 173 | Qed. 174 | 175 | Lemma subst_var : forall γ Γ n m A 176 | (HT : nth_error Γ n = Some A) 177 | (HΓ : tcmt γ Γ), 178 | Some [γ ! m](#(n+m)) = nth_error γ n /\ (nil ⊢ [γ ! m](#(n+m)) ::: A). 179 | Proof. 180 | induction γ; intros. 181 | destruct Γ; [destruct n |]; contradiction || discriminate. 182 | destruct Γ; [contradiction |]. 183 | destruct n; simpl in *. 184 | rewrite <- subst_gt; [| auto with arith]. 185 | unfold Specif.value; simpl; destruct (eq_nat_dec m m); 186 | inversion HT; subst; tauto. 187 | destruct HΓ as [Ht HΓ]; specialize (IHγ _ _ (S m) _ HT HΓ). 188 | rewrite plus_comm in IHγ; simpl in *; rewrite plus_comm in IHγ. 189 | destruct IHγ as [HEq HTy]; split. 190 | rewrite <- HEq; f_equal; symmetry. 191 | replace m with (length (@nil ty) + m) by reflexivity; eapply closed_sub; 192 | simpl; eassumption. 193 | replace m with (length (@nil ty) + m) by reflexivity. 194 | erewrite <- closed_sub; simpl in *; eassumption. 195 | Qed. 196 | 197 | (* Lemmas extending congruence to multistep reduction *) 198 | Lemma rec_cong_star : forall M M' M₀ M₁ 199 | (HR : M ↦* M'), 200 | rec M M₀ M₁ ↦* rec M' M₀ M₁. 201 | Proof. 202 | induction 1; [constructor |]. 203 | econstructor 2; [apply red_recC; eassumption | apply IHHR]. 204 | Qed. 205 | Lemma hd_cong_star : forall M M' 206 | (HRed : M ↦* M'), 207 | hd M ↦* hd M'. 208 | Proof. 209 | induction 1; [constructor |]. 210 | econstructor 2; [| apply IHHRed]. 211 | apply red_hdC; assumption. 212 | Qed. 213 | Lemma tl_cong_star : forall M M' 214 | (HRed : M ↦* M'), 215 | tl M ↦* tl M'. 216 | Proof. 217 | induction 1; [constructor |]. 218 | econstructor 2; [| apply IHHRed]. 219 | apply red_tlC; assumption. 220 | Qed. 221 | 222 | (* Progress lemma is not really needed to prove termination, 223 | but it's a good sanity check *) 224 | Lemma progress : forall M A 225 | (HT : nil ⊢ M ::: A), 226 | value M \/ exists N, M ↦ N. 227 | Proof. 228 | induction M; intros; simpl in *; inversion HT; subst; 229 | intuition eauto using value. 230 | (* var *) 231 | destruct n; discriminate. 232 | (* app *) 233 | right; specialize (IHM1 _ HTM); clear IHM2 HTN. 234 | destruct IHM1 as [HV | [N HRed]]. 235 | inversion HV; subst; try (inversion HTM; fail); []. 236 | eexists; apply red_β. 237 | exists (N @ M2); apply red_appC; assumption. 238 | (* rec *) 239 | right; specialize (IHM1 _ HTM); destruct IHM1 as [HV | [N HRed]]. 240 | inversion HV; subst; try (inversion HTM; fail). 241 | eexists; apply red_recz. 242 | eexists; apply red_recs. 243 | eexists; apply red_recC; eassumption. 244 | (* hd *) 245 | right; specialize (IHM _ HTM); destruct IHM as [HV | [N HRed]]. 246 | inversion HV; subst; try (inversion HTM; fail). 247 | eexists; apply red_hds. 248 | eexists; apply red_hdC; eassumption. 249 | (* tl *) 250 | right; specialize (IHM _ HTM); destruct IHM as [HV | [N HRed]]. 251 | inversion HV; subst; try (inversion HTM; fail). 252 | eexists; apply red_tls. 253 | eexists; apply red_tlC; eassumption. 254 | Qed. 255 | 256 | (* Type preservation is needed to use the head expansion lemma *) 257 | Lemma preservation : forall A M N 258 | (HT : nil ⊢ M ::: A) 259 | (HR : M ↦ N), 260 | nil ⊢ N ::: A. 261 | Proof. 262 | intros; remember (@nil ty) as Γ; generalize dependent N; 263 | induction HT; intros; inversion HR; subst. 264 | (* beta *) 265 | clear IHHT1 IHHT2 HR. 266 | inversion HT1; subst; clear HT1. 267 | eapply subst_types; simpl; [| eassumption]; simpl; tauto. 268 | (* cong app *) 269 | eapply tc_app; [apply IHHT1 |]; tauto. 270 | (* rec z *) 271 | assumption. 272 | (* rec s *) 273 | clear IHHT1 IHHT2 IHHT3 HR. 274 | inversion HT1; subst. 275 | eapply subst_types; [| simpl; eassumption]; simpl; intuition; []. 276 | apply tc_rec; assumption. 277 | (* cong rec *) 278 | apply tc_rec; [apply IHHT1 | |]; tauto. 279 | (* hd seed *) 280 | inversion HT; subst. 281 | eapply ssubst_type; simpl in *; eassumption. 282 | (* cong hd *) 283 | apply tc_hd; apply IHHT; tauto. 284 | (* tl seed *) 285 | inversion HT; subst. 286 | apply tc_seed; [eapply ssubst_type; simpl in *; eassumption | |]; assumption. 287 | (* cong tl *) 288 | apply tc_tl; apply IHHT; tauto. 289 | Qed. 290 | 291 | End Properties. 292 | 293 | Close Scope t_scope. 294 | Close Scope list_scope. 295 | -------------------------------------------------------------------------------- /Substitution.v: -------------------------------------------------------------------------------- 1 | (** This file defines the single and multiple substitutions. Since we only ever 2 | substitute closed terms, multiple substitution can be defined in terms of the 3 | single one. *) 4 | 5 | Require Export Syntax. 6 | Require Export Arith. 7 | Open Scope t_scope. 8 | Open Scope list_scope. 9 | 10 | Reserved Notation " [ γ ! n ] M " (at level 5). 11 | Reserved Notation " [ K ↑ n ] M " (at level 5). 12 | 13 | (* Substitution of index n for term K in M *) 14 | Fixpoint subt K n M := 15 | match M with 16 | | #m => if (eq_nat_dec n m) then K else M 17 | | λ A, M => λ A, [K ↑ S n]M 18 | | M @ N => [K ↑ n]M @ [K ↑ n]N 19 | | z => z 20 | | s M => s [K ↑ n]M 21 | | rec M M₀ M₁ => rec [K ↑ n]M [K ↑ n]M₀ [K ↑ S (S n)]M₁ 22 | | hd M => hd [K ↑ n]M 23 | | tl M => tl [K ↑ n]M 24 | | seed M M₀ M₁ => seed [K ↑ n]M [K ↑ S n]M₀ [K ↑ S n]M₁ 25 | | TT => TT 26 | | FF => FF 27 | end where " [ K ↑ n ] M " := (subt K n M) : t_scope. 28 | 29 | (* Substitution for the list γ, starting from index n in M *) 30 | Fixpoint sub (γ : list te) (n : nat) (M : te) := 31 | match γ with 32 | | nil => M 33 | | K :: γ => [K ↑ n][γ ! S n]M 34 | end where " [ γ ! n ] M " := (sub γ n M) : t_scope. 35 | 36 | (* Substitutions are composable if they don't miss an index *) 37 | Lemma subcomp : forall γ δ n M, 38 | ([γ ! n][δ ! n + length γ]M) = [γ ++ δ ! n]M. 39 | Proof. 40 | induction γ; intros; simpl in *; rewrite plus_comm; simpl; [tauto |]. 41 | rewrite plus_comm, IHγ; reflexivity. 42 | Qed. 43 | 44 | (* Boring lemmas about pushing substitutions through constructors *) 45 | 46 | Lemma subst_gt : forall γ n m 47 | (Hlt : n < m), 48 | #n = [γ ! m](#n). 49 | Proof. 50 | induction γ; simpl; intros; [reflexivity |]. 51 | rewrite <- IHγ; [| auto with arith]. 52 | simpl. 53 | destruct (eq_nat_dec m n); [| reflexivity]. 54 | subst; contradict Hlt; auto with arith. 55 | Qed. 56 | Lemma sub_lam : forall γ n A M, 57 | ([γ ! n]λ A, M) = λ A, [γ ! S n]M. 58 | Proof. 59 | induction γ; intros; simpl in *; [reflexivity |]. 60 | rewrite IHγ; simpl; reflexivity. 61 | Qed. 62 | Lemma sub_app : forall γ n M N, 63 | ([γ ! n](M @ N)) = [γ ! n]M @ [γ ! n]N. 64 | Proof. 65 | induction γ; intros; simpl in *; [reflexivity |]. 66 | rewrite IHγ; simpl; reflexivity. 67 | Qed. 68 | Lemma sub_z : forall γ n, 69 | ([γ ! n]z) = z. 70 | Proof. 71 | induction γ; intros; simpl in *; [reflexivity |]. 72 | rewrite IHγ; simpl; reflexivity. 73 | Qed. 74 | Lemma sub_s : forall γ n M, 75 | ([γ ! n](s M)) = s [γ ! n]M. 76 | Proof. 77 | induction γ; intros; simpl in *; [reflexivity |]. 78 | rewrite IHγ; simpl; reflexivity. 79 | Qed. 80 | Lemma sub_rec : forall γ n M M₀ M₁, 81 | ([γ ! n](rec M M₀ M₁)) = rec [γ ! n]M [γ ! n]M₀ [γ ! S (S n)]M₁. 82 | Proof. 83 | induction γ; intros; simpl in *; [reflexivity |]. 84 | rewrite IHγ; simpl; reflexivity. 85 | Qed. 86 | Lemma sub_hd : forall γ n M, 87 | ([γ ! n](hd M)) = hd [γ ! n]M. 88 | Proof. 89 | induction γ; intros; simpl in *; [reflexivity |]. 90 | rewrite IHγ; simpl; reflexivity. 91 | Qed. 92 | Lemma sub_tl : forall γ n M, 93 | ([γ ! n](tl M)) = tl [γ ! n]M. 94 | Proof. 95 | induction γ; intros; simpl in *; [reflexivity |]. 96 | rewrite IHγ; simpl; reflexivity. 97 | Qed. 98 | Lemma sub_seed : forall γ n M M₀ M₁, 99 | ([γ ! n](seed M M₀ M₁)) = seed [γ ! n]M [γ ! S n]M₀ [γ ! S n]M₁. 100 | Proof. 101 | induction γ; intros; simpl in *; [reflexivity |]. 102 | rewrite IHγ; simpl; reflexivity. 103 | Qed. 104 | Lemma sub_TT : forall γ n, 105 | ([γ ! n]TT) = TT. 106 | Proof. 107 | induction γ; intros; simpl in *; [reflexivity |]. 108 | rewrite IHγ; simpl; reflexivity. 109 | Qed. 110 | Lemma sub_FF : forall γ n, 111 | ([γ ! n]FF) = FF. 112 | Proof. 113 | induction γ; intros; simpl in *; [reflexivity |]. 114 | rewrite IHγ; simpl; reflexivity. 115 | Qed. 116 | 117 | Close Scope list_scope. 118 | Close Scope t_scope. -------------------------------------------------------------------------------- /Syntax.v: -------------------------------------------------------------------------------- 1 | (** This file defines the syntax of the System T extended with stream types. *) 2 | 3 | (* Types *) 4 | Inductive ty := 5 | | tbool : ty 6 | | tnat : ty 7 | | tarrow : ty -> ty -> ty 8 | | stream : ty -> ty. 9 | 10 | Notation " A '→' B " := (tarrow A B) (at level 30, right associativity) : t_scope. 11 | Notation " 'ω' " := tnat : t_scope. 12 | Notation " '2' " := tbool : t_scope. 13 | Open Scope t_scope. 14 | 15 | (* Terms *) 16 | (* The formalisation uses de Bruijn indices to deal with binders. *) 17 | Inductive te := 18 | | var : nat -> te 19 | | lam : ty -> te -> te 20 | | appl : te -> te -> te 21 | | z : te 22 | | s : te -> te 23 | | rec : te -> te -> te -> te 24 | | hd : te -> te 25 | | tl : te -> te 26 | | seed : te -> te -> te -> te 27 | | TT : te 28 | | FF : te. 29 | 30 | Notation "'λ' A , M " := (lam A M) (at level 50) : t_scope. 31 | Notation " # n " := (var n) (at level 20) : t_scope. 32 | Notation " M @ N " := (appl M N) (at level 40, left associativity) : t_scope. 33 | 34 | Close Scope t_scope. -------------------------------------------------------------------------------- /Termination.v: -------------------------------------------------------------------------------- 1 | (** This file contains the termination argument via logical relations *) 2 | 3 | Require Export Semantics. 4 | Open Scope t_scope. 5 | Open Scope list_scope. 6 | 7 | Section Definitions. 8 | 9 | (* Logical relation at 2 *) 10 | Inductive R2 : te -> Prop := 11 | | RT : forall M (HRed : M ↦* TT), R2 M 12 | | RF : forall M (HRed : M ↦* FF), R2 M. 13 | 14 | (* Logical relation at type ω *) 15 | Inductive Rω : te -> Prop := 16 | | Rω_z : forall M (HRed : M ↦* z), Rω M 17 | | Rω_s : forall M M' (HRed : M ↦* s M') (HR : Rω M'), Rω M. 18 | 19 | (* Logical relation at type stream, parameterised with the relation at element type *) 20 | CoInductive Rstream : (te -> Type) -> te -> Prop := 21 | | Rs : forall M P (HHD : P (hd M)) (HTL : Rstream P (tl M)), Rstream P M. 22 | 23 | (* The reducibility relation: has to be defined as a fixpoint, since Coq doesn't allow 24 | for negative occurrences of a relation in the A₁ → A₂ case *) 25 | Fixpoint R A M : Prop := 26 | match A with 27 | | 2 => R2 M 28 | | ω => Rω M 29 | | A₁ → A₂ => forall N (HTN : nil ⊢ N ::: A₁) (HR : R A₁ N), R A₂ (M @ N) 30 | | stream A => Rstream (R A) M 31 | end. 32 | 33 | (* Reduciblity for the whole contexts *) 34 | Fixpoint rctx γ Γ := 35 | match γ, Γ with 36 | | nil, nil => True 37 | | M :: γ, A :: Γ => R A M /\ rctx γ Γ 38 | | _, _ => False 39 | end. 40 | 41 | End Definitions. 42 | 43 | Section Termination_proof. 44 | 45 | Lemma head_expansion : forall A M N 46 | (HR : R A N) 47 | (HS : M ↦ N), 48 | R A M. 49 | Proof. 50 | induction A; intros; simpl in *. 51 | (* bool *) 52 | inversion HR; [ apply RT | apply RF ]; econstructor; eassumption. 53 | (* nat *) 54 | inversion HR. 55 | apply Rω_z; unfold steps; econstructor; eassumption. 56 | apply Rω_s with M'; [| assumption]; econstructor; eassumption. 57 | (* arr *) 58 | intros. 59 | eapply IHA2; [| apply red_appC; eassumption]. 60 | apply HR; assumption. 61 | (* stream *) 62 | (* We need coinduction to prove the property here, and a stronger lemma. *) 63 | assert (HT : exists K : te, (M ↦ K) /\ Rstream (R A) K) 64 | by (exists N; tauto). 65 | clear HS HR; destruct HT as [K [HRed HR]]; generalize dependent K; revert M. 66 | cofix; intros. 67 | inversion HR; subst; apply Rs. 68 | eapply IHA; [eassumption | apply red_hdC; eassumption]. 69 | eapply head_expansion; [| eassumption]. 70 | apply red_tlC; assumption. 71 | Qed. 72 | 73 | (* lifting of head expansion to reflexive transitive closure *) 74 | Lemma head_exp_star : forall A M N 75 | (HR : R A N) 76 | (HS : M ↦* N), 77 | R A M. 78 | Proof. 79 | intros; induction HS; [assumption |]. 80 | eapply head_expansion; [| eassumption]. 81 | tauto. 82 | Qed. 83 | 84 | Lemma types_red : forall Γ M A γ 85 | (HT : Γ ⊢ M ::: A) 86 | (HR : rctx γ Γ) 87 | (HΓ : tcmt γ Γ), 88 | R A [γ!0]M. 89 | Proof. 90 | intros; generalize dependent γ; induction HT; intros. 91 | (* var *) 92 | assert (HT := subst_var _ _ _ 0 _ HFind HΓ); rewrite plus_comm in HT; simpl in *. 93 | destruct HT as [HT _]. 94 | revert HT; generalize ([γ ! 0](#n)) as M; intros M; generalize dependent Γ; 95 | revert n; induction γ; simpl; destruct Γ; intros; try contradiction; [|]. 96 | destruct n; discriminate. 97 | destruct n; simpl in *. 98 | inversion HT; subst a. 99 | inversion HFind; subst t; tauto. 100 | destruct HR as [Ha HR]; destruct HΓ as [Ht HΓ]; eapply IHγ; eassumption. 101 | (* lam *) 102 | simpl; intros; rewrite sub_lam. 103 | eapply head_expansion; [| apply red_β]. 104 | rewrite subcomp; apply IHHT; simpl; tauto. 105 | (* appl *) 106 | rewrite sub_app; simpl in IHHT1. 107 | apply IHHT1; intuition; []. 108 | eapply subst_types; simpl; eassumption. 109 | (* z *) 110 | apply Rω_z; rewrite sub_z; constructor. 111 | (* s *) 112 | simpl in *; rewrite sub_s; eapply Rω_s; [constructor | intuition]. 113 | (* rec *) 114 | specialize (IHHT1 _ HR HΓ); specialize (IHHT2 _ HR HΓ); simpl in *. 115 | assert (HTR : nil ⊢ [γ ! 0]M ::: ω) by (eapply subst_types; simpl; eassumption). 116 | rewrite sub_rec; revert IHHT1 HTR; generalize ([γ ! 0]M); intros K HRK HTK; 117 | induction HRK. 118 | apply head_exp_star with (rec z [γ ! 0]M₀ [γ ! 2]M₁); 119 | [| apply rec_cong_star; assumption]. 120 | eapply head_expansion; [| apply red_recz]; assumption. 121 | eapply head_exp_star with (rec (s M') [γ ! 0]M₀ [γ ! 2]M₁); 122 | [| apply rec_cong_star; assumption]. 123 | eapply head_expansion; [| apply red_recs]. 124 | assert (HST : nil ⊢ s M' ::: ω). 125 | induction HRed; [assumption |]. 126 | eapply IHHRed, preservation; eassumption. 127 | inversion HST; subst. 128 | rewrite subcomp; apply IHHT3; simpl; intuition; []. 129 | apply tc_rec; [assumption | |]; eapply subst_types; simpl; eassumption. 130 | (* hd *) 131 | specialize (IHHT _ HR HΓ); inversion IHHT; subst. 132 | rewrite sub_hd; assumption. 133 | (* tl *) 134 | specialize (IHHT _ HR HΓ); inversion IHHT; subst. 135 | rewrite sub_tl; assumption. 136 | (* seed *) 137 | specialize (IHHT1 _ HR HΓ). 138 | assert (HTR : nil ⊢ [γ ! 0]M ::: A) by (eapply subst_types; simpl; eassumption). 139 | clear HT1; simpl; rewrite sub_seed. 140 | assert (exists K, (seed [γ ! 0]M [γ ! 1]M₀ [γ ! 1]M₁ ↦* seed K [γ ! 1]M₀ [γ ! 1]M₁) 141 | /\ (nil ⊢ K ::: A) /\ R A K). 142 | simpl; eexists; split; [constructor | split; assumption]. 143 | simpl in H; destruct H as [K [HRed [HTK HRK]]]; generalize dependent K. 144 | generalize (seed [γ ! 0]M [γ ! 1]M₀ [γ ! 1]M₁) as N. 145 | cofix; intros; apply Rs. 146 | eapply head_exp_star; [| eapply hd_cong_star; eassumption]. 147 | eapply head_expansion; [| apply red_hds]. 148 | change (R A [K :: nil ! 0]([γ ! 1]M₀)); rewrite subcomp. 149 | apply IHHT2; simpl; tauto. 150 | apply types_red with [K ↑ 0]([γ ! 1]M₁); clear types_red. 151 | apply clos_rt_rt1n; eapply rt_trans; 152 | [apply clos_rt1n_rt, tl_cong_star; eassumption |]. 153 | apply rt_step, red_tls. 154 | change (nil ⊢ [K :: nil ! 0]([γ ! 1]M₁) ::: A); rewrite subcomp. 155 | eapply subst_types; [| simpl; eassumption]; simpl; tauto. 156 | change (R A [K :: nil ! 0]([γ ! 1]M₁)); rewrite subcomp. 157 | apply IHHT3; simpl; tauto. 158 | (* TT *) 159 | rewrite sub_TT; apply RT; constructor. 160 | (* FF *) 161 | rewrite sub_FF; apply RF; constructor. 162 | Qed. 163 | 164 | Theorem termination : forall M (HT : nil ⊢ M ::: ω), 165 | exists V, value V /\ (M ↦* V). 166 | Proof. 167 | intros. 168 | apply types_red with (γ := nil) in HT; [| exact I | exact I]. 169 | inversion HT; subst. 170 | (* z *) 171 | exists z; split; [apply val_z | assumption]. 172 | (* s *) 173 | exists (s M'); split; [apply val_s | assumption]. 174 | Qed. 175 | 176 | End Termination_proof. 177 | --------------------------------------------------------------------------------