From 799a95e79c305a8872eaf67da6b4806aec13be93 Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Fri, 7 Jun 2024 14:10:47 +0200 Subject: [PATCH 01/15] reflection --- _CoqProject | 7 +- theories/reduction/Reduction.v | 306 +++++++++++++ theories/reduction/ReductionAndTransitivity.v | 252 +++++++++++ theories/reduction/ReductionConfluence.v | 64 +++ theories/reduction/ReductionProperties.v | 267 +++++++++++ theories/reduction/ReductionRho.v | 420 ++++++++++++++++++ theories/reduction/ReductionToCongruence.v | 179 ++++++++ 7 files changed, 1494 insertions(+), 1 deletion(-) create mode 100644 theories/reduction/Reduction.v create mode 100644 theories/reduction/ReductionAndTransitivity.v create mode 100644 theories/reduction/ReductionConfluence.v create mode 100644 theories/reduction/ReductionProperties.v create mode 100644 theories/reduction/ReductionRho.v create mode 100644 theories/reduction/ReductionToCongruence.v diff --git a/_CoqProject b/_CoqProject index e75baf3..113e8d4 100644 --- a/_CoqProject +++ b/_CoqProject @@ -20,6 +20,12 @@ theories/TermNotations.v theories/Typing.v theories/BasicMetaTheory.v +theories/reduction/Reduction.v +theories/reduction/ReductionProperties.v +theories/reduction/ReductionRho.v +theories/reduction/ReductionAndTransitivity.v + + theories/CScoping.v theories/CTyping.v theories/CCMetaTheory.v @@ -32,7 +38,6 @@ theories/Model.v theories/Admissible.v theories/FreeTheorem.v -theories/Examples.v theories/RTyping.v theories/Potential.v diff --git a/theories/reduction/Reduction.v b/theories/reduction/Reduction.v new file mode 100644 index 0000000..2a4c989 --- /dev/null +++ b/theories/reduction/Reduction.v @@ -0,0 +1,306 @@ +(* Definition of reduction rules which corresponds to the congruence relation *) +(* and proof that the system is confluent *) +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. +From GhostTT Require Export Univ TermNotations Typing. + +Import ListNotations. + +Set Default Goal Selector "!". + +Reserved Notation "Γ ⊨ u ↣ v" + (at level 80, u, v at next level, format "Γ ⊨ u ↣ v"). +Notation "s '··'" := (scons s ids) (at level 1, format "s ··") : subst_scope. +Notation "Γ ⊢ A ∷ m" := (scoping Γ A m) + (at level 80, A, m at next level, format "Γ ⊢ A ∷ m"). +Notation " [ Γ , s ] ⊢ A ∷ m" := (scoping (s::Γ) A m) + (at level 81, A, m at next level, format "[ Γ , s ] ⊢ A ∷ m"). + +Notation ℙ := mProp. +Notation 𝔾 := mGhost. +Notation 𝕋 := mType. +Notation 𝕂 := mKind. + +Notation "⋆" := + (lam ℙ bot (var 0)). + +(* ------------------------------------------------------------------------- *) +Section Properties_Star. + + Lemma type_star : ∀ Γ, Γ ⊢ ⋆ : top. + Proof. + intro Γ. + apply type_lam. + all: eauto using scope_bot, type_bot. + - apply scope_var; reflexivity. + - eapply type_var; reflexivity. + Qed. + + Lemma scope_star : ∀ Γ, Γ ⊢ ⋆ ∷ ℙ. + Proof. + intro. + apply scope_lam. + - apply scope_bot. + - apply scope_var; reflexivity. + Qed. + +End Properties_Star. + +(* ------------------------------------------------------------------------- *) +Section Rewriting. + (** rewriting rules **) + + Definition mode_eq : ∀ (x y : mode), { x = y } + { x ≠ y }. + Proof. + decide equality. + Defined. + + Definition red_lvl (m : mode) (i : level) : level := + if mode_eq m ℙ then 0 else i. + + + Inductive reduction (Γ : scope) : term → term → Prop := + + (* Computation rules *) + + | red_beta : + ∀ mx A t t' u u', + (mx::Γ) ⊨ t ↣ t' → + md Γ u = mx → + Γ ⊨ u ↣ u' → + Γ ⊨ app (lam mx A t) u ↣ t' <[ u' ·· ] + + | red_reveal_hide : + ∀ t P p t' p', + Γ ⊨ t ↣ t' → + Γ ⊨ p ↣ p' → + In (md Γ p) [ℙ;𝔾] → + Γ ⊨ reveal (hide t) P p ↣ app p' t' + + | red_if_true : + ∀ m P t f t', + Γ ⊨ t ↣ t' → + md Γ t = m → + Γ ⊨ tif m ttrue P t f ↣ t' + + | red_if_false : + ∀ m P t f f', + Γ ⊨ f ↣ f' → + md Γ f = m → + Γ ⊨ tif m tfalse P t f ↣ f' + + | red_nat_elim_zero : + ∀ m P z s z', + Γ ⊨ z ↣ z' → + md Γ z = m → + Γ ⊨ tnat_elim m tzero P z s ↣ z' + + | red_nat_elim_succ : + ∀ m P z s n P' z' s' n', + Γ ⊨ n ↣ n' → + Γ ⊨ P ↣ P' → + Γ ⊨ z ↣ z' → + Γ ⊨ s ↣ s' → + md Γ s = m → + Γ ⊨ tnat_elim m (tsucc n) P z s ↣ app (app s' n') (tnat_elim m n' P' z' s') + + | red_vec_elim_nil : + ∀ m A B P z s z', + Γ ⊨ z ↣ z' → + md Γ z = m → + Γ ⊨ tvec_elim m A (hide tzero) (tvnil B) P z s ↣ z' + + | red_vec_elim_cons : + ∀ m A a n n0 v P z s A' a' n' v' P' z' s', + Γ ⊨ A ↣ A' → + Γ ⊨ a ↣ a' → + Γ ⊨ n ↣ n' → + Γ ⊨ v ↣ v' → + Γ ⊨ P ↣ P' → + Γ ⊨ z ↣ z' → + Γ ⊨ s ↣ s' → + md Γ s = m → + Γ ⊨ tvec_elim m A n0 (tvcons a n v) P z s ↣ + app (app (app (app s' a') (glength A' n' v')) v') (tvec_elim m A' n' v' P' z' s') + + (* Congruence rules *) + + (* A rule to quotient away all levels of Prop, making it impredicatime *) + | red_Prop : + ∀ i, Γ ⊨ Sort ℙ i ↣ Sort ℙ 0 + + | red_Pi : + ∀ i j m mx A A' B B', + Γ ⊨ A ↣ A' → + (mx::Γ) ⊨ B ↣ B' → + Γ ⊨ Pi i j m mx A B ↣ Pi (red_lvl mx i) (red_lvl m j) m mx A' B' + + | red_Pi' : (* needed for red_subst *) + ∀ i j m mx A A' B B', + Γ ⊨ A ↣ A' → + (mx::Γ) ⊨ B ↣ B' → + Γ ⊨ Pi i j m mx A B ↣ Pi i j m mx A' B' + + | red_lam : + ∀ mx A A' t t', + Γ ⊨ A ↣ A' → + (mx::Γ) ⊨ t ↣ t' → + Γ ⊨ lam mx A t ↣ lam mx A' t' + + | red_app : + ∀ u u' v v', + Γ ⊨ u ↣ u' → + Γ ⊨ v ↣ v' → + Γ ⊨ app u v ↣ app u' v' + + | red_Erased : + ∀ A A', + Γ ⊨ A ↣ A' → + Γ ⊨ Erased A ↣ Erased A' + + | red_hide : + ∀ u u', + Γ ⊨ u ↣ u' → + Γ ⊨ hide u ↣ hide u' + + | red_reveal : + ∀ t t' P P' p p', + Γ ⊨ t ↣ t' → + Γ ⊨ P ↣ P' → + Γ ⊨ p ↣ p' → + Γ ⊨ reveal t P p ↣ reveal t' P' p' + + | red_Reveal : + ∀ t t' p p', + Γ ⊨ t ↣ t' → + Γ ⊨ p ↣ p' → + Γ ⊨ Reveal t p ↣ Reveal t' p' + + | red_gheq : + ∀ A A' u u' v v', + Γ ⊨ A ↣ A' → + Γ ⊨ u ↣ u' → + Γ ⊨ v ↣ v' → + Γ ⊨ gheq A u v ↣ gheq A' u' v' + + | red_if : + ∀ m b b' P P' t t' f f', + Γ ⊨ b ↣ b' → + Γ ⊨ P ↣ P' → + Γ ⊨ t ↣ t' → + Γ ⊨ f ↣ f' → + Γ ⊨ tif m b P t f ↣ tif m b' P' t' f' + + | red_succ : + ∀ n n', + Γ ⊨ n ↣ n' → + Γ ⊨ tsucc n ↣ tsucc n' + + | red_nat_elim : + ∀ m n n' P P' z z' s s', + Γ ⊨ n ↣ n' → + Γ ⊨ P ↣ P' → + Γ ⊨ z ↣ z' → + Γ ⊨ s ↣ s' → + Γ ⊨ tnat_elim m n P z s ↣ tnat_elim m n' P' z' s' + + | red_vec : + ∀ A A' n n', + Γ ⊨ A ↣ A' → + Γ ⊨ n ↣ n' → + Γ ⊨ tvec A n ↣ tvec A' n' + + | red_vnil : + ∀ A A', + Γ ⊨ A ↣ A' → + Γ ⊨ tvnil A ↣ tvnil A' + + | red_vcons : + ∀ a a' n n' v v', + Γ ⊨ a ↣ a' → + Γ ⊨ n ↣ n' → + Γ ⊨ v ↣ v' → + Γ ⊨ tvcons a n v ↣ tvcons a' n' v' + + | red_vec_elim : + ∀ m A A' n n' v v' P P' z z' s s', + Γ ⊨ A ↣ A' → + Γ ⊨ n ↣ n' → + Γ ⊨ v ↣ v' → + Γ ⊨ P ↣ P' → + Γ ⊨ z ↣ z' → + Γ ⊨ s ↣ s' → + Γ ⊨ tvec_elim m A n v P z s ↣ tvec_elim m A' n' v' P' z' s' + + | red_bot_elim : + ∀ m A A' p p', + Γ ⊨ A ↣ A' → + Γ ⊨ p ↣ p' → + Γ ⊨ bot_elim m A p ↣ bot_elim m A' p' + + (* Structural rule *) + + | red_refl : + ∀ u, + Γ ⊨ u ↣ u + + (* Proof irrelevance *) + + | red_irr : + ∀ p, + md Γ p = ℙ → + Γ ⊨ p ↣ ⋆ + + | red_toRev : + ∀ t t' p p' u u', + Γ ⊨ t ↣ t' → + Γ ⊨ p ↣ p' → + Γ ⊨ u ↣ u' → + Γ ⊨ toRev t p u ↣ toRev t' p' u' + + | red_fromRev : + ∀ t t' p p' u u', + Γ ⊨ t ↣ t' → + Γ ⊨ p ↣ p' → + Γ ⊨ u ↣ u' → + Γ ⊨ fromRev t p u ↣ fromRev t' p' u' + + | red_ghrefl : + ∀ A A' u u', + Γ ⊨ A ↣ A' → + Γ ⊨ u ↣ u' → + Γ ⊨ ghrefl A u ↣ ghrefl A' u' + + | red_ghcast : + ∀ A A' u u' v v' e e' P P' t t', + Γ ⊨ A ↣ A' → + Γ ⊨ u ↣ u' → + Γ ⊨ v ↣ v' → + Γ ⊨ e ↣ e' → + Γ ⊨ P ↣ P' → + Γ ⊨ t ↣ t' → + Γ ⊨ ghcast A u v e P t ↣ ghcast A' u' v' e' P' t' + + + where "Γ ⊨ u ↣ v" := (reduction Γ u v). + + +End Rewriting. +Notation "Γ ⊨ u ↣ v" := (reduction Γ u v). + +(* ------------------------------------------------------------------------- *) +(** rewriting automation **) +Create HintDb gtt_red discriminated. + +Hint Resolve red_beta red_reveal_hide red_if_true red_if_false red_nat_elim_zero + red_nat_elim_succ red_vec_elim_nil red_vec_elim_cons red_Prop red_Pi + red_lam red_app red_Erased red_hide red_reveal red_Reveal red_gheq + red_if red_succ red_nat_elim red_vec red_vnil red_vcons red_vec_elim + red_bot_elim red_refl red_irr + red_Pi' red_toRev red_fromRev red_ghrefl red_ghcast + : gtt_red. + +Ltac gred := + unshelve typeclasses eauto with gtt_scope gtt_red shelvedb ; shelve_unifiable. +(** end rewriting automation **) diff --git a/theories/reduction/ReductionAndTransitivity.v b/theories/reduction/ReductionAndTransitivity.v new file mode 100644 index 0000000..449b3eb --- /dev/null +++ b/theories/reduction/ReductionAndTransitivity.v @@ -0,0 +1,252 @@ +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval + TermMode Scoping BasicMetaTheory. +From GhostTT.reduction Require Import ReductionProperties. +From GhostTT.reduction Require Export Reduction. + +Import ListNotations. + +Set Default Goal Selector "!". + +Inductive reduction_trans (Γ : scope) (u v: term) : Prop := + | Refl: u = v → reduction_trans Γ u v + | Trans w : Γ ⊨ u ↣ w → reduction_trans Γ w v → reduction_trans Γ u v. + +Notation "Γ ⊨ u ↣* v" := (reduction_trans Γ u v) + (at level 80, u, v at next level, format "Γ ⊨ u ↣* v"). + +Lemma red_trans_direct {Γ : scope } {u v: term} : Γ ⊨ u ↣ v → Γ ⊨ u ↣* v. +Proof. + refine ( λ H, Trans Γ u v v H (Refl Γ v v _)). + reflexivity. +Qed. + +Lemma red_trans_trans {Γ : scope} {u v: term} : + ∀ w, Γ ⊨ u ↣* w → Γ ⊨ w ↣* v → Γ ⊨ u ↣* v. +Proof. + intros w red_u red_w. + induction red_u as [ u | u w w' red_u red_w' IHw]. + - subst u; exact red_w. + - eapply Trans; eauto. +Qed. + +(* reds deinitions *) + +Local Ltac end_things H:= + induction H in H |- *; + [subst; apply Refl; reflexivity | + eapply Trans; [ gred | eassumption]]. + +Lemma reds_beta (Γ : scope) (mx : mode) (A t t' u u' : term) : + mx :: Γ⊨t↣*t'→ md Γ u = mx → Γ⊨u↣*u' → Γ⊨app (lam mx A t) u↣*t' <[u'··]. +Proof. + intros red_t scope_u red_u. + induction red_u. + - subst. induction red_t. + * subst. apply red_trans_direct. gred; reflexivity. + * eapply Trans; [ | eauto]. gred. + - eapply Trans; [ | erewrite red_md in scope_u; eauto]. gred. + +Qed. + +Lemma reds_reveal_hide (Γ : scope) (mp : mode) (t P p t' p' : term): + Γ⊨t↣*t' → Γ⊨p↣*p' → In (md Γ p) [ℙ;𝔾] → + Γ⊨reveal (hide t) P p↣*app p' t'. +Proof. + intros red_t red_p Hscope. + induction red_t. + - subst. induction red_p. + * subst. apply red_trans_direct. gred. + * eapply Trans. + + apply red_reveal; gred. + + erewrite red_md in Hscope; eauto. + - eapply Trans; [ | eauto]. gred. +Qed. + +(* Lemma reds_if_true *) +(* Lemma reds_if_false *) +(* Lemma reds_nat_elim_zero *) +(* Lemma reds_nat_elim_succ *) +(* Lemma reds_vec_elim_nil *) +(* Lemma reds_vec_elim_cons *) + +Lemma reds_Prop (Γ : scope) (i : level): Γ⊨Sort ℙ i↣*Sort ℙ 0. +Proof. + apply red_trans_direct. gred. +Qed. + +Lemma reds_Pi (Γ : scope) (i j : level) (m mx : mode) (A A' B B' : term) : + Γ⊨A↣*A' → mx :: Γ⊨B↣*B' → Γ⊨Pi i j m mx A B↣*Pi (red_lvl mx i) (red_lvl m j) m mx A' B'. +Proof. + intros red_A red_B. + induction red_A. + - subst. induction red_B. + * subst. apply red_trans_direct. gred. + * eapply Trans; [gred | eassumption]. + - eapply Trans; [gred | eassumption]. +Qed. + +Lemma reds_lam (Γ : scope) (mx : mode) (A A' t t' : term) : + Γ⊨A↣*A' → mx :: Γ⊨t↣*t' → Γ⊨lam mx A t↣*lam mx A' t'. +Proof. + intros red_A red_t. + induction red_A. + - subst. end_things red_t. + - eapply Trans; [ gred | eassumption]. +Qed. + +Lemma reds_app (Γ : scope) (u u' v v' : term) : + Γ⊨u↣*u' → Γ⊨v↣*v' → Γ⊨app u v↣*app u' v'. +Proof. + intros red_u red_v. + induction red_u. + - subst. end_things red_v. + - eapply Trans; [ gred | eassumption]. +Qed. + +Lemma reds_Erased (Γ : scope) (A A' : term) : + Γ⊨A↣*A' → Γ⊨Erased A↣*Erased A'. +Proof. +intro red_A; end_things red_A. +Qed. + +Lemma reds_hide (Γ : scope) (A A' : term) : + Γ⊨A↣*A' → Γ⊨hide A↣*hide A'. +Proof. +intro red_A; end_things red_A. +Qed. + +Lemma reds_reveal (Γ : scope) (t t' P P' p p' : term) : + Γ⊨t↣*t' → Γ⊨P↣*P' → Γ⊨p↣*p' → Γ⊨reveal t P p↣*reveal t' P' p'. +Proof. + intros red_t red_P red_p. + induction red_t. + - subst. induction red_P. + * subst. end_things red_p. + * eapply Trans; [ gred | eassumption]. + - eapply Trans; [ gred | eassumption]. +Qed. + +Lemma reds_Reveal (Γ : scope) (t t' p p' : term) : + Γ⊨t↣*t' → Γ⊨p↣*p' → Γ⊨Reveal t p↣*Reveal t' p'. +Proof. + intros red_t red_p. + induction red_t. + - subst. end_things red_p. + - eapply Trans; [ gred | eassumption]. +Qed. + +Lemma reds_gheq (Γ : scope) (A A' u u' v v' : term): + Γ⊨A↣*A' → Γ⊨u↣*u' → Γ⊨v↣*v' → Γ⊨gheq A u v↣*gheq A' u' v'. +Proof. + intros red_A red_u red_v. + induction red_A. + - subst. induction red_u. + * subst. end_things red_v. + * eapply Trans; [ gred | eassumption]. + - eapply Trans; [ gred | eassumption]. +Qed. + +Lemma reds_if (Γ : scope) (m : mode) (b b' P P' t t' f f' : term) : + Γ⊨b↣*b' → Γ⊨P↣*P' → Γ⊨t↣*t' → Γ⊨f↣*f' → Γ⊨tif m b P t f↣*tif m b' P' t' f'. +Proof. + intros red_b red_P red_t red_f. + induction red_b. + - subst. induction red_P. + * subst. induction red_t. + + subst. end_things red_f. + + eapply Trans; [ gred | eassumption]. + * eapply Trans; [ gred | eassumption]. + - eapply Trans; [ gred | eassumption]. +Qed. + + +Lemma reds_succ (Γ : scope) (n n' : term): + Γ⊨n↣*n' → Γ⊨tsucc n↣*tsucc n'. +Proof. + intro red_A; end_things red_A. +Qed. + +Lemma reds_nat_elim (Γ : scope) (m : mode) (n n' P P' z z' s s' : term) : + Γ⊨n↣*n' → Γ⊨P↣*P' → Γ⊨z↣*z' → Γ⊨s↣*s' → Γ⊨tnat_elim m n P z s↣*tnat_elim m n' P' z' s'. +Proof. + intros red_n red_P red_z red_s. + induction red_n. + - subst. induction red_P. + * subst. induction red_z. + + subst. end_things red_s. + + eapply Trans; [ gred | eassumption]. + * eapply Trans; [ gred | eassumption]. + - eapply Trans; [ gred | eassumption]. +Qed. + +Lemma reds_vec (Γ : scope) (A A' n n' : term) : + Γ⊨A↣*A' → Γ⊨n↣*n' → Γ⊨tvec A n↣*tvec A' n'. +Proof. + intros red_A red_n. + induction red_A. + - subst. end_things red_n. + - eapply Trans; [ gred | eassumption]. +Qed. + +Lemma reds_vnil (Γ : scope) (A A' : term) : + Γ⊨A↣*A' → Γ⊨tvnil A↣*tvnil A'. +Proof. + intro red_A; end_things red_A. +Qed. + +Lemma reds_vcons (Γ : scope) (a a' n n' v v' : term) : + Γ⊨a↣*a' → Γ⊨n↣*n' → Γ⊨v↣*v' → Γ⊨tvcons a n v↣*tvcons a' n' v'. +Proof. + intros red_a red_n red_v. + induction red_a. + - subst. induction red_n. + * subst. end_things red_v. + * eapply Trans; [ gred | eassumption]. + - eapply Trans; [ gred | eassumption]. +Qed. + +Lemma reds_vec_elim (Γ : scope) (m : mode) (A A' n n' v v' P P' z z' s s' : term): + Γ⊨A↣*A' → Γ⊨n↣*n' → Γ⊨v↣*v' → Γ⊨P↣*P' → Γ⊨z↣*z' → Γ⊨s↣*s' + → Γ⊨tvec_elim m A n v P z s↣*tvec_elim m A' n' v' P' z' s'. +Proof. + intros red_A red_n red_v red_P red_z red_s. + induction red_A. + - subst. induction red_n. + * subst. induction red_v. + + subst. induction red_P. + ++ subst. induction red_z. + +++ subst. end_things red_s. + +++ eapply Trans; [ gred | eassumption]. + ++ eapply Trans; [ gred | eassumption]. + + eapply Trans; [ gred | eassumption]. + * eapply Trans; [ gred | eassumption]. + - eapply Trans; [ gred | eassumption]. +Qed. + +Lemma reds_bot_elim (Γ : scope) (m : mode) (A A' p p' : term) : + Γ⊨A↣*A' → Γ⊨p↣*p' → Γ⊨bot_elim m A p↣*bot_elim m A' p'. +Proof. + intros red_A red_p. + induction red_A. + - subst. end_things red_p. + - eapply Trans; [ gred | eassumption]. +Qed. + + +(* ------------------------------------------------------------------------- *) +(** rewriting automation **) +Create HintDb gtt_reds discriminated. + +Hint Resolve reds_beta reds_reveal_hide (*reds_if_true reds_if_false reds_nat_elim_zero*) + (* reds_nat_elim_succ reds_vec_elim_nil reds_vec_elim_cons*) reds_Prop reds_Pi + reds_lam reds_app reds_Erased reds_hide reds_reveal reds_Reveal reds_gheq + reds_if reds_succ reds_nat_elim reds_vec reds_vnil reds_vcons reds_vec_elim + reds_bot_elim (*reds_irr*) + : gtt_reds. + +Ltac greds := + unshelve typeclasses eauto with gtt_scope gtt_reds gtt_red shelvedb ; shelve_unifiable. +(** end rewriting automation **) + diff --git a/theories/reduction/ReductionConfluence.v b/theories/reduction/ReductionConfluence.v new file mode 100644 index 0000000..fd67cc0 --- /dev/null +++ b/theories/reduction/ReductionConfluence.v @@ -0,0 +1,64 @@ +(* Definition of reduction rules which corresponds to the congruence relation *) +(* and proof that the system is confluent *) +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. +From GhostTT.reduction Require Import ReductionProperties ReductionRho. +From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. + +Import ListNotations. + +Set Default Goal Selector "!". + +Theorem reduction_diamond (Γ : scope) (t u v: term) : + Γ ⊨ t ↣ u → Γ ⊨ t ↣ v → + ∃ w, Γ ⊨ u ↣ w ∧ Γ ⊨ v ↣ w. +Proof. + intros red_tu red_tv. + exists (ρ Γ t). + eauto using reduction_triangle. +Qed. + +Theorem reduction_local_confluence (Γ : scope) (t u v: term) : + Γ ⊨ t ↣ u → Γ ⊨ t ↣ v → + ∃ w, Γ ⊨ u ↣* w ∧ Γ ⊨ v ↣* w. +Proof. + intros red_t0 red_t1. + exists (ρ Γ t). + split; apply red_trans_direct; eauto using reduction_triangle. +Qed. + +Theorem reduction_confluence (Γ : scope) (t u v: term) : + Γ ⊨ t ↣* u → Γ ⊨ t ↣* v → + ∃ w, Γ ⊨ u ↣* w ∧ Γ ⊨ v ↣* w. +Proof. + intros red_tu red_tv. + induction red_tu as [t u red_tu | t u1 u0 red_tu red_u0 IH] in u, v, red_tu, red_tv |- *. + - induction red_tv as [t v red_tv | t v1 v0 red_tv red_v0 IH'] in u, v, red_tu, red_tv |- *. + * subst. exists v; split; constructor; reflexivity. + * subst t. + exists v1. + split. + + apply (Trans Γ u v1 v0); assumption. + + constructor; reflexivity. + - induction red_tv as [t v red_tv | t v1 v0 red_tv red_v0 IH'] in u0, u1, v, red_tu, red_tv, red_u0, IH |- *. + * subst t. + exists u1. + split. + + constructor; reflexivity. + + apply (Trans Γ v u1 u0); assumption. + * assert (∃ w, (Γ⊨(ρ Γ t)↣*w) ∧ Γ⊨v1↣*w) as H. + { eapply IH'; eauto. + + eauto using reduction_triangle; eauto. + + constructor; reflexivity. + + intros. + eexists. split; [eassumption | apply red_trans_direct; gred]. } + clear IH'. + destruct H as [w_v [red_ρt red_v1]]. + specialize (Trans Γ _ w_v _ (reduction_triangle Γ _ _ red_tu) red_ρt) as red_u0'. + specialize (IH w_v red_u0'). + destruct IH as [w [red_u1' red_v1']]. + exists w; split. + + assumption. + + apply (red_trans_trans w_v); assumption. +Qed. diff --git a/theories/reduction/ReductionProperties.v b/theories/reduction/ReductionProperties.v new file mode 100644 index 0000000..06ebefb --- /dev/null +++ b/theories/reduction/ReductionProperties.v @@ -0,0 +1,267 @@ +(* Definition of reduction rules which corresponds to the congruence relation *) +(* and proof that the system is confluent *) +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. +From GhostTT.reduction Require Export Reduction. + +Import ListNotations. + +Set Default Goal Selector "!". + +Lemma md_ren' {Γ Δ :scope} {t: term} {ρ: nat → nat} (e : ∀ n, nth (ρ n) Γ 𝕋 = nth n Δ 𝕋) : + md Δ t = md Γ (ρ ⋅ t). +Proof. + induction t in Γ, Δ, t, ρ, e|- *. + all: cbn; eauto. + - cbn. match goal with H: ∀ _ _, _ → _ |- _ => + eapply H; eauto end. + intro n; destruct n; cbn; auto. + - match goal with H: ∀ _ _, _ → _ |- _ => + erewrite H; eauto end. +Qed. + +Lemma md_up_term {Γ : scope} {m: mode} {σ : nat → term} {n : nat} : + md (m::Γ) (up_term σ (S n)) = md Γ (σ n). +Proof. + asimpl; ssimpl. + unfold shift. + symmetry. + apply md_ren'. + induction n; eauto. +Qed. + +Lemma md_subst' {Γ Δ :scope} {t: term} {σ: nat → term} (e : ∀ n, md Γ (σ n) = nth n Δ 𝕋) : + md Δ t = md Γ (t<[σ]). +Proof. + induction t in Γ, Δ, t, σ, e|- *. + all: cbn; eauto. + - match goal with H: ∀ _ _, _ → _ |- _ => + erewrite H; eauto end. + intro n; destruct n; eauto. + erewrite md_up_term. auto. + - match goal with H: ∀ _ _, _ → _ |- _ => + erewrite H; eauto end. +Qed. + +Lemma red_md : + ∀ Γ t t', Γ ⊨ t ↣ t' → md Γ t = md Γ t'. +Proof. + intros Γ t t' red_t. + induction red_t in red_t |- *. + all: try solve [cbn; congruence]. + - cbn in *. eapply eq_trans; eauto. + erewrite md_subst'; eauto. + intro n. destruct n; eauto. + cbn. subst. auto. + - cbn. match goal with H: In _ _ |- _ => + destruct H as [ H0 | [ H0 |]] end. + 3: contradiction. + all: rewrite <- H0. + all: congruence. + - match goal with H: md ?Γ ?p = _ |- md ?Γ (reveal _ _ ?p) = _ => + cbn; rewrite H; reflexivity end. +Qed. + +Lemma glenght_red_subst (A n v : term) (σ : nat → term) : + (glength A n v)<[σ] = glength (A<[σ]) (n<[σ]) (v<[σ]). +Proof. + change (tvec_elim 𝔾 (A <[ σ]) (n <[ σ]) (v <[ σ]) + (lam 𝔾 (Erased tnat) + (lam 𝕋 ((tvec (S ⋅ A) (var 0))<[up_term σ]) (Erased tnat)) + ) + (hide tzero) + (lam 𝕋 (A<[σ]) + (lam 𝔾 (Erased tnat) + (lam 𝕋 (tvec (S ⋅ S ⋅ A) (var 0) <[up_term (up_term σ)]) + (lam 𝔾 (Erased tnat) + (gS (var 0)) + <[(up_term (up_term (up_term σ)))]) + ) + ) + ) + = glength (A<[σ]) (n<[σ]) (v<[σ])). + unfold glength. + repeat f_equal. + all: asimpl; reflexivity. +Qed. + +Lemma red_lam_inv {Γ : scope} {mx md_t : mode} {A t u: term} : + (Γ⊨lam mx A t↣ u ) → md Γ (lam mx A t) = md_t → (md_t ≠ ℙ) → + ( ∃ A' t', u = lam mx A' t' ∧ Γ⊨A↣A' ∧ mx :: Γ⊨t↣t'). +Proof. + intros prf_red scope_t not_Prop. + remember (lam mx A t) as lam_t eqn:e0. + remember u as u0 eqn:e1. + induction prf_red. + all: try solve [inversion e0]. + - inversion e0. + inversion e1; subst. + eauto. + - exists A, t; auto using red_refl. + - subst. contradiction. +Qed. + +Lemma red_hide_inv (Γ : scope) (t0 t' : term) (red_hide : Γ⊨hide t0 ↣t' ) : ∃ t0', t' = hide t0' ∧ Γ ⊨ t0 ↣ t0'. +Proof. + inversion red_hide; subst. + - eauto. + - eauto using red_refl. + - cbn in *. + match goal with | HC : 𝔾 = ℙ |- _ => inversion HC end. +Qed. + +Lemma red_succ_inv (Γ : scope) (n t' : term) (red_succ : Γ⊨tsucc n ↣t' ) : ∃ n', t' = tsucc n' ∧ Γ ⊨ n ↣ n'. +Proof. + inversion red_succ; subst. + - eauto. + - eauto using red_refl. + - cbn in *. + match goal with | HC : 𝕋 = ℙ |- _ => inversion HC end. +Qed. + +Lemma red_nil_inv (Γ : scope) (A t' : term) (red_nil : Γ ⊨ tvnil A ↣ t' ) : + ∃ A', t' = tvnil A' ∧ Γ ⊨ A ↣ A'. +Proof. + inversion red_nil; subst. + - eauto. + - eauto using red_refl. + - cbn in *. + match goal with | HC : 𝕋 = ℙ |- _ => inversion HC end. +Qed. + +Lemma red_cons_inv (Γ : scope) (a n v t' : term) (red_cons : Γ ⊨ tvcons a n v ↣ t' ) : + ∃ a' n' v', t' = tvcons a' n' v' ∧ Γ ⊨ a ↣ a' ∧ Γ ⊨ n ↣ n' ∧ Γ ⊨ v ↣ v'. +Proof. + inversion red_cons; subst. + - do 3 eexists; eauto. + - do 3 eexists; eauto using red_refl. + - cbn in *. + match goal with | HC : 𝕋 = ℙ |- _ => inversion HC end. +Qed. + +Lemma red_ren (Γ Δ : scope) (ρ: nat → nat) (t t': term) : + (∀ n, nth (ρ n) Γ 𝕋 = nth n Δ 𝕋) → + Δ ⊨ t ↣ t' → Γ ⊨ ρ ⋅ t ↣ ρ ⋅ t'. +Proof. + intros Hscope red_t. + induction red_t in Γ, Δ, ρ, Hscope, t, t', red_t |- *. + all: try solve [gred; erewrite <- md_ren'; eauto]. + - assert (∀ (t' u' : term), ((up_ren ρ) ⋅ t') <[ (ρ ⋅ u')··] = ρ ⋅ t' <[ u'··]) as er. + { intros. ssimpl. reflexivity. } + rewrite <- er. + gred; eauto. + + intro n; destruct n; cbn; auto. + + erewrite <- md_ren'; eauto. + - asimpl. + repeat rewrite (rinstInst'_term ρ). + rewrite glenght_red_subst. + repeat rewrite <- (rinstInst'_term ρ). + gred; erewrite <- md_ren'; eauto. + - gred. intro n; destruct n; cbn; auto. + - gred. intro n; destruct n; cbn; auto. + - gred. intro n; destruct n; cbn; auto. +Qed. + +Lemma up_subst_red (Γ : scope) (m : mode) (σ σ' : nat → term) : + (∀ n, Γ ⊨ σ n ↣ σ' n) → + (∀ n, m::Γ ⊨ up_term σ n ↣ up_term σ' n). +Proof. + intros Hyp n. + destruct n. + - ssimpl; gred. + - asimpl; ssimpl. + eapply (red_ren); eauto. + intro n0; destruct n0; cbn; auto. +Qed. + +Lemma red_subst_r (Γ : scope) (t : term) (σ σ' : nat → term) : + (∀ n, Γ ⊨ σ n ↣ σ' n) → + Γ ⊨ t <[σ] ↣ t <[σ']. +Proof. + intro red_σ. + induction t in Γ, σ, σ', red_σ |- *. + all: gred. + all: eauto using up_subst_red. +Qed. + +Lemma red_subst (Γ Δ : scope) (t t' : term) (σ σ' : nat → term) : + (∀ n, md Γ (σ n) = nth n Δ 𝕋) → + (∀ n, Γ ⊨ σ n ↣ σ' n) → + Δ ⊨ t ↣ t' → + Γ ⊨ t <[σ] ↣ t' <[σ']. +Proof. + intros Hscope red_σ red_t. + remember Δ as Δ0 eqn:e. + induction red_t in Γ, Δ, e, Hscope, σ, σ', red_σ, t, t', red_t |- *. + all: try solve [ gred; erewrite <- md_subst'; eauto using up_subst_red]. + - assert (∀ t' u', (t' <[ up_term σ']) <[ (u' <[ σ'])··] = (t' <[ u'··]) <[ σ']) as er. + { intros. ssimpl; apply ext_term. intro n; destruct n. + all: asimpl; reflexivity. } + rewrite <- er. + gred; eauto using up_subst_red. + * intro n; destruct n; auto. + rewrite md_up_term. cbn; auto. + * erewrite <- md_subst'; eauto. + - asimpl. erewrite glenght_red_subst. + gred; erewrite <- md_subst'; eauto using scoping_subst. + - gred; eauto using up_subst_red. + intro n; destruct n; [ |rewrite md_up_term]; auto. + - gred; eauto using up_subst_red. + intro n; destruct n; [ |rewrite md_up_term]; auto. + - gred; eauto using up_subst_red. + intro n; destruct n; [ |rewrite md_up_term]; auto. + - apply red_subst_r. assumption. +Qed. + +Ltac red_lam_inv_auto A' t' e red_A' red_t':= + match goal with + | red_lam : ?Γ⊨lam ?m ?A ?t ↣?u |- _ => + eapply red_lam_inv in red_lam; eauto; + destruct red_lam as [A' [t' [e [red_A' red_t']]]]; + try subst u + end. + +Ltac red_hide_inv_auto t0' e:= + match goal with + | red_hide : ?Γ⊨hide ?t0 ↣?t' |- _ => + apply red_hide_inv in red_hide; + destruct red_hide as [t0' [e red_hide]]; + try subst t' + end. + +Ltac red_succ_inv_auto n' e:= + match goal with + | red_succ : ?Γ⊨tsucc ?t0 ↣?t' |- _ => + apply red_succ_inv in red_succ; + destruct red_succ as [n' [e red_succ]]; + try subst t' + end. + +Ltac red_nil_inv_auto A' e:= + match goal with + | red_nil : ?Γ⊨tvnil ?A ↣?t' |- _ => + apply red_nil_inv in red_nil; + destruct red_nil as [A' [e red_nil]]; + try subst t' + end. + +Ltac red_conv_inv_auto a' n' v' e red_a' red_n' red_v':= + match goal with + | red_cons : ?Γ⊨tvcons ?a ?n ?v ↣?t' |- _ => + apply red_cons_inv in red_cons; + destruct red_cons as [a' [n' [v' [e [red_a' [red_n' red_v']]]]]]; + try subst t' + end. + +Ltac red_basic_inv := + let e := fresh "e" in + match goal with + | H : ?Γ ⊨ tzero ↣ ?t |- _ => + inversion H + | H : ?Γ ⊨ ttrue ↣ ?t |- _ => + inversion H + | H : ?Γ ⊨ tfalse ↣ ?t |- _ => + inversion H + end; subst. + diff --git a/theories/reduction/ReductionRho.v b/theories/reduction/ReductionRho.v new file mode 100644 index 0000000..0c5fc43 --- /dev/null +++ b/theories/reduction/ReductionRho.v @@ -0,0 +1,420 @@ +(* Definition of reduction rules which corresponds to the congruence relation *) +(* and proof that the system is confluent *) +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. +From GhostTT.reduction Require Import ReductionProperties. +From GhostTT.reduction Require Export Reduction. + +Import ListNotations. + +Set Default Goal Selector "!". + +From Equations Require Import Equations. +Derive NoConfusion Subterm for term. + +Section views_terms. + + (* term_view_app *) + Local Definition discr_term_lam (t : term) : Prop := + match t with + | lam _ _ _ => False + | _ => True + end. + Inductive term_view_app_i : term → term → Type := + | term_lam m A t u: term_view_app_i (lam m A t) u + | term_not_lam t u: discr_term_lam t → term_view_app_i t u. + + Equations term_view_app (t u: term) : term_view_app_i t u := + term_view_app (lam m A t) u := term_lam m A t u; + term_view_app t u := term_not_lam t u I. + + + (* term_view_hide *) + Local Definition discr_term_hide (t : term) : Prop := + match t with + | hide _ => False + | _ => True + end. + Inductive term_view_reveal_i : term → term → term → Type := + | term_hide t P p: term_view_reveal_i (hide t) P p + | term_not_hide t P p: discr_term_hide t → term_view_reveal_i t P p. + + Equations term_view_reveal (t P p : term) : term_view_reveal_i t P p := + term_view_reveal (hide t) P p := term_hide t P p ; + term_view_reveal t P p := term_not_hide t P p I. + + + (* term_view_if *) + Local Definition discr_term_true_false (t : term) : Prop := + match t with + | ttrue | tfalse => False + | _ => True + end. + Inductive term_view_if_i : term → term → term → term → Type := + | term_true P t f: term_view_if_i ttrue P t f + | term_false P t f: term_view_if_i tfalse P t f + | term_not_true_not_false b P t f: discr_term_true_false b → term_view_if_i b P t f. + + Equations term_view_if (b P t f : term) : term_view_if_i b P t f := + term_view_if ttrue P t f := term_true P t f ; + term_view_if tfalse P t f := term_false P t f ; + term_view_if b P t f := term_not_true_not_false b P t f I. + + + (* term_view_nat_elim *) + Local Definition discr_term_zero_succ (t : term) : Prop := + match t with + | tzero | tsucc _ => False + | _ => True + end. + Inductive term_view_nat_elim_i : term → term → term → term → Type := + | term_zero P z s: term_view_nat_elim_i tzero P z s + | term_succ n P z s: term_view_nat_elim_i (tsucc n) P z s + | term_not_zero_not_succ n P z s: discr_term_zero_succ n → term_view_nat_elim_i n P z s. + + Equations term_view_nat_elim (n P z s : term) : term_view_nat_elim_i n P z s := + term_view_nat_elim tzero P z s := term_zero P z s; + term_view_nat_elim (tsucc n) P z s := term_succ n P z s ; + term_view_nat_elim n P z s := term_not_zero_not_succ n P z s I. + + + (* term_view_vec_elim *) + Local Definition discr_term_vnil_vcons (t : term) : Prop := + match t with + | tvnil _ | tvcons _ _ _ => False + | _ => True + end. + Local Definition discr_term_hide_zero (t : term) : Prop := + match t with + | hide tzero => False + | _ => True + end. + Local Definition discr_term_for_vec_elim (n t : term) : Prop := discr_term_vnil_vcons t ∨ discr_term_hide_zero n. + Inductive term_view_vec_elim_i : term → term → term → term → term → term → Type := + | term_vnil A B P z s: term_view_vec_elim_i A (hide tzero) (tvnil B) P z s + | term_vcons A n1 a n0 v P z s: term_view_vec_elim_i A n1 (tvcons a n0 v) P z s + | term_other_vec_elim A n v P z s: discr_term_for_vec_elim n v → term_view_vec_elim_i A n v P z s. + Inductive sub_term_view_vec_elim_i : term → Type := + | sub_term_hide_zero : sub_term_view_vec_elim_i (hide tzero) + | sub_term_not_hide t : discr_term_hide_zero t → sub_term_view_vec_elim_i t. + + Equations sub_term_view_vec_elim n : sub_term_view_vec_elim_i n := + sub_term_view_vec_elim (hide tzero) := sub_term_hide_zero; + sub_term_view_vec_elim n := sub_term_not_hide n I. + + Equations term_view_vec_elim (A n v P z s : term) : term_view_vec_elim_i A n v P z s := + term_view_vec_elim A n (tvnil B) P z s with sub_term_view_vec_elim n := { + | sub_term_hide_zero := term_vnil A B P z s; + | sub_term_not_hide n Hn := term_other_vec_elim A n (tvnil B) P z s (or_intror Hn)}; + term_view_vec_elim A n1 (tvcons a n0 v) P z s := term_vcons A n1 a n0 v P z s ; + term_view_vec_elim A n v P z s := term_other_vec_elim A n v P z s (or_introl I). + +End views_terms. + + +(* ------------------------------------------------------------------------- *) +Section Rho. + + (* ------------------------------------------------------------------------- *) + (** Definition of rho **) + + Equations ρ (Γ : scope) (t : term) : term + by wf t term_subterm := + + (* cast *) + ρ Γ (ghcast A u v e P t) := + if mode_eq (md Γ t) ℙ then ⋆ + else ghcast (ρ Γ A) (ρ Γ u) (ρ Γ v) (ρ Γ e) (ρ Γ P) (ρ Γ t); + (* some cases when t ∷ ℙ *) + ρ Γ (var n) := + if mode_eq (md Γ (var n)) ℙ then ⋆ else var n; + ρ Γ (toRev _ _ _) := ⋆; + ρ Γ (fromRev _ _ _) := ⋆; + ρ Γ (ghrefl _ _) := ⋆; + + (* Beta reduction and app *) + ρ Γ (app t u) with (term_view_app t u) := { + | (term_lam mx A t1 u):= + if mode_eq (md Γ (lam mx A t1)) ℙ then ⋆ + else if mode_eq (md Γ u) mx then + (ρ (mx::Γ) t1) <[ (ρ Γ u)··] + else app (ρ Γ (lam mx A t1)) (ρ Γ u); + | (term_not_lam t u Ht) := + if mode_eq (md Γ t) ℙ then ⋆ + else app (ρ Γ t) (ρ Γ u)}; + + (* Reveal hide and reveal *) + ρ Γ (reveal t P p) with (term_view_reveal t P p):= { + | term_hide t P p := + if mode_eq (md Γ (reveal (hide t) P p)) ℙ then ⋆ + else app (ρ Γ p) (ρ Γ t); + | term_not_hide t P p Ht := + if mode_eq (md Γ (reveal t P p)) ℙ then ⋆ + else reveal (ρ Γ t) (ρ Γ P) (ρ Γ p)}; + + (* Sort ℙ i *) + ρ Γ (Sort m i) := Sort m (red_lvl m i); + + (* Cases where context changes *) + (* red _Pi_ℙ_ℙ *) + ρ Γ (Pi i j m mx A B) := + Pi (red_lvl mx i) (red_lvl m j) m mx (ρ Γ A) (ρ (mx::Γ) B); + (* red_lam *) + ρ Γ (lam mx A t) := + if mode_eq (md (mx::Γ) t) ℙ then ⋆ + else lam mx (ρ Γ A) (ρ (mx::Γ) t); + + (* One variable recursive cases *) + (* Erased *) + ρ Γ (Erased t) := Erased (ρ Γ t); + (* hide *) + ρ Γ (hide t) := hide (ρ Γ t); + (* tsucc *) + ρ Γ (tsucc n) := tsucc (ρ Γ n); + (* tvnil *) + ρ Γ (tvnil A) := tvnil (ρ Γ A); + + (* Two variables recursive cases *) + (* Reveal *) + ρ Γ (Reveal t P) := Reveal (ρ Γ t) (ρ Γ P); + (* vec *) + ρ Γ (tvec A n) := tvec (ρ Γ A) (ρ Γ n); + (* bot_elim *) + ρ Γ (bot_elim m A p) := + if mode_eq m ℙ then ⋆ + else bot_elim m (ρ Γ A) (ρ Γ p); + + (* Three variables recursive cases *) + (* vcons *) + ρ Γ (tvcons a n v) := tvcons (ρ Γ a) (ρ Γ n) (ρ Γ v); + (* gheq *) + ρ Γ (gheq A u v) := gheq (ρ Γ A) (ρ Γ u) (ρ Γ v); + + (* Four variables recursive cases *) + (* if *) + ρ Γ (tif m b P t f) with term_view_if b P t f := { + | term_true P t f := + if mode_eq m ℙ then ⋆ + else if mode_eq (md Γ t) m then ρ Γ t + else tif m ttrue (ρ Γ P) (ρ Γ t) (ρ Γ f); + | term_false P t f := + if mode_eq m ℙ then ⋆ + else if mode_eq (md Γ f) m then ρ Γ f + else tif m tfalse (ρ Γ P) (ρ Γ t) (ρ Γ f); + | term_not_true_not_false b P t f Hb := + if mode_eq m ℙ then ⋆ + else tif m (ρ Γ b) (ρ Γ P) (ρ Γ t) (ρ Γ f)}; + (* nat_elim *) + ρ Γ (tnat_elim m n P z s) with term_view_nat_elim n P z s := { + | term_zero P z s := + if mode_eq m ℙ then ⋆ + else if mode_eq (md Γ z) m then ρ Γ z + else tnat_elim m tzero (ρ Γ P) (ρ Γ z) (ρ Γ s); + | term_succ n P z s := + if mode_eq m ℙ then ⋆ + else if mode_eq (md Γ s) m then + let n' := (ρ Γ n) in let s' := (ρ Γ s) in + app (app s' n') (tnat_elim m n' (ρ Γ P) (ρ Γ z) s') + else tnat_elim m (tsucc (ρ Γ n)) (ρ Γ P) (ρ Γ z) (ρ Γ s); + | term_not_zero_not_succ n P z s Hn := + if mode_eq m ℙ then ⋆ + else tnat_elim m (ρ Γ n) (ρ Γ P) (ρ Γ z) (ρ Γ s)}; + + (* Six variables recursive cases *) + (* vec_elim *) + ρ Γ (tvec_elim m A n0 v P z s) with term_view_vec_elim A n0 v P z s := { + | term_vnil A B P z s := + if mode_eq m ℙ then ⋆ + else if mode_eq (md Γ z) m then ρ Γ z + else tvec_elim m (ρ Γ A) (hide tzero) (tvnil (ρ Γ B)) (ρ Γ P) (ρ Γ z) (ρ Γ s); + | term_vcons A n0 a n v P z s := + if mode_eq m ℙ then ⋆ + else if mode_eq (md Γ s) m then + let A' := ρ Γ A in let n' := ρ Γ n in + let v' := ρ Γ v in let s' := ρ Γ s in + app (app (app (app s' (ρ Γ a)) (glength A' n' v')) v') (tvec_elim m A' n' v' (ρ Γ P) (ρ Γ z) s') + else tvec_elim m (ρ Γ A) (ρ Γ n0) (tvcons (ρ Γ a) (ρ Γ n) (ρ Γ v)) (ρ Γ P) (ρ Γ z) (ρ Γ s); + | term_other_vec_elim A n0 v P z s Hvn := + if mode_eq m ℙ then ⋆ + else tvec_elim m (ρ Γ A) (ρ Γ n0) (ρ Γ v) (ρ Γ P) (ρ Γ z) (ρ Γ s)}; + + (* zero variables recursive cases *) + ρ Γ tbool := tbool; + ρ Γ ttrue := ttrue; + ρ Γ tfalse := tfalse; + ρ Γ tnat := tnat; + ρ Γ tzero := tzero; + ρ Γ bot := bot. + +End Rho. +(* ------------------------------------------------------------------------- *) +(** properties of ρ **) + +Lemma ρ_of_ℙ (Γ : scope) (p : term) : + md Γ p = ℙ → ρ Γ p = ⋆. +Proof. + intro prf_p. + induction p. + all: try reflexivity. + all: simp ρ. + all: try solve [inversion prf_p]. + 8: case term_view_vec_elim in *; simp ρ. + 7: case term_view_nat_elim in *; simp ρ. + 6: case term_view_if in *; simp ρ. + 4: case term_view_reveal in *; simp ρ. + 3: case term_view_app in *; simp ρ. + all: case (mode_eq _ _); [reflexivity |contradiction ]. +Qed. + +Theorem rho_one_step : ∀ Γ t, Γ ⊨ t ↣ ρ Γ t. +Proof. + intros Γ t. + induction t in Γ |- *. + all: simp ρ. + all: try case (mode_eq _ _) as [ e | ne ]; gred. + all: try reflexivity. + - unfold red_lvl. + case (mode_eq _ _) as [ e | ne ]; subst; gred. + - case term_view_app as [m A t | ]; simp ρ. + 1: assert ( Γ⊨lam m A t↣ρ Γ (lam m A t)) as red_lam; eauto. + 1: simp ρ in red_lam. + all: case (mode_eq _ _) as [ e | ne ]; gred. + all: case (mode_eq _ _) as [ e' | ne' ]; gred. + red_lam_inv_auto A' t' e red_A' red_t'. + noconf e. + assumption. + - case term_view_reveal as [t P p | ]; simp ρ. + all: case (mode_eq _ _) as [ e | ne ]; gred. + * assert (Γ⊨hide t↣ρ Γ (hide t)) as H; eauto. + simp ρ in H. + red_hide_inv_auto t' e. + noconf e. + assumption. + * cbn in ne. case (md _ _) in *. + all: try contradiction. + right; left. reflexivity. + - case term_view_if as [ | | ]; simp ρ. + all: repeat case (mode_eq _ _) as [ | ]. + all: try solve [gred]. + all: apply red_if; gred. + - case term_view_nat_elim as [ |n | ]; simp ρ. + all: repeat case (mode_eq _ _) as [ | ]; cbn. + 3: eapply red_nat_elim; gred. + all: gred. + assert (Γ⊨tsucc n↣ρ Γ (tsucc n)) as H; eauto. + simp ρ in H. + red_succ_inv_auto n' e'. + noconf e'. + assumption. + - case term_view_vec_elim as [ | A n a n0 v| ]; simp ρ. + all:repeat case (mode_eq _ _) as [ | ]; cbn. + 3: eapply red_vec_elim; gred. + all: try solve [gred]. + assert (Γ⊨tvcons a n0 v↣ρ Γ (tvcons a n0 v)) as H; eauto. + simp ρ in H. + red_conv_inv_auto a' n' v' e' red_a' red_n' red_v'. + cbn; gred. + all: congruence. +Qed. + +Lemma reduction_triangle (Γ : scope) (t u : term) : + Γ ⊨ t ↣ u → Γ ⊨ u ↣ (ρ Γ t). +Proof. + intros red_t. + induction red_t in t, Γ, u, red_t |- *. + all: simp ρ. + all: try solve [gred; try case (mode_eq _ _); eauto using rho_one_step; gred]. + all: try solve [rewrite ρ_of_ℙ; gred]. + all: try solve [repeat case (mode_eq _ _) as [ | ]; + try contradiction; solve [cbn; gred; cbn; erewrite <- red_md; eauto; congruence]]. + (* red_beta*) + - repeat (case (mode_eq _ _) as [ | ]); gred. + * cbn in *. erewrite <- md_subst'; eauto. + + erewrite <- red_md; eassumption. + + intro n; destruct n; cbn; eauto. + erewrite <- red_md; eassumption. + * eapply red_subst; eauto. + + intro n0; destruct n0; cbn; eauto. + erewrite <- red_md; eassumption. + + intro n0; destruct n0; cbn; gred. + * exfalso; eauto. + (* red_reveal_hide *) + - case (mode_eq _ _) as [ | ]; gred. + erewrite <- red_md; gred. + (* red_vec_elim_cons *) + - repeat case (mode_eq _ _) as [ | ]. + * gred; cbn. erewrite <- red_md; eauto; congruence. + * cbn; gred. + + eapply red_ren; eauto. + intro n0'; destruct n0'; eauto. + + eapply (red_ren _ (𝕋::_)); cbn; eauto. + eapply red_ren; eauto. + all: intro n0'; destruct n0'; eauto. + * contradiction. + (* red_app *) + - case term_view_app as [ mx0 A u v | u v d] in *. + all: simp ρ in *. + all: repeat case (mode_eq _ _) as [ | ]; gred. + * cbn. erewrite <- red_md; eauto. cbn. assumption. + * red_lam_inv_auto A' t' e' red_A red_u. + match goal with H: _ ⊨ lam _ _ _ ↣ _ |- _ => + eapply red_lam_inv in H; eauto; + [ destruct H as [A'' [t'' [ Hu'' [red_A' red_u']]]] | ] end. + + apply sym_eq in Hu''; inversion Hu''; subst. + gred. symmetry; apply red_md. assumption. + + cbn. erewrite <- red_md; eauto. + * cbn; erewrite <- red_md; eauto. + (* red_reveal *) + - case term_view_reveal as [t | ]. + all: simp ρ in *. + all: repeat case (mode_eq _ _) as [ | ]; gred. + 1,3: cbn in *; erewrite <- red_md; eauto. + red_hide_inv_auto t0' e. + gred; eauto using red_md. + * red_hide_inv_auto t0 e. + noconf e. + assumption. + * cbn in *. erewrite <- red_md; eauto. + right; left. case (md _ _) in *; try contradiction. + reflexivity. + (* red_if *) + - simp ρ in *. + case term_view_if as [ | | ] in *. + all: simp ρ in *. + all: repeat case (mode_eq _ _) as [ | ]; gred. + all: red_basic_inv. + all: try (cbn in *; match goal with H: _ = ℙ |- _ => + inversion H end). + all: gred; symmetry; eauto using red_md. + (* red_nat_elim *) + - simp ρ in *. + case term_view_nat_elim as [ |n| n] in *. + all: simp ρ in *. + all: repeat case (mode_eq _ _) as [ | ]; cbn; gred. + * red_basic_inv. + all: try (cbn in *; match goal with H: _ = ℙ |- _ => + inversion H end). + gred; symmetry; eauto using red_md. + * red_succ_inv_auto n'' e'. + cbn; gred. + + red_succ_inv_auto n' e''. + noconf e''. + assumption. + + erewrite <- red_md; eauto. + (* red_vec_elim *) + - case term_view_vec_elim as [ |A n v P z s| ] in *. + all: simp ρ in *. + all: repeat case (mode_eq _ _) as [ | ]; cbn; gred. + * red_hide_inv_auto zero' e'. + red_basic_inv. + all: try (cbn in *; match goal with H: _ = ℙ |- _ => + inversion H end). + red_nil_inv_auto A0 e'. + eapply red_vec_elim_nil; eauto. + erewrite <- red_md; eauto. + * red_conv_inv_auto a1 n1 v1 e' red_a1 red_n1 red_v1. + red_conv_inv_auto a2 n2 v2 e' red_a2 red_n2 red_v2. + noconf e'. + gred; erewrite <- red_md; eauto. +Qed. diff --git a/theories/reduction/ReductionToCongruence.v b/theories/reduction/ReductionToCongruence.v new file mode 100644 index 0000000..1915d61 --- /dev/null +++ b/theories/reduction/ReductionToCongruence.v @@ -0,0 +1,179 @@ +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval + TermMode Scoping BasicMetaTheory. +From GhostTT.reduction Require Export ReductionProperties ReductionConfluence. +From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. + +Import ListNotations. + +Set Default Goal Selector "!". + +Lemma conv_subst_r (Γ : context) (Δ : scope) (m : mode) (σ σ' : nat → term) (t : term) : + sscoping (sc Γ) σ Δ → + Δ ⊢ t∷m → (∀ n, Γ ⊢ σ n ≡ σ' n) → Γ ⊢ t <[ σ] ≡ t <[ σ']. +Proof. + intros Hscope scope_t conv_σ. + induction t in Γ, Δ, Hscope, m, scope_t, σ, σ', conv_σ |- *. + all: inversion scope_t; subst. + all: try solve [gconv]. + - asimpl; gconv; unfold ueq; eauto using sscoping_shift. + admit. (*lemma 1*) + - asimpl; gconv; eauto using sscoping_shift. + admit. (*lemma 1*) + - asimpl; gconv. eapply conv_irr. + * gscope; eauto using scoping_subst. + * gscope; eauto using scoping_subst. + all:admit. (*lemma 2*) + - asimpl; gconv. eapply conv_irr. + * gscope; eauto using scoping_subst. + * gscope; eauto using scoping_subst. + all:admit. (*lemma 2*) + - asimpl; gconv. eapply conv_irr. + * gscope; eauto using scoping_subst. + * gscope; eauto using scoping_subst. + all:admit. (*lemma 2*) + - admit. (*real issue to be discussed*) +Admitted. + + + +Theorem reduction_to_conversion : + ∀ Γ m t t', (sc Γ) ⊢ t∷m → (sc Γ) ⊨ t ↣ t' → Γ ⊢ t ≡ t'. +Proof. (* + intros Γ m t t' scope_t red_t. + remember (sc Γ) as Γ0 eqn:e. + induction red_t in Γ, Γ0, e, m, scope_t, t', red_t |- *. + all: try solve [inversion scope_t; gconv]. + - scope_sub_inv. eapply conv_trans. + * eapply conv_beta; eauto. + * eapply conv_trans. + 1: eapply conv_subst_r; eauto using sscoping_one. + 2: eapply conv_subst; eauto. + + intro n; destruct n; gconv; reflexivity. + + Unshelve. + 2: exact ((mx, u')::Γ). + cbn; eapply sscoping_one; eauto using red_md. + + eauto. + - scope_sub_inv. + eapply conv_trans. + * eapply reveal_hide; eauto. + * gconv; auto. + - inversion scope_t; subst. + eapply conv_trans; solve [gconv | eauto]. + - inversion scope_t; subst. + eapply conv_trans; solve [gconv | eauto]. + - inversion scope_t; subst. + eapply conv_trans; solve [gconv | eauto]. + - inversion scope_t; subst. + eapply conv_trans. + * eapply conv_nat_elim_succ; auto. + * gconv; auto. + - inversion scope_t; subst. + eapply conv_trans. + * admit. + * eauto. + - inversion scope_t; subst. + eapply conv_trans. + * eapply conv_vec_elim_cons; auto. + * gconv; auto. + + eapply conv_ren; eauto using rtyping_S. + + do 2 (eapply conv_ren; eauto using rtyping_S). + - inversion scope_t; subst; gconv; auto. + + destruct mx; unfold ueq; eauto. + + destruct m0; unfold ueq; eauto. + - inversion scope_t; subst; gconv; auto. + + destruct mx; unfold ueq; eauto. + + destruct m0; unfold ueq; eauto. + - inversion scope_t; subst; gconv; auto. + - subst; apply conv_irr; auto using scope_star. + - subst; apply conv_irr; auto using scope_star. + - case m in *; inversion scope_t; subst. + apply conv_irr; auto using scope_star. + gscope; eauto using red_md. + - case m in *; inversion scope_t; subst. + apply conv_irr; auto using scope_star. + gscope; eauto using red_md. + - case m in *; inversion scope_t; subst. + apply conv_irr; auto using scope_star. + gscope; eauto using red_md. + - admit.*) +Admitted. + + +Local Ltac conversion_to_reduction_exists := + match goal with + | H : ∃ _, _ ∧ _ |- + ∃ w, ?Γ ⊨ ?c _ ↣* w ∧ ?Γ ⊨ ?c _ ↣* w => + let w := fresh "w" in + destruct H as [ w [ ]]; + exists (c w) + | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _ |- + ∃ w, ?Γ ⊨ ?c _ _ ↣* w ∧ ?Γ ⊨ ?c _ _↣* w => + let w0 := fresh "w0" in let w1 := fresh "w1" in + destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; + exists (c w0 w1) + | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _, H2 : ∃ _, _ ∧ _ |- + ∃ w, ?Γ ⊨ ?c _ _ _ ↣* w ∧ ?Γ ⊨ ?c _ _ _ ↣* w => + let w0 := fresh "w0" in let w1 := fresh "w1" in + let w2 := fresh "w2" in + destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; + destruct H2 as [ w2 [ ]]; + exists (c w0 w1 w2) + | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _, H2 : ∃ _, _ ∧ _, H3 : ∃ _, _ ∧ _ |- + ∃ w, ?Γ ⊨ ?c _ _ _ _ ↣* w ∧ ?Γ ⊨ ?c _ _ _ _ ↣* w => + let w0 := fresh "w0" in let w1 := fresh "w1" in + let w2 := fresh "w2" in let w3 := fresh "w3" in + destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; + destruct H2 as [ w2 [ ]]; destruct H3 as [ w3 [ ]]; + exists (c w0 w1 w2 w3) + | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _, H2 : ∃ _, _ ∧ _, + H3 : ∃ _, _ ∧ _ , H4 : ∃ _, _ ∧ _, H5 : ∃ _, _ ∧ _ |- + ∃ w, ?Γ ⊨ tvec_elim ?m _ _ _ _ _ _ ↣* w ∧ ?Γ ⊨ tvec_elim ?m _ _ _ _ _ _↣* w => + let w0 := fresh "w0" in let w1 := fresh "w1" in + let w2 := fresh "w2" in let w3 := fresh "w3" in + let w4 := fresh "w4" in let w5 := fresh "w5" in + destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; + destruct H2 as [ w2 [ ]]; destruct H3 as [ w3 [ ]]; + destruct H4 as [ w4 [ ]]; destruct H5 as [ w5 [ ]]; + exists (tvec_elim m w0 w1 w2 w3 w4 w5) + end. + +Theorem conversion_to_reduction : + ∀ Γ t t', Γ ⊢ t ≡ t' → ∃ u, (sc Γ) ⊨ t ↣* u ∧ (sc Γ) ⊨ t' ↣* u. +Proof. + intros Γ t t' conv_t. + induction conv_t in conv_t |- *. + all: try solve [match goal with | |- ∃ _, _ ∧ (_ ⊨ ?c ↣* _) => exists c end; split; apply red_trans_direct; gred; eauto using scoping_md]. + all: try solve [conversion_to_reduction_exists; split; greds]. + - match goal with | |- ∃ _, _ ∧ (_ ⊨ ?c ↣* _) => exists c end; split; greds. + all: try (apply red_trans_direct; gred). + erewrite scoping_md; eauto. + - exists (Sort ℙ 0); split; greds. + - match goal with + | H0: ∃ _, _ ∧ _, H1 : ∃ _, _∧ _, + Hi : ueq ?m ?i ?i', Hj : ueq ?m' ?j ?j' |- _ => + destruct H0 as [ w0 [ ]]; + destruct H1 as [ w1 [ ]]; + exists (Pi (red_lvl m i) (red_lvl m' j) m' m w0 w1); + destruct Hi, Hj + end. + all: split; subst; unfold red_lvl; cbn; greds. + - match goal with + | H: ∃ _, _ ∧ _ |- _ => destruct H as [w []] end. + eauto. + - match goal with + | H0: ∃ _, _ ∧ _, H1 : ∃ _, _∧ _ |- _ => + destruct H0 as [ w0 [ ]]; + destruct H1 as [ w1 [ ]] + end. + assert (∃ w: term, (sc Γ⊨w0↣*w) ∧ sc Γ⊨w1↣*w) as Hw. + { eauto using reduction_confluence. } + destruct Hw as [wf []]. + exists wf. + split; eauto using red_trans_trans. + - exists ⋆; split; apply red_trans_direct; gred. + all: eauto using scoping_md. +Qed. + +Print Assumptions conversion_to_reduction. From 86842cf48a464c8f4dc54bd19b9275f6977d5c61 Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Wed, 12 Jun 2024 11:08:04 +0200 Subject: [PATCH 02/15] reflectiontoconversion --- _CoqProject | 2 + theories/Typing.v | 5 +- theories/reduction/ReductionProperties.v | 33 ++++ theories/reduction/ReductionToCongruence.v | 190 +++++++++------------ 4 files changed, 120 insertions(+), 110 deletions(-) diff --git a/_CoqProject b/_CoqProject index 113e8d4..06fd308 100644 --- a/_CoqProject +++ b/_CoqProject @@ -24,6 +24,8 @@ theories/reduction/Reduction.v theories/reduction/ReductionProperties.v theories/reduction/ReductionRho.v theories/reduction/ReductionAndTransitivity.v +theories/reduction/ReductionConfluence.v +theories/reduction/ReductionToCongruence.v theories/CScoping.v diff --git a/theories/Typing.v b/theories/Typing.v index 8e25a8a..6f8d79c 100644 --- a/theories/Typing.v +++ b/theories/Typing.v @@ -86,13 +86,14 @@ Inductive conversion (Γ : context) : term → term → Prop := Γ ⊢ tnat_elim m (tsucc n) P z s ≡ app (app s n) (tnat_elim m n P z s) | conv_vec_elim_nil : - ∀ m A P z s, + ∀ m A B P z s, m ≠ mKind → cscoping Γ A mKind → + cscoping Γ B mKind → cscoping Γ P mKind → cscoping Γ z m → cscoping Γ s m → - Γ ⊢ tvec_elim m A (hide tzero) (tvnil A) P z s ≡ z + Γ ⊢ tvec_elim m A (hide tzero) (tvnil B) P z s ≡ z | conv_vec_elim_cons : ∀ m A a n n' v P z s, diff --git a/theories/reduction/ReductionProperties.v b/theories/reduction/ReductionProperties.v index 06ebefb..e058659 100644 --- a/theories/reduction/ReductionProperties.v +++ b/theories/reduction/ReductionProperties.v @@ -63,6 +63,39 @@ Proof. cbn; rewrite H; reflexivity end. Qed. +Lemma red_scope : + ∀ Γ m t t', Γ ⊨ t ↣ t' → Γ ⊢ t∷m → Γ ⊢ t'∷m. +Proof. + intros Γ m t t' red_t scope_t. + induction red_t in Γ, m, t, t', red_t, scope_t |- *. + all: try solve [inversion scope_t; gscope]. + - inversion scope_t. + match goal with H : _ ⊢ lam _ _ _∷_ |- _ => + inversion H; subst end. + eapply scoping_subst; eauto. + eapply sscoping_one. + erewrite scoping_md; eauto. + - inversion scope_t. + match goal with H : _ ⊢ hide _∷_ |- _ => + inversion H; subst end. + gscope. + - inversion scope_t. + match goal with H : _ ⊢ tsucc _∷_ |- _ => + inversion H; subst end. + gscope. + - inversion scope_t. + match goal with H : _ ⊢ tvcons _ _ _∷_ |- _ => + inversion H; subst end. + gscope; eauto. + * intro H; inversion H. + * eapply scoping_ren; eauto using rscoping_S. + * eapply scoping_ren; eauto using rscoping_S. + eapply scoping_ren; eauto using rscoping_S. + * right; left. reflexivity. + - erewrite scoping_md in *; eauto. + subst. apply scope_star. +Qed. + Lemma glenght_red_subst (A n v : term) (σ : nat → term) : (glength A n v)<[σ] = glength (A<[σ]) (n<[σ]) (v<[σ]). Proof. diff --git a/theories/reduction/ReductionToCongruence.v b/theories/reduction/ReductionToCongruence.v index 1915d61..3647324 100644 --- a/theories/reduction/ReductionToCongruence.v +++ b/theories/reduction/ReductionToCongruence.v @@ -7,136 +7,112 @@ From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. Import ListNotations. +Notation "Γ ⊨ u ε↣ v" := (Γ ⊨ ε|u| ↣ ε|v|) + (at level 80, u, v at next level, format "Γ ⊨ u ε↣ v"). + Set Default Goal Selector "!". Lemma conv_subst_r (Γ : context) (Δ : scope) (m : mode) (σ σ' : nat → term) (t : term) : - sscoping (sc Γ) σ Δ → - Δ ⊢ t∷m → (∀ n, Γ ⊢ σ n ≡ σ' n) → Γ ⊢ t <[ σ] ≡ t <[ σ']. + sscoping (sc Γ) σ Δ → sscoping (sc Γ) σ' Δ → + Δ ⊢ ε|t|∷m → (∀ n, Γ ⊢ σ n ≡ σ' n) → Γ ⊢ ε|t| <[σ] ≡ ε|t| <[σ']. Proof. - intros Hscope scope_t conv_σ. - induction t in Γ, Δ, Hscope, m, scope_t, σ, σ', conv_σ |- *. + intros Hscope Hscope' scope_t conv_σ. + induction t in Γ, Δ, Hscope, Hscope', m, scope_t, σ, σ', conv_σ |- *. all: inversion scope_t; subst. all: try solve [gconv]. - - asimpl; gconv; unfold ueq; eauto using sscoping_shift. - admit. (*lemma 1*) - - asimpl; gconv; eauto using sscoping_shift. - admit. (*lemma 1*) - - asimpl; gconv. eapply conv_irr. - * gscope; eauto using scoping_subst. - * gscope; eauto using scoping_subst. - all:admit. (*lemma 2*) - - asimpl; gconv. eapply conv_irr. - * gscope; eauto using scoping_subst. - * gscope; eauto using scoping_subst. - all:admit. (*lemma 2*) - - asimpl; gconv. eapply conv_irr. - * gscope; eauto using scoping_subst. - * gscope; eauto using scoping_subst. - all:admit. (*lemma 2*) - - admit. (*real issue to be discussed*) -Admitted. - + 3-5 : eauto using conv_irr, scoping_subst. + all: cbn; gconv; unfold ueq; eauto using sscoping_shift. + all: intro n; destruct n; gconv. + all: cbn; ssimpl; eauto using conv_ren, rtyping_S. +Qed. Theorem reduction_to_conversion : - ∀ Γ m t t', (sc Γ) ⊢ t∷m → (sc Γ) ⊨ t ↣ t' → Γ ⊢ t ≡ t'. -Proof. (* + ∀ Γ m t t', (sc Γ) ⊢ t∷m → (sc Γ) ⊨ t ↣ t' → Γ ⊢ t ε≡ t'. +Proof. intros Γ m t t' scope_t red_t. remember (sc Γ) as Γ0 eqn:e. - induction red_t in Γ, Γ0, e, m, scope_t, t', red_t |- *. - all: try solve [inversion scope_t; gconv]. - - scope_sub_inv. eapply conv_trans. - * eapply conv_beta; eauto. - * eapply conv_trans. - 1: eapply conv_subst_r; eauto using sscoping_one. - 2: eapply conv_subst; eauto. - + intro n; destruct n; gconv; reflexivity. - + Unshelve. - 2: exact ((mx, u')::Γ). - cbn; eapply sscoping_one; eauto using red_md. - + eauto. - - scope_sub_inv. + induction red_t in Γ, Γ0, e, t, m, scope_t, t', red_t |- *. + 28:{ apply conv_irr; [| apply scope_star]. + apply scoping_castrm. + erewrite scoping_md in *; eauto. congruence. } + 9, 27: solve [gconv]. + all: inversion scope_t; subst. + 26-28: solve [apply conv_irr; apply scoping_castrm; subst; gscope; eauto using red_scope]. + 11-26: solve [gconv; cbn; f_equal; assumption]. + 3-5: solve [subst; eapply conv_trans; [| eauto]; gconv; eauto using scoping_castrm]. + 6-7: solve [unfold red_lvl; repeat case (mode_eq _ ℙ) as [ | ]; subst; gconv; unfold ueq; eauto]. + - cbn in *; inversion H2; subst. + erewrite scoping_md in *; eauto. eapply conv_trans. - * eapply reveal_hide; eauto. - * gconv; auto. - - inversion scope_t; subst. - eapply conv_trans; solve [gconv | eauto]. - - inversion scope_t; subst. - eapply conv_trans; solve [gconv | eauto]. - - inversion scope_t; subst. - eapply conv_trans; solve [gconv | eauto]. - - inversion scope_t; subst. + * eauto using conv_beta, scoping_castrm. + * eapply conv_trans. + + eapply (conv_subst _ (_,,(_, ε|var 0|))). + ++ eauto using sscoping_one, scoping_castrm. + ++ eauto. + + erewrite castrm_subst. asimpl. + eapply conv_subst_r. + ++ eauto using sscoping_one, scoping_castrm. + ++ eauto using sscoping_one, scoping_castrm, red_scope. + ++ apply scoping_castrm. + eapply red_scope; eauto. + ++ intro n; destruct n; cbn; gconv. reflexivity. + - cbn in *; inversion H4; subst. eapply conv_trans. - * eapply conv_nat_elim_succ; auto. - * gconv; auto. - - inversion scope_t; subst. + * eapply reveal_hide; eauto using scoping_castrm. + * eauto using cong_app. + - cbn in *; inversion H7; subst. eapply conv_trans. - * admit. - * eauto. - - inversion scope_t; subst. + * eapply conv_nat_elim_succ; eauto using scoping_castrm. + * gconv; eauto. + - cbn in *; inversion H11; subst. + eapply conv_trans; [ | eauto]. + gconv; eauto using scoping_castrm. + - cbn in *; inversion H11; subst. eapply conv_trans. - * eapply conv_vec_elim_cons; auto. - * gconv; auto. - + eapply conv_ren; eauto using rtyping_S. - + do 2 (eapply conv_ren; eauto using rtyping_S). - - inversion scope_t; subst; gconv; auto. - + destruct mx; unfold ueq; eauto. - + destruct m0; unfold ueq; eauto. - - inversion scope_t; subst; gconv; auto. - + destruct mx; unfold ueq; eauto. - + destruct m0; unfold ueq; eauto. - - inversion scope_t; subst; gconv; auto. - - subst; apply conv_irr; auto using scope_star. - - subst; apply conv_irr; auto using scope_star. - - case m in *; inversion scope_t; subst. - apply conv_irr; auto using scope_star. - gscope; eauto using red_md. - - case m in *; inversion scope_t; subst. - apply conv_irr; auto using scope_star. - gscope; eauto using red_md. - - case m in *; inversion scope_t; subst. - apply conv_irr; auto using scope_star. - gscope; eauto using red_md. - - admit.*) -Admitted. + * eapply conv_vec_elim_cons; eauto using scoping_castrm. + * gconv; eauto. + + erewrite castrm_ren. eauto using conv_ren, rtyping_S. + + do 2 erewrite castrm_ren. eauto using conv_ren, rtyping_S. +Qed. Local Ltac conversion_to_reduction_exists := match goal with | H : ∃ _, _ ∧ _ |- ∃ w, ?Γ ⊨ ?c _ ↣* w ∧ ?Γ ⊨ ?c _ ↣* w => - let w := fresh "w" in - destruct H as [ w [ ]]; - exists (c w) + let w := fresh "w" in + destruct H as [ w [ ]]; + exists (c w) | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _ |- - ∃ w, ?Γ ⊨ ?c _ _ ↣* w ∧ ?Γ ⊨ ?c _ _↣* w => - let w0 := fresh "w0" in let w1 := fresh "w1" in - destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; - exists (c w0 w1) + ∃ w, ?Γ ⊨ ?c _ _ ↣* w ∧ ?Γ ⊨ ?c _ _↣* w => + let w0 := fresh "w0" in let w1 := fresh "w1" in + destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; + exists (c w0 w1) | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _, H2 : ∃ _, _ ∧ _ |- - ∃ w, ?Γ ⊨ ?c _ _ _ ↣* w ∧ ?Γ ⊨ ?c _ _ _ ↣* w => - let w0 := fresh "w0" in let w1 := fresh "w1" in - let w2 := fresh "w2" in - destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; - destruct H2 as [ w2 [ ]]; - exists (c w0 w1 w2) + ∃ w, ?Γ ⊨ ?c _ _ _ ↣* w ∧ ?Γ ⊨ ?c _ _ _ ↣* w => + let w0 := fresh "w0" in let w1 := fresh "w1" in + let w2 := fresh "w2" in + destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; + destruct H2 as [ w2 [ ]]; + exists (c w0 w1 w2) | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _, H2 : ∃ _, _ ∧ _, H3 : ∃ _, _ ∧ _ |- - ∃ w, ?Γ ⊨ ?c _ _ _ _ ↣* w ∧ ?Γ ⊨ ?c _ _ _ _ ↣* w => - let w0 := fresh "w0" in let w1 := fresh "w1" in - let w2 := fresh "w2" in let w3 := fresh "w3" in - destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; - destruct H2 as [ w2 [ ]]; destruct H3 as [ w3 [ ]]; - exists (c w0 w1 w2 w3) + ∃ w, ?Γ ⊨ ?c _ _ _ _ ↣* w ∧ ?Γ ⊨ ?c _ _ _ _ ↣* w => + let w0 := fresh "w0" in let w1 := fresh "w1" in + let w2 := fresh "w2" in let w3 := fresh "w3" in + destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; + destruct H2 as [ w2 [ ]]; destruct H3 as [ w3 [ ]]; + exists (c w0 w1 w2 w3) | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _, H2 : ∃ _, _ ∧ _, H3 : ∃ _, _ ∧ _ , H4 : ∃ _, _ ∧ _, H5 : ∃ _, _ ∧ _ |- - ∃ w, ?Γ ⊨ tvec_elim ?m _ _ _ _ _ _ ↣* w ∧ ?Γ ⊨ tvec_elim ?m _ _ _ _ _ _↣* w => - let w0 := fresh "w0" in let w1 := fresh "w1" in - let w2 := fresh "w2" in let w3 := fresh "w3" in - let w4 := fresh "w4" in let w5 := fresh "w5" in - destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; - destruct H2 as [ w2 [ ]]; destruct H3 as [ w3 [ ]]; - destruct H4 as [ w4 [ ]]; destruct H5 as [ w5 [ ]]; - exists (tvec_elim m w0 w1 w2 w3 w4 w5) + ∃ w, ?Γ ⊨ tvec_elim ?m _ _ _ _ _ _ ↣* w ∧ ?Γ ⊨ tvec_elim ?m _ _ _ _ _ _↣* w => + let w0 := fresh "w0" in let w1 := fresh "w1" in + let w2 := fresh "w2" in let w3 := fresh "w3" in + let w4 := fresh "w4" in let w5 := fresh "w5" in + destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; + destruct H2 as [ w2 [ ]]; destruct H3 as [ w3 [ ]]; + destruct H4 as [ w4 [ ]]; destruct H5 as [ w5 [ ]]; + exists (tvec_elim m w0 w1 w2 w3 w4 w5) end. Theorem conversion_to_reduction : @@ -164,8 +140,8 @@ Proof. eauto. - match goal with | H0: ∃ _, _ ∧ _, H1 : ∃ _, _∧ _ |- _ => - destruct H0 as [ w0 [ ]]; - destruct H1 as [ w1 [ ]] + destruct H0 as [ w0 [ ]]; + destruct H1 as [ w1 [ ]] end. assert (∃ w: term, (sc Γ⊨w0↣*w) ∧ sc Γ⊨w1↣*w) as Hw. { eauto using reduction_confluence. } @@ -175,5 +151,3 @@ Proof. - exists ⋆; split; apply red_trans_direct; gred. all: eauto using scoping_md. Qed. - -Print Assumptions conversion_to_reduction. From 17b5d139d76af38ab734b26372b5b6c019456b1f Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Wed, 12 Jun 2024 17:34:36 +0200 Subject: [PATCH 03/15] Injectivity Pi --- _CoqProject | 2 + theories/reduction/Injectivity.v | 118 ++++++++++++++++++ theories/reduction/ReductionAndTransitivity.v | 76 ++++++++++- theories/reduction/ReductionProperties.v | 41 +++++- theories/reduction/ReductionToCongruence.v | 10 ++ 5 files changed, 239 insertions(+), 8 deletions(-) create mode 100644 theories/reduction/Injectivity.v diff --git a/_CoqProject b/_CoqProject index 06fd308..08a2a5c 100644 --- a/_CoqProject +++ b/_CoqProject @@ -26,6 +26,7 @@ theories/reduction/ReductionRho.v theories/reduction/ReductionAndTransitivity.v theories/reduction/ReductionConfluence.v theories/reduction/ReductionToCongruence.v +theories/reduction/Injectivity.v theories/CScoping.v @@ -40,6 +41,7 @@ theories/Model.v theories/Admissible.v theories/FreeTheorem.v +theories/Examples.v theories/RTyping.v theories/Potential.v diff --git a/theories/reduction/Injectivity.v b/theories/reduction/Injectivity.v new file mode 100644 index 0000000..80be86d --- /dev/null +++ b/theories/reduction/Injectivity.v @@ -0,0 +1,118 @@ +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval + TermMode Scoping BasicMetaTheory. +From GhostTT.reduction Require Import ReductionToCongruence. +From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. + +Import ListNotations. + +Set Default Goal Selector "!". + +Lemma glength_castrm {A n v: term} : + ε|glength A n v| = glength ε|A| ε|n| ε|v| . +Proof. + cbn. unfold glength. repeat f_equal. + + eauto using castrm_ren. + + eapply eq_trans; eauto using castrm_ren. + f_equal. eauto using castrm_ren. +Qed. + +Lemma red_castrm {Γ : scope} {t t' : term} : + Γ ⊨t ↣ t' → Γ ⊨ t ε↣ t'. +Proof. + intro red_t. + induction red_t. + 1: erewrite castrm_subst; asimpl. + all: gred. + all: try (erewrite <- md_castrm; assumption). + remember (glength _ _ _) as t eqn:e. + apply (f_equal castrm) in e. + rewrite glength_castrm in e. + cbn; subst. + rewrite e. + gred. + erewrite <- md_castrm. + reflexivity. +Qed. + +Theorem injectivity_lam {Γ : context} {m md_t md_t': mode} {A A' t t': term} : + md_t ≠ ℙ → + (sc Γ) ⊢ A∷𝕂 → m::(sc Γ)⊢t∷md_t → + (sc Γ) ⊢ A'∷𝕂 → m::(sc Γ)⊢t'∷md_t → + Γ ⊢ lam m A t ≡ lam m A' t' → + Γ ⊢ A ε≡ A' ∧ (m,A)::Γ ⊢ t ε≡ t'. +Proof. + intros not_Prop scope_A scope_t scope_A' scope_t' H. + apply conversion_to_reduction in H. + destruct H as [w [red1 red2]]. + inversion red1. + - inversion red2 as [e|]; subst. + * inversion e. split; gconv. + * apply reds_lam_inv in red2 as [* [* [e []]]]. + 2: cbn; erewrite scoping_md; eauto. + inversion e. + split; apply conv_sym; subst. + all: eapply reductions_to_conversion; cbn; eauto. + - inversion red2; subst. + * apply reds_lam_inv in red1 as [* [* [e []]]]. + 2: cbn; erewrite scoping_md; eauto. + inversion e. + split; eapply reductions_to_conversion; cbn; eauto. + * apply reds_lam_inv in red1 as [* [* [e [ ]]]]. + 2: cbn; erewrite scoping_md; eauto. + apply reds_lam_inv in red2 as [* [* [e'[ ]]]]. + 2: cbn; erewrite scoping_md; eauto. + subst; inversion e'; subst. + split; eapply conv_trans. + 2,4: apply conv_sym. + all: eapply reductions_to_conversion; eauto. +Qed. + +Theorem injectivity_Pi {Γ : context} {i i' j j': level} {m m' mx mx': mode} {A A' B B': term} : + (sc Γ) ⊢ A∷𝕂 → mx::(sc Γ)⊢B∷𝕂 → + (sc Γ) ⊢ A'∷𝕂 → mx'::(sc Γ)⊢B'∷𝕂 → + Γ ⊢ Pi i j m mx A B ≡ Pi i' j' m' mx' A' B' → + Γ ⊢ A ε≡ A' ∧ (mx,A)::Γ ⊢ B ε≡ B'. +Proof. + intros scope_A scope_B scope_A' scope_B' H. + apply conversion_to_reduction in H. + destruct H as [w [red1 red2]]. + inversion red1. + - inversion red2 as [e|]; subst. + * inversion e. split; gconv. + * apply reds_Pi_inv in red2 as [* [* [* [* [e [ ]]]]]]. + inversion e. + split; apply conv_sym; subst. + all: eapply reductions_to_conversion; cbn; eauto. + - inversion red2; subst. + * apply reds_Pi_inv in red1 as [* [* [* [* [e [ ]]]]]]. + inversion e. + eauto using reductions_to_conversion. + * apply reds_Pi_inv in red1 as [* [* [* [* [e [ ]]]]]]. + apply reds_Pi_inv in red2 as [* [* [* [* [e'[ ]]]]]]. + subst; inversion e'; subst. + split; eapply conv_trans. + 2,4: apply conv_sym. + all: eapply reductions_to_conversion; eauto. +Qed. + +Theorem injectivity_Sort {Γ : context} {i i': level} {m m' : mode} : + Γ ⊢ Sort m i ≡ Sort m' i' → m' = m. +Proof. + intros H. + apply conversion_to_reduction in H. + destruct H as [w [red1 red2]]. + inversion red1. + - inversion red2 as [e|]; subst. + * inversion e. reflexivity. + * apply reds_Sort_inv in red2 as [* e]. + inversion e. reflexivity. + - inversion red2; subst. + * apply reds_Sort_inv in red1 as [* e]. + inversion e. reflexivity. + * apply reds_Sort_inv in red1 as [* e]. + apply reds_Sort_inv in red2 as [* e']. + subst; inversion e'. + reflexivity. +Qed. diff --git a/theories/reduction/ReductionAndTransitivity.v b/theories/reduction/ReductionAndTransitivity.v index 449b3eb..b9936c2 100644 --- a/theories/reduction/ReductionAndTransitivity.v +++ b/theories/reduction/ReductionAndTransitivity.v @@ -8,7 +8,7 @@ From GhostTT.reduction Require Export Reduction. Import ListNotations. Set Default Goal Selector "!". - +(* Definition *) Inductive reduction_trans (Γ : scope) (u v: term) : Prop := | Refl: u = v → reduction_trans Γ u v | Trans w : Γ ⊨ u ↣ w → reduction_trans Γ w v → reduction_trans Γ u v. @@ -16,6 +16,7 @@ Inductive reduction_trans (Γ : scope) (u v: term) : Prop := Notation "Γ ⊨ u ↣* v" := (reduction_trans Γ u v) (at level 80, u, v at next level, format "Γ ⊨ u ↣* v"). +(* Usefull properties *) Lemma red_trans_direct {Γ : scope } {u v: term} : Γ ⊨ u ↣ v → Γ ⊨ u ↣* v. Proof. refine ( λ H, Trans Γ u v v H (Refl Γ v v _)). @@ -31,6 +32,16 @@ Proof. - eapply Trans; eauto. Qed. +Corollary reds_scope (Γ : scope) (m: mode) (t t': term) : + Γ ⊨ t ↣* t' → Γ⊢t∷m → Γ⊢t'∷m. +Proof. + intros reds_t scope_t. + induction reds_t. + - subst; assumption. + - eauto using red_scope. +Qed. + + (* reds deinitions *) Local Ltac end_things H:= @@ -39,7 +50,7 @@ Local Ltac end_things H:= eapply Trans; [ gred | eassumption]]. Lemma reds_beta (Γ : scope) (mx : mode) (A t t' u u' : term) : - mx :: Γ⊨t↣*t'→ md Γ u = mx → Γ⊨u↣*u' → Γ⊨app (lam mx A t) u↣*t' <[u'··]. + mx :: Γ⊨t↣*t'→ md Γ u = mx → Γ⊨u↣*u' → Γ⊨app (lam mx A t) u↣*t' <[u'··]. Proof. intros red_t scope_u red_u. induction red_u. @@ -108,13 +119,13 @@ Qed. Lemma reds_Erased (Γ : scope) (A A' : term) : Γ⊨A↣*A' → Γ⊨Erased A↣*Erased A'. Proof. -intro red_A; end_things red_A. + intro red_A; end_things red_A. Qed. Lemma reds_hide (Γ : scope) (A A' : term) : Γ⊨A↣*A' → Γ⊨hide A↣*hide A'. Proof. -intro red_A; end_things red_A. + intro red_A; end_things red_A. Qed. Lemma reds_reveal (Γ : scope) (t t' P P' p p' : term) : @@ -250,3 +261,60 @@ Ltac greds := unshelve typeclasses eauto with gtt_scope gtt_reds gtt_red shelvedb ; shelve_unifiable. (** end rewriting automation **) + +(* reds inversions *) + +Lemma reds_lam_inv {Γ : scope} {m : mode} {A t u: term} : + Γ⊨lam m A t↣* u → md Γ (lam m A t) ≠ ℙ → + (∃ A' t', u = lam m A' t' ∧ Γ ⊨ A ↣* A' ∧ m::Γ ⊨ t ↣* t'). +Proof. + intros red_lam not_Prop. + remember (lam m A t) as t0 eqn:e0. + induction red_lam as [|w u v H red_v IH] in A, t, t0, e0, red_lam, not_Prop. + - exists A, t. + subst; repeat split; eauto using red_trans_direct, red_refl. + - subst. + red_lam_inv_auto A'' t'' e red_A red_t. + assert (md Γ (lam m A'' t'') ≠ ℙ) as not_Prop'. + { cbn in *; intro H. apply not_Prop. + erewrite red_md; eauto. } + specialize (IH _ _ eq_refl not_Prop'). + destruct IH as [A' [t' [e [red_A'' red_t'']]]]. + exists A', t'. repeat split. + * assumption. + * eapply red_trans_trans; eauto using red_trans_direct. + * eapply red_trans_trans; eauto using red_trans_direct. +Qed. + +Lemma reds_Pi_inv {Γ : scope} {i j: level} {m mx : mode} {A B t: term} : + Γ⊨Pi i j m mx A B↣* t → + (∃ A' B' i' j', t = Pi i' j' m mx A' B' ∧ Γ ⊨ A ↣* A' ∧ mx::Γ ⊨ B ↣* B'). +Proof. + intro red_Pi. + remember (Pi i j m mx A B) as t0 eqn:e0. + induction red_Pi as [|w u v H red_v IH] in A, B, i, j, t0, e0, red_Pi. + - exists A, B, i, j. + subst; repeat split; eauto using red_trans_direct, red_refl. + - subst. + apply red_Pi_inv in H. + destruct H as [A''[B''[i''[j''[e []]]]]]. + specialize (IH _ _ _ _ e). + destruct IH as [A' [B' [i' [j' [e' [red_A'' red_B'']]]]]]. + exists A', B', i', j'. repeat split. + * assumption. + * eapply red_trans_trans; eauto using red_trans_direct. + * eapply red_trans_trans; eauto using red_trans_direct. +Qed. + +Lemma reds_Sort_inv {Γ : scope} {i: level} {m : mode} {t: term} : + Γ⊨Sort m i ↣* t → ∃ i', t= Sort m i'. +Proof. + intro red_sort. + remember (Sort m i) as t0 eqn:e0. + induction red_sort as [|w u v H red_v IH] in i, t0, e0, red_sort. + - subst; eauto. + - subst. + apply red_Sort_inv in H. + destruct H as [i' e]. + eauto. +Qed. diff --git a/theories/reduction/ReductionProperties.v b/theories/reduction/ReductionProperties.v index e058659..0acf244 100644 --- a/theories/reduction/ReductionProperties.v +++ b/theories/reduction/ReductionProperties.v @@ -119,14 +119,14 @@ Proof. all: asimpl; reflexivity. Qed. -Lemma red_lam_inv {Γ : scope} {mx md_t : mode} {A t u: term} : - (Γ⊨lam mx A t↣ u ) → md Γ (lam mx A t) = md_t → (md_t ≠ ℙ) → +Lemma red_lam_inv {Γ : scope} {mx : mode} {A t u: term} : + (Γ⊨lam mx A t↣ u ) → md Γ (lam mx A t) ≠ ℙ → ( ∃ A' t', u = lam mx A' t' ∧ Γ⊨A↣A' ∧ mx :: Γ⊨t↣t'). Proof. - intros prf_red scope_t not_Prop. + intros red_lam not_Prop. remember (lam mx A t) as lam_t eqn:e0. remember u as u0 eqn:e1. - induction prf_red. + induction red_lam. all: try solve [inversion e0]. - inversion e0. inversion e1; subst. @@ -135,6 +135,31 @@ Proof. - subst. contradiction. Qed. + +Lemma red_Pi_inv {Γ : scope} {i j: level} {m mx : mode} {A B t: term} : + Γ⊨Pi i j m mx A B↣ t → + (∃ A' B' i' j', t = Pi i' j' m mx A' B' ∧ Γ ⊨ A ↣ A' ∧ mx::Γ ⊨ B ↣ B'). +Proof. + intro red_Pi. + inversion red_Pi; subst. + - do 4 eexists; eauto. + - do 4 eexists; eauto using red_refl. + - do 4 eexists; eauto using red_refl. + - cbn in *. + match goal with | HC : 𝕂 = ℙ |- _ => inversion HC end. +Qed. + +Lemma red_Sort_inv {Γ: scope} {i: level} {m: mode} {t: term} : + Γ ⊨ Sort m i ↣ t → ∃ i', t = Sort m i'. +Proof. + intro red_sort. + inversion red_sort. + - eauto. + - eauto. + - cbn in *. + match goal with | HC : 𝕂 = ℙ |- _ => inversion HC end. +Qed. + Lemma red_hide_inv (Γ : scope) (t0 t' : term) (red_hide : Γ⊨hide t0 ↣t' ) : ∃ t0', t' = hide t0' ∧ Γ ⊨ t0 ↣ t0'. Proof. inversion red_hide; subst. @@ -255,6 +280,14 @@ Ltac red_lam_inv_auto A' t' e red_A' red_t':= try subst u end. +Ltac red_Pi_inv_auto A' B' i' j' e red_A' red_B':= + match goal with + | red_Pi : ?Γ⊨Pi ?i ?j ?m ?mx ?A ?B ↣?u |- _ => + eapply red_Pi_inv in red_Pi; eauto; + destruct red_lam as [A' [B' [i' [j' [e [red_A' red_B']]]]]]; + try subst u + end. + Ltac red_hide_inv_auto t0' e:= match goal with | red_hide : ?Γ⊨hide ?t0 ↣?t' |- _ => diff --git a/theories/reduction/ReductionToCongruence.v b/theories/reduction/ReductionToCongruence.v index 3647324..b0e8c87 100644 --- a/theories/reduction/ReductionToCongruence.v +++ b/theories/reduction/ReductionToCongruence.v @@ -76,6 +76,16 @@ Proof. + do 2 erewrite castrm_ren. eauto using conv_ren, rtyping_S. Qed. +Corollary reductions_to_conversion : + ∀ Γ m t t', (sc Γ) ⊢ t∷m → (sc Γ) ⊨ t ↣* t' → Γ ⊢ t ε≡ t'. +Proof. + intros Γ m t t' scope_t red_t. + induction red_t. + - subst; gconv. + - eapply conv_trans. + * eauto using reduction_to_conversion. + * eauto using red_scope. +Qed. Local Ltac conversion_to_reduction_exists := match goal with From 432f4a769d606f46ed0b8f64a83f8dcd9fd97201 Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Thu, 13 Jun 2024 10:44:08 +0200 Subject: [PATCH 04/15] fix Param --- theories/Param.v | 4 +-- theories/reduction/SubjectReduction.v | 46 +++++++++++++++++++++++++++ 2 files changed, 48 insertions(+), 2 deletions(-) create mode 100644 theories/reduction/SubjectReduction.v diff --git a/theories/Param.v b/theories/Param.v index 0160d7e..75c9fcf 100644 --- a/theories/Param.v +++ b/theories/Param.v @@ -2196,8 +2196,8 @@ Proof. eapply scoping_epm_lift. 1: escope. 1: reflexivity. apply csc_param_ctx. * escope. all: eauto using csc_param_ctx. - - cbn. eapply param_scoping in H0, H1, H2, H3. - rewrite <- csc_param_ctx in H0, H1, H2, H3. + - cbn. eapply param_scoping in H0, H1, H2, H3, H4. + rewrite <- csc_param_ctx in H0, H1, H2, H3, H4. destruct m. + contradiction. + cbn in *. eapply cconv_irr. diff --git a/theories/reduction/SubjectReduction.v b/theories/reduction/SubjectReduction.v new file mode 100644 index 0000000..770e039 --- /dev/null +++ b/theories/reduction/SubjectReduction.v @@ -0,0 +1,46 @@ +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval + TermMode Scoping BasicMetaTheory. +From GhostTT.reduction Require Import ReductionProperties Injectivity. +From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. + +Import ListNotations. + +Set Default Goal Selector "!". + +Lemma scoping_type {Γ: context} {A t: term}: + wf Γ → Γ ⊢ t:A → sc Γ⊢ A∷𝕂. +Proof. + intros wfΓ type_t. + specialize (validity Γ _ _ wfΓ type_t) as [scope_t [i type_A]]. + specialize (validity Γ _ _ wfΓ type_A) as [scope_A [j type_scope]]. + ttinv type_scope type_scope. + apply injectivity_Sort in type_scope. + rewrite type_scope in *. + assumption. +Qed. + + + +Theorem subjet_reduction (Γ: context) (A t t': term): + wf Γ → (sc Γ) ⊨ t ↣ t' → Γ ⊢ t:A → mdc Γ t ≠ ℙ → Γ⊢ t':A. +Proof. + intros wfΓ red_t type_t not_Prop. + remember (sc Γ) as Γ0 eqn:e0. + induction red_t in Γ, Γ0, e0, wfΓ, A, red_t, type_t, not_Prop. + all: gtype. + - admit. + - specialize (validity Γ _ _ wfΓ type_t) as [scope_t [i type_A]]. + ttinv type_t type_t'. + destruct_exists type_t'. + repeat destruct type_t' as [H0 [ H1 [ H2 [H3 [H4 [H5 [H6[]]]]]]]]. + + eapply type_app. + * admit. + * exact (scoping_type wfΓ type_t). + * eapply red_scope. + 2: exact (proj1 (validity Γ _ _ wfΓ type_t)). + subst; gred. + * + From ded96a2ac81114221071274e1aa211011d904ab8 Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Tue, 18 Jun 2024 10:36:07 +0200 Subject: [PATCH 05/15] =?UTF-8?q?=09modifi=C3=A9=C2=A0:=20=20=20=20=20=20?= =?UTF-8?q?=20=20=20theories/Scoping.v=20=09modifi=C3=A9=C2=A0:=20=20=20?= =?UTF-8?q?=20=20=20=20=20=20theories/reduction/Injectivity.v=20=09modifi?= =?UTF-8?q?=C3=A9=C2=A0:=20=20=20=20=20=20=20=20=20theories/reduction/Redu?= =?UTF-8?q?ctionProperties.v=20=09modifi=C3=A9=C2=A0:=20=20=20=20=20=20=20?= =?UTF-8?q?=20=20theories/reduction/ReductionToCongruence.v=20=09modifi?= =?UTF-8?q?=C3=A9=C2=A0:=20=20=20=20=20=20=20=20=20theories/reduction/Subj?= =?UTF-8?q?ectReduction.v?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- theories/Scoping.v | 2 +- theories/reduction/Injectivity.v | 53 ++++- theories/reduction/ReductionProperties.v | 16 +- theories/reduction/ReductionToCongruence.v | 1 + theories/reduction/SubjectReduction.v | 256 +++++++++++++++++++-- 5 files changed, 298 insertions(+), 30 deletions(-) diff --git a/theories/Scoping.v b/theories/Scoping.v index 86f2d27..eb11b23 100644 --- a/theories/Scoping.v +++ b/theories/Scoping.v @@ -19,7 +19,7 @@ Inductive scoping (Γ : scope) : term → mode → Prop := nth_error Γ x = Some m → scoping Γ (var x) m -| scpoe_sort : +| scope_sort : ∀ m i, scoping Γ (Sort m i) mKind diff --git a/theories/reduction/Injectivity.v b/theories/reduction/Injectivity.v index 80be86d..c511b09 100644 --- a/theories/reduction/Injectivity.v +++ b/theories/reduction/Injectivity.v @@ -3,12 +3,28 @@ From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. From GhostTT.reduction Require Import ReductionToCongruence. +From GhostTT Require Import Model. From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. Import ListNotations. Set Default Goal Selector "!". +Lemma castrm_castrm {t : term} : + ε| ε|t| | = ε|t| . +Proof. + induction t. + all: cbn; congruence. +Qed. + +Lemma castrm_castrm_conv {Γ : context} {t t': term} : + Γ ⊢ ε|t| ε≡ ε|t'| → Γ ⊢ t ε≡ t'. +Proof. + intro conv. + do 2 rewrite castrm_castrm in conv. + exact conv. +Qed. + Lemma glength_castrm {A n v: term} : ε|glength A n v| = glength ε|A| ε|n| ε|v| . Proof. @@ -36,14 +52,16 @@ Proof. reflexivity. Qed. + Theorem injectivity_lam {Γ : context} {m md_t md_t': mode} {A A' t t': term} : md_t ≠ ℙ → - (sc Γ) ⊢ A∷𝕂 → m::(sc Γ)⊢t∷md_t → - (sc Γ) ⊢ A'∷𝕂 → m::(sc Γ)⊢t'∷md_t → + (sc Γ) ⊢ lam m A t∷md_t → + (sc Γ) ⊢ lam m A' t'∷md_t → Γ ⊢ lam m A t ≡ lam m A' t' → Γ ⊢ A ε≡ A' ∧ (m,A)::Γ ⊢ t ε≡ t'. Proof. - intros not_Prop scope_A scope_t scope_A' scope_t' H. + intros not_Prop scope_lam scope_lam' H. + inversion scope_lam; inversion scope_lam'; subst. apply conversion_to_reduction in H. destruct H as [w [red1 red2]]. inversion red1. @@ -52,7 +70,7 @@ Proof. * apply reds_lam_inv in red2 as [* [* [e []]]]. 2: cbn; erewrite scoping_md; eauto. inversion e. - split; apply conv_sym; subst. + split; apply conv_sym. all: eapply reductions_to_conversion; cbn; eauto. - inversion red2; subst. * apply reds_lam_inv in red1 as [* [* [e []]]]. @@ -70,12 +88,13 @@ Proof. Qed. Theorem injectivity_Pi {Γ : context} {i i' j j': level} {m m' mx mx': mode} {A A' B B': term} : - (sc Γ) ⊢ A∷𝕂 → mx::(sc Γ)⊢B∷𝕂 → - (sc Γ) ⊢ A'∷𝕂 → mx'::(sc Γ)⊢B'∷𝕂 → + sc Γ ⊢ Pi i j m mx A B ∷ 𝕂 → + sc Γ ⊢ Pi i' j' m' mx' A' B'∷ 𝕂 → Γ ⊢ Pi i j m mx A B ≡ Pi i' j' m' mx' A' B' → Γ ⊢ A ε≡ A' ∧ (mx,A)::Γ ⊢ B ε≡ B'. Proof. - intros scope_A scope_B scope_A' scope_B' H. + intros scope_Pi scope_Pi' H. + inversion scope_Pi; inversion scope_Pi'; subst. apply conversion_to_reduction in H. destruct H as [w [red1 red2]]. inversion red1. @@ -97,6 +116,26 @@ Proof. all: eapply reductions_to_conversion; eauto. Qed. +Corollary injectivity_Pi_castrm {Γ : context} {i i' j j': level} {m m' mx mx': mode} {A A' B B': term} : + sc Γ ⊢ Pi i j m mx A B ∷ 𝕂 → + sc Γ ⊢ Pi i' j' m' mx' A' B'∷ 𝕂 → + Γ ⊢ Pi i j m mx A B ε≡ Pi i' j' m' mx' A' B' → + Γ ⊢ A ε≡ A' ∧ (mx,A)::Γ ⊢ B ε≡ B'. +Proof. + intros scope_Pi scope_Pi' H. + cbn in H. + apply injectivity_Pi in H as [conv_A conv_B]. + - apply castrm_castrm_conv in conv_A. + apply castrm_castrm_conv in conv_B. + split. + * assumption. + * eapply conv_upto; eauto. + - inversion scope_Pi. + gconv; eauto using scoping_castrm. + - inversion scope_Pi'. + gconv; eauto using scoping_castrm. +Qed. + Theorem injectivity_Sort {Γ : context} {i i': level} {m m' : mode} : Γ ⊢ Sort m i ≡ Sort m' i' → m' = m. Proof. diff --git a/theories/reduction/ReductionProperties.v b/theories/reduction/ReductionProperties.v index 0acf244..65707eb 100644 --- a/theories/reduction/ReductionProperties.v +++ b/theories/reduction/ReductionProperties.v @@ -44,10 +44,10 @@ Proof. erewrite H; eauto end. Qed. -Lemma red_md : - ∀ Γ t t', Γ ⊨ t ↣ t' → md Γ t = md Γ t'. +Lemma red_md {Γ : scope} {t t' : term} : + Γ ⊨ t ↣ t' → md Γ t = md Γ t'. Proof. - intros Γ t t' red_t. + intro red_t. induction red_t in red_t |- *. all: try solve [cbn; congruence]. - cbn in *. eapply eq_trans; eauto. @@ -63,10 +63,10 @@ Proof. cbn; rewrite H; reflexivity end. Qed. -Lemma red_scope : - ∀ Γ m t t', Γ ⊨ t ↣ t' → Γ ⊢ t∷m → Γ ⊢ t'∷m. +Lemma red_scope {Γ : scope} {m : mode} {t t' : term} : + Γ ⊨ t ↣ t' → Γ ⊢ t∷m → Γ ⊢ t'∷m. Proof. - intros Γ m t t' red_t scope_t. + intros red_t scope_t. induction red_t in Γ, m, t, t', red_t, scope_t |- *. all: try solve [inversion scope_t; gscope]. - inversion scope_t. @@ -160,8 +160,10 @@ Proof. match goal with | HC : 𝕂 = ℙ |- _ => inversion HC end. Qed. -Lemma red_hide_inv (Γ : scope) (t0 t' : term) (red_hide : Γ⊨hide t0 ↣t' ) : ∃ t0', t' = hide t0' ∧ Γ ⊨ t0 ↣ t0'. +Lemma red_hide_inv {Γ : scope} {t0 t' : term} : + Γ⊨hide t0 ↣t' → ∃ t0', t' = hide t0' ∧ Γ ⊨ t0 ↣ t0'. Proof. + intro red_hide. inversion red_hide; subst. - eauto. - eauto using red_refl. diff --git a/theories/reduction/ReductionToCongruence.v b/theories/reduction/ReductionToCongruence.v index b0e8c87..94cb23c 100644 --- a/theories/reduction/ReductionToCongruence.v +++ b/theories/reduction/ReductionToCongruence.v @@ -87,6 +87,7 @@ Proof. * eauto using red_scope. Qed. + Local Ltac conversion_to_reduction_exists := match goal with | H : ∃ _, _ ∧ _ |- diff --git a/theories/reduction/SubjectReduction.v b/theories/reduction/SubjectReduction.v index 770e039..01a44f9 100644 --- a/theories/reduction/SubjectReduction.v +++ b/theories/reduction/SubjectReduction.v @@ -2,7 +2,8 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. -From GhostTT.reduction Require Import ReductionProperties Injectivity. +From GhostTT Require Import Model. +From GhostTT.reduction Require Import ReductionProperties ReductionToCongruence Injectivity. From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. Import ListNotations. @@ -21,7 +22,31 @@ Proof. assumption. Qed. +Lemma type_lam_Pi_inv {Γ : context} {m mx: mode} {i j : level} {f A B : term}: + wf Γ → Γ ⊢ f : Pi i j m mx A B → (∃ A t, f = lam mx A t) ∨ (∃ n, f = var n). +Proof. + intros wfΓ type_f. + induction f. + 2-3: ttinv type_f Hconv; destruct_exists Hconv; repeat destruct Hconv as [_ Hconv]; admit. + - eauto. + - ttinv type_f Hconv; destruct_exists Hconv. + repeat destruct Hconv as [_ Hconv]. + specialize (scoping_type wfΓ type_f) as scope_Pi. + apply scope_pi_inv in scope_Pi. + destruct scope_Pi as [scope_A [scope_B _]]. + cbn in Hconv. + apply injectivity_Pi in Hconv; eauto using scoping_castrm. + * left. do 2 eexists. admit. +Abort. +Ltac ttinv_destruct h HN:= + ttinv h HN; destruct_exists HN; + let rec destruct_conj HN := + lazymatch type of HN with + | _ ∧ _ => let H := fresh "H" in + destruct HN as [H HN]; destruct_conj HN + |_ => idtac end + in destruct_conj HN. Theorem subjet_reduction (Γ: context) (A t t': term): wf Γ → (sc Γ) ⊨ t ↣ t' → Γ ⊢ t:A → mdc Γ t ≠ ℙ → Γ⊢ t':A. @@ -29,18 +54,219 @@ Proof. intros wfΓ red_t type_t not_Prop. remember (sc Γ) as Γ0 eqn:e0. induction red_t in Γ, Γ0, e0, wfΓ, A, red_t, type_t, not_Prop. - all: gtype. - - admit. - - specialize (validity Γ _ _ wfΓ type_t) as [scope_t [i type_A]]. - ttinv type_t type_t'. - destruct_exists type_t'. - repeat destruct type_t' as [H0 [ H1 [ H2 [H3 [H4 [H5 [H6[]]]]]]]]. - - eapply type_app. - * admit. - * exact (scoping_type wfΓ type_t). - * eapply red_scope. - 2: exact (proj1 (validity Γ _ _ wfΓ type_t)). - subst; gred. - * + all: subst; cbn in *. + (** Easy cases **) + (* refl *) + 27: auto. + (* t of mode ℙ *) + 27-30: contradiction. + + (** Extra hypothesis **) + all: specialize (validity Γ _ _ wfΓ type_t) as [scope_t [i_u type_A]]. + all: specialize (scoping_type wfΓ type_t) as scope_A. + all: ttinv_destruct type_t conv_A. + + (** Computations **) + (* Beta_red *) + 1: admit. + (* reveal_hide *) + 1: admit. + (* Sort *) + 7: {econstructor. 4: gtype. all: gtype. } + (* Pi *) + 7: admit. + (* recursive computation *) + (* end case : if_true if_false nat_elim_zero *) + 1-3: solve [ refine (type_conv Γ i_u _ _ A _ _ scope_A _ _ conv_A type_A); + [gscope |eauto using red_scope | gtype; reflexivity]]. + (* end_case : vec_elim_nil *) + 2: admit. + (* recursion case *) + 1-2: admit. + + (* contexts *) + (* conv_A correct *) + 4,5,10: solve [refine (type_conv Γ i_u _ _ A _ _ scope_A _ _ conv_A type_A); + [gscope | gscope; eauto using red_scope | + gtype; eauto using red_scope; + erewrite scoping_md; [ |eassumption]; + intro HC; inversion HC]]. + 1: { refine (type_conv Γ i_u _ _ A _ _ (scoping_type wfΓ type_t) _ _ conv_A type_A); + [gscope | gscope; eauto using red_scope | ]. + gtype; eauto using red_scope. + all: try (erewrite scoping_md;[|eassumption]; intro HC; inversion HC). + + econstructor; eauto using red_scope. + apply IHred_t1; eauto. + all: try (erewrite scoping_md;[|eassumption]; intro HC; inversion HC). + + admit. (* Γ ≡ Δ → Γ ⊢ B:A → Δ ⊢ B:A *) + } + (* change type via type_conv and cons_A*) + 1: assert (Γ ⊢ Pi x x0 x1 mx A' x2 ε≡ A) as conv_A'. + 1:{ eapply conv_trans. + 2: eassumption. + cbn; gconv. + 2-3: right; reflexivity. + apply conv_sym. + eapply reduction_to_conversion in red_t1; eauto. } + 2: assert (Γ ⊢ x4 <[ v'··] ε≡ A) as conv_A'. + 2:{ eapply conv_trans. + 2: eassumption. + do 2 rewrite castrm_subst. + eapply (conv_subst_r Γ (x::sc Γ)). + + apply sscoping_castrm. + eauto using sscoping_one, red_scope. + + apply sscoping_castrm. + eauto using sscoping_one, red_scope. + + apply scoping_castrm. eassumption. + + intro n; destruct n; cbn; gconv. + apply conv_sym. + eapply reduction_to_conversion in red_t2; eauto. } + 2: { refine (type_conv Γ i_u _ _ A _ _ scope_A _ _ conv_A' type_A); gscope; eauto using red_scope. + + admit. + + admit. + + eapply type_app. + ++ eassumption. + ++ eauto using red_scope. + ++ eauto using red_scope. + ++ gscope. + ++ gtype. reflexivity. + ++ admit. + ++ assumption. + ++ assumption. } + - + refine (type_conv Γ i_u _ _ A _ _ scope_A _ _ conv_A' type_A). + 1,2: cbn; gscope; eauto using red_scope. + * eapply red_scope; eauto. + erewrite scoping_md; eauto. + * eapply type_lam; gscope; eauto using red_scope. + + rewrite (scoping_md _ A0 𝕂); eauto; intro HC; inversion HC. + +admit. + + econstructor; eauto using red_scope. + apply IHred_t1; eauto. + rewrite (scoping_md _ A0 𝕂); eauto; intro HC; inversion HC. + + admit. + - refine (type_conv Γ i_u _ _ A _ _ (scoping_type wfΓ type_t) _ _ conv_A type_A); [gscope | gscope; eauto using red_scope | ]. + * cbn; destruct (mdc Γ p) in *; eauto. + * admit. + * apply type_reveal. gtype; eauto using red_scope. + erewrite scoping_md; [intro HC| eassumption]. inversion HC. + - refine (type_conv Γ i_u _ _ A _ _ (scoping_type wfΓ type_t) _ _ conv_A type_A); [gscope | gscope; eauto using red_scope | ]. + gtype; eauto using red_scope. + erewrite scoping_md; eauto; intro HC; inversion HC. + + + + * + eapply type_lam. + econstructor; gscope; eauto using red_scope. + * gtype. + * gscope. + * cbn in *; inversion scope_t; gscope. + all: eauto using red_scope. + * admit. + - admit. + - admit. + - specialize (validity Γ _ _ wfΓ type_t) as [_ [i' type_A]]. + ttinv type_t type_t'. + destruct_exists type_t'; destruct_conj type_t'. + eapply type_conv. + 4: apply type_sort. + * constructor. + * eauto using scoping_type. + * constructor. + * assumption. + * eassumption. + - specialize (validity Γ _ _ wfΓ type_t) as [_ [i' type_A]]. + ttinv type_t type_t'. + destruct_exists type_t'; destruct_conj type_t'. + eapply type_conv. + 4: {constructor; subst; eauto using red_scope. + + apply IHred_t1; eauto. + ++ admit. + ++ admit. + + apply IHred_t2; eauto. + ++ eapply wf_cons; eauto; admit. + ++ admit. + ++ admit. + } + * constructor. + * eauto using scoping_type. + * constructor; subst;eauto using red_scope. + * cbn in *. + assert (umax mx m (red_lvl mx i) (red_lvl m j) = umax mx m i j) as e. + { destruct m, mx; cbn. + all: reflexivity. } + rewrite e; assumption. + * eassumption. + - admit. + - specialize (validity Γ _ _ wfΓ type_t) as [_ [i type_A]]. + ttinv type_t type_t'. + destruct_exists type_t'; destruct_conj type_t'. + eapply type_conv. + 4: { gtype; eauto using red_scope. + + erewrite scoping_md; [ | eassumption]. + intro HC; inversion HC. + + admit. + + admit. + + admit. } + * gscope. eauto using red_scope. + * eauto using scoping_type. + * gscope; eauto using red_scope. + * eapply conv_trans. + + cbn. apply cong_Pi. + ++ apply conv_sym. eauto using reduction_to_conversion. + ++ gconv. + ++ right. reflexivity. + ++ right. reflexivity. + + eassumption. + * cbn in type_A. + erewrite scoping_md in *; eauto. + gscope. + - specialize (validity Γ _ _ wfΓ type_t) as [_ [i type_A]]. + ttinv type_t type_t'. + destruct_exists type_t'; destruct_conj type_t'. + admit. + - specialize (validity Γ _ _ wfΓ type_t) as [_ [i type_A]]. + ttinv type_t type_t'. + destruct_exists type_t'; destruct_conj type_t'. + eapply type_conv. + 4: { gtype; eauto using red_scope. + erewrite scoping_md; eauto. } + all: try eassumption. + * gscope. + * eauto using scoping_type. + * gscope; eauto using red_scope. + - specialize (validity Γ _ _ wfΓ type_t) as [_ [i type_A]]. + ttinv type_t type_t'. + destruct_exists type_t'; destruct_conj type_t'. + eapply type_conv. + 4: { gtype; eauto using red_scope. + erewrite scoping_md; [|eauto]. + intro HC; inversion HC. } + all: try eassumption. + * gscope. + * eauto using scoping_type. + * gscope; eauto using red_scope. + - specialize (validity Γ _ _ wfΓ type_t) as [_ [i type_A]]. + ttinv type_t type_t'. + destruct_exists type_t'; destruct_conj type_t'. + eapply type_conv. + 4: { gtype; eauto using red_scope; admit. } + all: try eassumption. + * gscope; eauto using red_scope. + * eauto using scoping_type. + * gscope; eauto using red_scope; admit. + * admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. +Admitted. From 04de14a9cde8ecc6a74bd2d00c0467466bbb6560 Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Tue, 18 Jun 2024 17:14:00 +0200 Subject: [PATCH 06/15] =?UTF-8?q?r=C3=A9organisation?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- _CoqProject | 17 +- theories/reduction/Injectivity.v | 3 +- theories/reduction/Notations.v | 37 +++ theories/reduction/SubjectReduction.v | 7 +- .../Confluence.v} | 4 +- .../Properties.v} | 4 +- .../reduction/{ => multisteps}/Reduction.v | 29 +- .../{ => multisteps}/ReductionToCongruence.v | 7 +- .../{ReductionRho.v => multisteps/Rho.v} | 5 +- .../Transitivity.v} | 9 +- theories/reduction/onestep/Reduction.v | 267 ++++++++++++++++++ 11 files changed, 338 insertions(+), 51 deletions(-) create mode 100644 theories/reduction/Notations.v rename theories/reduction/{ReductionConfluence.v => multisteps/Confluence.v} (93%) rename theories/reduction/{ReductionProperties.v => multisteps/Properties.v} (98%) rename theories/reduction/{ => multisteps}/Reduction.v (89%) rename theories/reduction/{ => multisteps}/ReductionToCongruence.v (95%) rename theories/reduction/{ReductionRho.v => multisteps/Rho.v} (99%) rename theories/reduction/{ReductionAndTransitivity.v => multisteps/Transitivity.v} (97%) create mode 100644 theories/reduction/onestep/Reduction.v diff --git a/_CoqProject b/_CoqProject index 08a2a5c..b5aa0b7 100644 --- a/_CoqProject +++ b/_CoqProject @@ -20,13 +20,13 @@ theories/TermNotations.v theories/Typing.v theories/BasicMetaTheory.v -theories/reduction/Reduction.v -theories/reduction/ReductionProperties.v -theories/reduction/ReductionRho.v -theories/reduction/ReductionAndTransitivity.v -theories/reduction/ReductionConfluence.v -theories/reduction/ReductionToCongruence.v -theories/reduction/Injectivity.v +theories/reduction/Notations.v +theories/reduction/multisteps/Reduction.v +theories/reduction/multisteps/Properties.v +theories/reduction/multisteps/Rho.v +theories/reduction/multisteps/Transitivity.v +theories/reduction/multisteps/Confluence.v +theories/reduction/multisteps/ReductionToCongruence.v theories/CScoping.v @@ -38,6 +38,9 @@ theories/Revival.v theories/Param.v theories/Model.v + +theories/reduction/Injectivity.v + theories/Admissible.v theories/FreeTheorem.v diff --git a/theories/reduction/Injectivity.v b/theories/reduction/Injectivity.v index c511b09..5a6e211 100644 --- a/theories/reduction/Injectivity.v +++ b/theories/reduction/Injectivity.v @@ -2,9 +2,8 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. -From GhostTT.reduction Require Import ReductionToCongruence. +From GhostTT.reduction.multisteps Require Import ReductionToCongruence. From GhostTT Require Import Model. -From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. Import ListNotations. diff --git a/theories/reduction/Notations.v b/theories/reduction/Notations.v new file mode 100644 index 0000000..24d6ff6 --- /dev/null +++ b/theories/reduction/Notations.v @@ -0,0 +1,37 @@ +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import BasicAST Scoping TermNotations. + + + +(** General notations **) +(* Substitution of var 0 *) +Notation "s '··'" := (scons s ids) (at level 1, format "s ··") : subst_scope. +(* A is of scope m *) +Notation "Γ ⊢ A ∷ m" := (scoping Γ A m) + (at level 80, A, m at next level, format "Γ ⊢ A ∷ m"). + +(*Mode abreviations*) +Notation ℙ := mProp. +Notation 𝔾 := mGhost. +Notation 𝕋 := mType. +Notation 𝕂 := mKind. + +Notation "⊤" := top. +Notation "⊥" := bot. + +(** Notation for reductions **) +(* Step by step reduction *) +Reserved Notation "Γ ⊨ u ↣ v" + (at level 80, u, v at next level, format "Γ ⊨ u ↣ v"). +Reserved Notation "Γ ⊨ u ↣* v" + (at level 80, u, v at next level, format "Γ ⊨ u ↣* v"). +Reserved Notation "Γ ⊨ u ε↣ v" + (at level 80, u, v at next level, format "Γ ⊨ u ε↣ v"). + + + +(* Multi-step reduction *) +Reserved Notation "Γ ⊨ u ⇶ v" + (at level 80, u, v at next level, format "Γ ⊨ u ⇶ v"). + diff --git a/theories/reduction/SubjectReduction.v b/theories/reduction/SubjectReduction.v index 01a44f9..f713e7f 100644 --- a/theories/reduction/SubjectReduction.v +++ b/theories/reduction/SubjectReduction.v @@ -48,8 +48,11 @@ Ltac ttinv_destruct h HN:= |_ => idtac end in destruct_conj HN. -Theorem subjet_reduction (Γ: context) (A t t': term): - wf Γ → (sc Γ) ⊨ t ↣ t' → Γ ⊢ t:A → mdc Γ t ≠ ℙ → Γ⊢ t':A. + + +Theorem subjet_reduction (Γ: context) (A t t': term) + (wfΓ : wf Γ) ( red_t : sc Γ ⊨ t ↣ t') (type_t : Γ ⊢ t:A) + no_red_irr red_t → Γ⊢ t':A. Proof. intros wfΓ red_t type_t not_Prop. remember (sc Γ) as Γ0 eqn:e0. diff --git a/theories/reduction/ReductionConfluence.v b/theories/reduction/multisteps/Confluence.v similarity index 93% rename from theories/reduction/ReductionConfluence.v rename to theories/reduction/multisteps/Confluence.v index fd67cc0..1c96c50 100644 --- a/theories/reduction/ReductionConfluence.v +++ b/theories/reduction/multisteps/Confluence.v @@ -3,8 +3,8 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. -From GhostTT.reduction Require Import ReductionProperties ReductionRho. -From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. +From GhostTT.reduction.multisteps Require Import Rho Transitivity. +From GhostTT.reduction.multisteps Require Export Reduction. Import ListNotations. diff --git a/theories/reduction/ReductionProperties.v b/theories/reduction/multisteps/Properties.v similarity index 98% rename from theories/reduction/ReductionProperties.v rename to theories/reduction/multisteps/Properties.v index 65707eb..83d6a76 100644 --- a/theories/reduction/ReductionProperties.v +++ b/theories/reduction/multisteps/Properties.v @@ -3,10 +3,10 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. -From GhostTT.reduction Require Export Reduction. +From GhostTT.reduction Require Import Notations. +From GhostTT.reduction.multisteps Require Export Reduction. Import ListNotations. - Set Default Goal Selector "!". Lemma md_ren' {Γ Δ :scope} {t: term} {ρ: nat → nat} (e : ∀ n, nth (ρ n) Γ 𝕋 = nth n Δ 𝕋) : diff --git a/theories/reduction/Reduction.v b/theories/reduction/multisteps/Reduction.v similarity index 89% rename from theories/reduction/Reduction.v rename to theories/reduction/multisteps/Reduction.v index 2a4c989..f851217 100644 --- a/theories/reduction/Reduction.v +++ b/theories/reduction/multisteps/Reduction.v @@ -1,34 +1,17 @@ -(* Definition of reduction rules which corresponds to the congruence relation *) -(* and proof that the system is confluent *) +(* Definition of multistep reduction rules which corresponds to the congruence relation *) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. -From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl TermMode Scoping BasicMetaTheory. From GhostTT Require Export Univ TermNotations Typing. +From GhostTT.reduction Require Export Notations. Import ListNotations. - Set Default Goal Selector "!". -Reserved Notation "Γ ⊨ u ↣ v" - (at level 80, u, v at next level, format "Γ ⊨ u ↣ v"). -Notation "s '··'" := (scons s ids) (at level 1, format "s ··") : subst_scope. -Notation "Γ ⊢ A ∷ m" := (scoping Γ A m) - (at level 80, A, m at next level, format "Γ ⊢ A ∷ m"). -Notation " [ Γ , s ] ⊢ A ∷ m" := (scoping (s::Γ) A m) - (at level 81, A, m at next level, format "[ Γ , s ] ⊢ A ∷ m"). - -Notation ℙ := mProp. -Notation 𝔾 := mGhost. -Notation 𝕋 := mType. -Notation 𝕂 := mKind. - -Notation "⋆" := - (lam ℙ bot (var 0)). - (* ------------------------------------------------------------------------- *) Section Properties_Star. - Lemma type_star : ∀ Γ, Γ ⊢ ⋆ : top. + Lemma type_star : ∀ Γ, Γ ⊢ ⋆ : ⊤. Proof. intro Γ. apply type_lam. @@ -290,7 +273,7 @@ End Rewriting. Notation "Γ ⊨ u ↣ v" := (reduction Γ u v). (* ------------------------------------------------------------------------- *) -(** rewriting automation **) +(* multi-step rewriting automation *) Create HintDb gtt_red discriminated. Hint Resolve red_beta red_reveal_hide red_if_true red_if_false red_nat_elim_zero @@ -303,4 +286,4 @@ Hint Resolve red_beta red_reveal_hide red_if_true red_if_false red_nat_elim_zero Ltac gred := unshelve typeclasses eauto with gtt_scope gtt_red shelvedb ; shelve_unifiable. -(** end rewriting automation **) +(** end multi-step rewriting automation **) diff --git a/theories/reduction/ReductionToCongruence.v b/theories/reduction/multisteps/ReductionToCongruence.v similarity index 95% rename from theories/reduction/ReductionToCongruence.v rename to theories/reduction/multisteps/ReductionToCongruence.v index 94cb23c..fd7b163 100644 --- a/theories/reduction/ReductionToCongruence.v +++ b/theories/reduction/multisteps/ReductionToCongruence.v @@ -2,13 +2,12 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. -From GhostTT.reduction Require Export ReductionProperties ReductionConfluence. -From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. +From GhostTT.reduction.multisteps Require Export Properties Confluence. +From GhostTT.reduction.multisteps Require Export Reduction Transitivity. Import ListNotations. -Notation "Γ ⊨ u ε↣ v" := (Γ ⊨ ε|u| ↣ ε|v|) - (at level 80, u, v at next level, format "Γ ⊨ u ε↣ v"). +Notation "Γ ⊨ u ε↣ v" := (Γ ⊨ ε|u| ↣ ε|v|). Set Default Goal Selector "!". diff --git a/theories/reduction/ReductionRho.v b/theories/reduction/multisteps/Rho.v similarity index 99% rename from theories/reduction/ReductionRho.v rename to theories/reduction/multisteps/Rho.v index 0c5fc43..9d73f13 100644 --- a/theories/reduction/ReductionRho.v +++ b/theories/reduction/multisteps/Rho.v @@ -3,10 +3,7 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. -From GhostTT.reduction Require Import ReductionProperties. -From GhostTT.reduction Require Export Reduction. - -Import ListNotations. +From GhostTT.reduction.multisteps Require Import Properties. Set Default Goal Selector "!". diff --git a/theories/reduction/ReductionAndTransitivity.v b/theories/reduction/multisteps/Transitivity.v similarity index 97% rename from theories/reduction/ReductionAndTransitivity.v rename to theories/reduction/multisteps/Transitivity.v index b9936c2..90335f1 100644 --- a/theories/reduction/ReductionAndTransitivity.v +++ b/theories/reduction/multisteps/Transitivity.v @@ -2,19 +2,18 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. -From GhostTT.reduction Require Import ReductionProperties. -From GhostTT.reduction Require Export Reduction. +From GhostTT.reduction.multisteps Require Import Properties. +From GhostTT.reduction.multisteps Require Export Reduction. Import ListNotations. - Set Default Goal Selector "!". + (* Definition *) Inductive reduction_trans (Γ : scope) (u v: term) : Prop := | Refl: u = v → reduction_trans Γ u v | Trans w : Γ ⊨ u ↣ w → reduction_trans Γ w v → reduction_trans Γ u v. -Notation "Γ ⊨ u ↣* v" := (reduction_trans Γ u v) - (at level 80, u, v at next level, format "Γ ⊨ u ↣* v"). +Notation "Γ ⊨ u ↣* v" := (reduction_trans Γ u v). (* Usefull properties *) Lemma red_trans_direct {Γ : scope } {u v: term} : Γ ⊨ u ↣ v → Γ ⊨ u ↣* v. diff --git a/theories/reduction/onestep/Reduction.v b/theories/reduction/onestep/Reduction.v new file mode 100644 index 0000000..5ebfb9b --- /dev/null +++ b/theories/reduction/onestep/Reduction.v @@ -0,0 +1,267 @@ +(* Definition of reduction rules which corresponds to the congruence relation *) +(* and proof that the system is confluent *) +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. +From GhostTT Require Export Univ TermNotations Typing. +From GhostTT.reduction Require Export Notations. + +Import ListNotations. +Set Default Goal Selector "!". + +(* ------------------------------------------------------------------------- *) +Section Rewriting. + (** rewriting rules **) + + Definition mode_eq : ∀ (x y : mode), { x = y } + { x ≠ y }. + Proof. + decide equality. + Defined. + + Definition red_lvl (m : mode) (i : level) : level := + if mode_eq m ℙ then 0 else i. + + Inductive reduction (Γ : scope) : term → term → Prop := + + (* Computation rules *) + + | red_beta : + ∀ mx A t t' u u', + (mx::Γ) ⊨ t ⤖ t' → + md Γ u = mx → + Γ ⊨ u ⤖ u' → + Γ ⊨ app (lam mx A t) u ⤖ t' <[ u' ·· ] + + | red_reveal_hide : + ∀ t P p t' p', + Γ ⊨ t ↣ t' → + Γ ⊨ p ↣ p' → + In (md Γ p) [ℙ;𝔾] → + Γ ⊨ reveal (hide t) P p ↣ app p' t' + + | red_if_true : + ∀ m P t f t', + Γ ⊨ t ↣ t' → + md Γ t = m → + Γ ⊨ tif m ttrue P t f ↣ t' + + | red_if_false : + ∀ m P t f f', + Γ ⊨ f ↣ f' → + md Γ f = m → + Γ ⊨ tif m tfalse P t f ↣ f' + + | red_nat_elim_zero : + ∀ m P z s z', + Γ ⊨ z ↣ z' → + md Γ z = m → + Γ ⊨ tnat_elim m tzero P z s ↣ z' + + | red_nat_elim_succ : + ∀ m P z s n P' z' s' n', + Γ ⊨ n ↣ n' → + Γ ⊨ P ↣ P' → + Γ ⊨ z ↣ z' → + Γ ⊨ s ↣ s' → + md Γ s = m → + Γ ⊨ tnat_elim m (tsucc n) P z s ↣ app (app s' n') (tnat_elim m n' P' z' s') + + | red_vec_elim_nil : + ∀ m A B P z s z', + Γ ⊨ z ↣ z' → + md Γ z = m → + Γ ⊨ tvec_elim m A (hide tzero) (tvnil B) P z s ↣ z' + + | red_vec_elim_cons : + ∀ m A a n n0 v P z s A' a' n' v' P' z' s', + Γ ⊨ A ↣ A' → + Γ ⊨ a ↣ a' → + Γ ⊨ n ↣ n' → + Γ ⊨ v ↣ v' → + Γ ⊨ P ↣ P' → + Γ ⊨ z ↣ z' → + Γ ⊨ s ↣ s' → + md Γ s = m → + Γ ⊨ tvec_elim m A n0 (tvcons a n v) P z s ↣ + app (app (app (app s' a') (glength A' n' v')) v') (tvec_elim m A' n' v' P' z' s') + + (* Congruence rules *) + + (* A rule to quotient away all levels of Prop, making it impredicatime *) + | red_Prop : + ∀ i, Γ ⊨ Sort ℙ i ↣ Sort ℙ 0 + + | red_Pi : + ∀ i j m mx A A' B B', + Γ ⊨ A ↣ A' → + (mx::Γ) ⊨ B ↣ B' → + Γ ⊨ Pi i j m mx A B ↣ Pi (red_lvl mx i) (red_lvl m j) m mx A' B' + + | red_Pi' : (* needed for red_subst *) + ∀ i j m mx A A' B B', + Γ ⊨ A ↣ A' → + (mx::Γ) ⊨ B ↣ B' → + Γ ⊨ Pi i j m mx A B ↣ Pi i j m mx A' B' + + | red_lam : + ∀ mx A A' t t', + Γ ⊨ A ↣ A' → + (mx::Γ) ⊨ t ↣ t' → + Γ ⊨ lam mx A t ↣ lam mx A' t' + + | red_app : + ∀ u u' v v', + Γ ⊨ u ↣ u' → + Γ ⊨ v ↣ v' → + Γ ⊨ app u v ↣ app u' v' + + | red_Erased : + ∀ A A', + Γ ⊨ A ↣ A' → + Γ ⊨ Erased A ↣ Erased A' + + | red_hide : + ∀ u u', + Γ ⊨ u ↣ u' → + Γ ⊨ hide u ↣ hide u' + + | red_reveal : + ∀ t t' P P' p p', + Γ ⊨ t ↣ t' → + Γ ⊨ P ↣ P' → + Γ ⊨ p ↣ p' → + Γ ⊨ reveal t P p ↣ reveal t' P' p' + + | red_Reveal : + ∀ t t' p p', + Γ ⊨ t ↣ t' → + Γ ⊨ p ↣ p' → + Γ ⊨ Reveal t p ↣ Reveal t' p' + + | red_gheq : + ∀ A A' u u' v v', + Γ ⊨ A ↣ A' → + Γ ⊨ u ↣ u' → + Γ ⊨ v ↣ v' → + Γ ⊨ gheq A u v ↣ gheq A' u' v' + + | red_if : + ∀ m b b' P P' t t' f f', + Γ ⊨ b ↣ b' → + Γ ⊨ P ↣ P' → + Γ ⊨ t ↣ t' → + Γ ⊨ f ↣ f' → + Γ ⊨ tif m b P t f ↣ tif m b' P' t' f' + + | red_succ : + ∀ n n', + Γ ⊨ n ↣ n' → + Γ ⊨ tsucc n ↣ tsucc n' + + | red_nat_elim : + ∀ m n n' P P' z z' s s', + Γ ⊨ n ↣ n' → + Γ ⊨ P ↣ P' → + Γ ⊨ z ↣ z' → + Γ ⊨ s ↣ s' → + Γ ⊨ tnat_elim m n P z s ↣ tnat_elim m n' P' z' s' + + | red_vec : + ∀ A A' n n', + Γ ⊨ A ↣ A' → + Γ ⊨ n ↣ n' → + Γ ⊨ tvec A n ↣ tvec A' n' + + | red_vnil : + ∀ A A', + Γ ⊨ A ↣ A' → + Γ ⊨ tvnil A ↣ tvnil A' + + | red_vcons : + ∀ a a' n n' v v', + Γ ⊨ a ↣ a' → + Γ ⊨ n ↣ n' → + Γ ⊨ v ↣ v' → + Γ ⊨ tvcons a n v ↣ tvcons a' n' v' + + | red_vec_elim : + ∀ m A A' n n' v v' P P' z z' s s', + Γ ⊨ A ↣ A' → + Γ ⊨ n ↣ n' → + Γ ⊨ v ↣ v' → + Γ ⊨ P ↣ P' → + Γ ⊨ z ↣ z' → + Γ ⊨ s ↣ s' → + Γ ⊨ tvec_elim m A n v P z s ↣ tvec_elim m A' n' v' P' z' s' + + | red_bot_elim : + ∀ m A A' p p', + Γ ⊨ A ↣ A' → + Γ ⊨ p ↣ p' → + Γ ⊨ bot_elim m A p ↣ bot_elim m A' p' + + (* Structural rule *) + + | red_refl : + ∀ u, + Γ ⊨ u ↣ u + + (* Proof irrelevance *) + + | red_irr : + ∀ p, + md Γ p = ℙ → + Γ ⊨ p ↣ ⋆ + + | red_toRev : + ∀ t t' p p' u u', + Γ ⊨ t ↣ t' → + Γ ⊨ p ↣ p' → + Γ ⊨ u ↣ u' → + Γ ⊨ toRev t p u ↣ toRev t' p' u' + + | red_fromRev : + ∀ t t' p p' u u', + Γ ⊨ t ↣ t' → + Γ ⊨ p ↣ p' → + Γ ⊨ u ↣ u' → + Γ ⊨ fromRev t p u ↣ fromRev t' p' u' + + | red_ghrefl : + ∀ A A' u u', + Γ ⊨ A ↣ A' → + Γ ⊨ u ↣ u' → + Γ ⊨ ghrefl A u ↣ ghrefl A' u' + + | red_ghcast : + ∀ A A' u u' v v' e e' P P' t t', + Γ ⊨ A ↣ A' → + Γ ⊨ u ↣ u' → + Γ ⊨ v ↣ v' → + Γ ⊨ e ↣ e' → + Γ ⊨ P ↣ P' → + Γ ⊨ t ↣ t' → + Γ ⊨ ghcast A u v e P t ↣ ghcast A' u' v' e' P' t' + + + where "Γ ⊨ u ↣ v" := (reduction Γ u v). + + +End Rewriting. +Notation "Γ ⊨ u ↣ v" := (reduction Γ u v). + +(* ------------------------------------------------------------------------- *) +(** rewriting automation **) +Create HintDb gtt_red discriminated. + +Hint Resolve red_beta red_reveal_hide red_if_true red_if_false red_nat_elim_zero + red_nat_elim_succ red_vec_elim_nil red_vec_elim_cons red_Prop red_Pi + red_lam red_app red_Erased red_hide red_reveal red_Reveal red_gheq + red_if red_succ red_nat_elim red_vec red_vnil red_vcons red_vec_elim + red_bot_elim red_refl red_irr + red_Pi' red_toRev red_fromRev red_ghrefl red_ghcast + : gtt_red. + +Ltac gred := + unshelve typeclasses eauto with gtt_scope gtt_red shelvedb ; shelve_unifiable. +(** end rewriting automation **) From bed85317eea951bd0912cd4df688b4f487894eaf Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Wed, 19 Jun 2024 14:09:52 +0200 Subject: [PATCH 07/15] visio --- theories/reduction/onestep/Reduction.v | 282 ++++--------------------- 1 file changed, 45 insertions(+), 237 deletions(-) diff --git a/theories/reduction/onestep/Reduction.v b/theories/reduction/onestep/Reduction.v index 5ebfb9b..c0571ba 100644 --- a/theories/reduction/onestep/Reduction.v +++ b/theories/reduction/onestep/Reduction.v @@ -4,264 +4,72 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. From GhostTT Require Export Univ TermNotations Typing. -From GhostTT.reduction Require Export Notations. +From GhostTT.reduction Require Export Notations Utils. Import ListNotations. Set Default Goal Selector "!". (* ------------------------------------------------------------------------- *) -Section Rewriting. - (** rewriting rules **) +Reserved Notation "t ↣ t'" (at level 70). - Definition mode_eq : ∀ (x y : mode), { x = y } + { x ≠ y }. - Proof. - decide equality. - Defined. +Inductive reduction : term → term → Prop := - Definition red_lvl (m : mode) (i : level) : level := - if mode_eq m ℙ then 0 else i. + (* Computation rules *) - Inductive reduction (Γ : scope) : term → term → Prop := + | red_beta : + ∀ mx A t u, + app (lam mx A t) u ↣ t <[ u ·· ] - (* Computation rules *) + | red_reveal_hide : + ∀ t P p, + reveal (hide t) P p ↣ app p t - | red_beta : - ∀ mx A t t' u u', - (mx::Γ) ⊨ t ⤖ t' → - md Γ u = mx → - Γ ⊨ u ⤖ u' → - Γ ⊨ app (lam mx A t) u ⤖ t' <[ u' ·· ] + | red_if_true : + ∀ m P t f, tif m ttrue P t f ↣ t - | red_reveal_hide : - ∀ t P p t' p', - Γ ⊨ t ↣ t' → - Γ ⊨ p ↣ p' → - In (md Γ p) [ℙ;𝔾] → - Γ ⊨ reveal (hide t) P p ↣ app p' t' + | red_if_false : + ∀ m P t f, tif m tfalse P t f ↣ f - | red_if_true : - ∀ m P t f t', - Γ ⊨ t ↣ t' → - md Γ t = m → - Γ ⊨ tif m ttrue P t f ↣ t' + | red_nat_elim_zero : + ∀ m P z s, tnat_elim m tzero P z s ↣ z - | red_if_false : - ∀ m P t f f', - Γ ⊨ f ↣ f' → - md Γ f = m → - Γ ⊨ tif m tfalse P t f ↣ f' + | red_nat_elim_succ : + ∀ m P z s n, + tnat_elim m (tsucc n) P z s ↣ app (app s n) (tnat_elim m n P z s) - | red_nat_elim_zero : - ∀ m P z s z', - Γ ⊨ z ↣ z' → - md Γ z = m → - Γ ⊨ tnat_elim m tzero P z s ↣ z' + | red_vec_elim_nil : + ∀ m A B P z s, + tvec_elim m A (hide tzero) (tvnil B) P z s ↣ z - | red_nat_elim_succ : - ∀ m P z s n P' z' s' n', - Γ ⊨ n ↣ n' → - Γ ⊨ P ↣ P' → - Γ ⊨ z ↣ z' → - Γ ⊨ s ↣ s' → - md Γ s = m → - Γ ⊨ tnat_elim m (tsucc n) P z s ↣ app (app s' n') (tnat_elim m n' P' z' s') + | red_vec_elim_cons : + ∀ m A a n n0 v P z s, + tvec_elim m A n0 (tvcons a n v) P z s ↣ + app (app (app (app s a) (glength A n v)) v) (tvec_elim m A n v P z s) - | red_vec_elim_nil : - ∀ m A B P z s z', - Γ ⊨ z ↣ z' → - md Γ z = m → - Γ ⊨ tvec_elim m A (hide tzero) (tvnil B) P z s ↣ z' + | red_Prop : + ∀ i, Sort ℙ i ↣ Sort ℙ 0 - | red_vec_elim_cons : - ∀ m A a n n0 v P z s A' a' n' v' P' z' s', - Γ ⊨ A ↣ A' → - Γ ⊨ a ↣ a' → - Γ ⊨ n ↣ n' → - Γ ⊨ v ↣ v' → - Γ ⊨ P ↣ P' → - Γ ⊨ z ↣ z' → - Γ ⊨ s ↣ s' → - md Γ s = m → - Γ ⊨ tvec_elim m A n0 (tvcons a n v) P z s ↣ - app (app (app (app s' a') (glength A' n' v')) v') (tvec_elim m A' n' v' P' z' s') + | red_Pi : + ∀ i j m mx A B, + Pi i j m mx A B ↣ Pi (red_lvl mx i) (red_lvl m j) m mx A B - (* Congruence rules *) + | red_subterm : ∀ u u' C, + u ↣ u' → + C <[u··] ↣ C <[u'··] - (* A rule to quotient away all levels of Prop, making it impredicatime *) - | red_Prop : - ∀ i, Γ ⊨ Sort ℙ i ↣ Sort ℙ 0 + (* (* Structural rule *) - | red_Pi : - ∀ i j m mx A A' B B', - Γ ⊨ A ↣ A' → - (mx::Γ) ⊨ B ↣ B' → - Γ ⊨ Pi i j m mx A B ↣ Pi (red_lvl mx i) (red_lvl m j) m mx A' B' + | red_refl : + ∀ u, + Γ ⊨ u ↣ u *) - | red_Pi' : (* needed for red_subst *) - ∀ i j m mx A A' B B', - Γ ⊨ A ↣ A' → - (mx::Γ) ⊨ B ↣ B' → - Γ ⊨ Pi i j m mx A B ↣ Pi i j m mx A' B' + (* (* Proof irrelevance *) - | red_lam : - ∀ mx A A' t t', - Γ ⊨ A ↣ A' → - (mx::Γ) ⊨ t ↣ t' → - Γ ⊨ lam mx A t ↣ lam mx A' t' + | red_irr : + ∀ p, + md Γ p = ℙ → + Γ ⊨ p ↣ ⋆ *) - | red_app : - ∀ u u' v v', - Γ ⊨ u ↣ u' → - Γ ⊨ v ↣ v' → - Γ ⊨ app u v ↣ app u' v' + where "u ↣ v" := (reduction u v). - | red_Erased : - ∀ A A', - Γ ⊨ A ↣ A' → - Γ ⊨ Erased A ↣ Erased A' - - | red_hide : - ∀ u u', - Γ ⊨ u ↣ u' → - Γ ⊨ hide u ↣ hide u' - - | red_reveal : - ∀ t t' P P' p p', - Γ ⊨ t ↣ t' → - Γ ⊨ P ↣ P' → - Γ ⊨ p ↣ p' → - Γ ⊨ reveal t P p ↣ reveal t' P' p' - - | red_Reveal : - ∀ t t' p p', - Γ ⊨ t ↣ t' → - Γ ⊨ p ↣ p' → - Γ ⊨ Reveal t p ↣ Reveal t' p' - - | red_gheq : - ∀ A A' u u' v v', - Γ ⊨ A ↣ A' → - Γ ⊨ u ↣ u' → - Γ ⊨ v ↣ v' → - Γ ⊨ gheq A u v ↣ gheq A' u' v' - - | red_if : - ∀ m b b' P P' t t' f f', - Γ ⊨ b ↣ b' → - Γ ⊨ P ↣ P' → - Γ ⊨ t ↣ t' → - Γ ⊨ f ↣ f' → - Γ ⊨ tif m b P t f ↣ tif m b' P' t' f' - - | red_succ : - ∀ n n', - Γ ⊨ n ↣ n' → - Γ ⊨ tsucc n ↣ tsucc n' - - | red_nat_elim : - ∀ m n n' P P' z z' s s', - Γ ⊨ n ↣ n' → - Γ ⊨ P ↣ P' → - Γ ⊨ z ↣ z' → - Γ ⊨ s ↣ s' → - Γ ⊨ tnat_elim m n P z s ↣ tnat_elim m n' P' z' s' - - | red_vec : - ∀ A A' n n', - Γ ⊨ A ↣ A' → - Γ ⊨ n ↣ n' → - Γ ⊨ tvec A n ↣ tvec A' n' - - | red_vnil : - ∀ A A', - Γ ⊨ A ↣ A' → - Γ ⊨ tvnil A ↣ tvnil A' - - | red_vcons : - ∀ a a' n n' v v', - Γ ⊨ a ↣ a' → - Γ ⊨ n ↣ n' → - Γ ⊨ v ↣ v' → - Γ ⊨ tvcons a n v ↣ tvcons a' n' v' - - | red_vec_elim : - ∀ m A A' n n' v v' P P' z z' s s', - Γ ⊨ A ↣ A' → - Γ ⊨ n ↣ n' → - Γ ⊨ v ↣ v' → - Γ ⊨ P ↣ P' → - Γ ⊨ z ↣ z' → - Γ ⊨ s ↣ s' → - Γ ⊨ tvec_elim m A n v P z s ↣ tvec_elim m A' n' v' P' z' s' - - | red_bot_elim : - ∀ m A A' p p', - Γ ⊨ A ↣ A' → - Γ ⊨ p ↣ p' → - Γ ⊨ bot_elim m A p ↣ bot_elim m A' p' - - (* Structural rule *) - - | red_refl : - ∀ u, - Γ ⊨ u ↣ u - - (* Proof irrelevance *) - - | red_irr : - ∀ p, - md Γ p = ℙ → - Γ ⊨ p ↣ ⋆ - - | red_toRev : - ∀ t t' p p' u u', - Γ ⊨ t ↣ t' → - Γ ⊨ p ↣ p' → - Γ ⊨ u ↣ u' → - Γ ⊨ toRev t p u ↣ toRev t' p' u' - - | red_fromRev : - ∀ t t' p p' u u', - Γ ⊨ t ↣ t' → - Γ ⊨ p ↣ p' → - Γ ⊨ u ↣ u' → - Γ ⊨ fromRev t p u ↣ fromRev t' p' u' - - | red_ghrefl : - ∀ A A' u u', - Γ ⊨ A ↣ A' → - Γ ⊨ u ↣ u' → - Γ ⊨ ghrefl A u ↣ ghrefl A' u' - - | red_ghcast : - ∀ A A' u u' v v' e e' P P' t t', - Γ ⊨ A ↣ A' → - Γ ⊨ u ↣ u' → - Γ ⊨ v ↣ v' → - Γ ⊨ e ↣ e' → - Γ ⊨ P ↣ P' → - Γ ⊨ t ↣ t' → - Γ ⊨ ghcast A u v e P t ↣ ghcast A' u' v' e' P' t' - - - where "Γ ⊨ u ↣ v" := (reduction Γ u v). - - -End Rewriting. -Notation "Γ ⊨ u ↣ v" := (reduction Γ u v). - -(* ------------------------------------------------------------------------- *) -(** rewriting automation **) -Create HintDb gtt_red discriminated. - -Hint Resolve red_beta red_reveal_hide red_if_true red_if_false red_nat_elim_zero - red_nat_elim_succ red_vec_elim_nil red_vec_elim_cons red_Prop red_Pi - red_lam red_app red_Erased red_hide red_reveal red_Reveal red_gheq - red_if red_succ red_nat_elim red_vec red_vnil red_vcons red_vec_elim - red_bot_elim red_refl red_irr - red_Pi' red_toRev red_fromRev red_ghrefl red_ghcast - : gtt_red. - -Ltac gred := - unshelve typeclasses eauto with gtt_scope gtt_red shelvedb ; shelve_unifiable. -(** end rewriting automation **) + (* ------------------------------------------------------------------------- *) From e42c324dcb7acc6421014fdb7f8956666073d8d8 Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Wed, 26 Jun 2024 13:09:22 +0200 Subject: [PATCH 08/15] toward subject reduction --- _CoqProject | 16 +- theories/Admissible.v | 62 +-- theories/Model.v | 268 +--------- theories/reduction/Injectivity.v | 77 ++- theories/reduction/Model.v | 290 +++++++++++ theories/reduction/Notations.v | 33 +- theories/reduction/SubjectReduction.v | 275 ---------- theories/reduction/Utils.v | 78 +++ theories/reduction/multisteps/Confluence.v | 14 +- theories/reduction/multisteps/Properties.v | 147 ++---- theories/reduction/multisteps/Reduction.v | 490 ++++++++---------- .../multisteps/ReductionToCongruence.v | 30 +- theories/reduction/multisteps/Rho.v | 14 +- theories/reduction/multisteps/Transitivity.v | 92 ++-- theories/reduction/onestep/Properties.v | 59 +++ theories/reduction/onestep/Reduction.v | 25 +- theories/reduction/onestep/SubjectReduction.v | 341 ++++++++++++ theories/reduction/wrapping/Core.v | 161 ++++++ theories/reduction/wrapping/Properties.v | 103 ++++ 19 files changed, 1535 insertions(+), 1040 deletions(-) create mode 100644 theories/reduction/Model.v delete mode 100644 theories/reduction/SubjectReduction.v create mode 100644 theories/reduction/Utils.v create mode 100644 theories/reduction/onestep/Properties.v create mode 100644 theories/reduction/onestep/SubjectReduction.v create mode 100644 theories/reduction/wrapping/Core.v create mode 100644 theories/reduction/wrapping/Properties.v diff --git a/_CoqProject b/_CoqProject index b5aa0b7..539dbb3 100644 --- a/_CoqProject +++ b/_CoqProject @@ -20,14 +20,27 @@ theories/TermNotations.v theories/Typing.v theories/BasicMetaTheory.v + theories/reduction/Notations.v +theories/reduction/Utils.v + theories/reduction/multisteps/Reduction.v +theories/reduction/multisteps/Transitivity.v theories/reduction/multisteps/Properties.v theories/reduction/multisteps/Rho.v -theories/reduction/multisteps/Transitivity.v theories/reduction/multisteps/Confluence.v theories/reduction/multisteps/ReductionToCongruence.v +theories/reduction/Injectivity.v +theories/reduction/Model.v + +theories/reduction/wrapping/Core.v +theories/reduction/wrapping/Properties.v + +theories/reduction/onestep/Reduction.v +theories/reduction/onestep/Properties.v + + theories/CScoping.v theories/CTyping.v @@ -39,7 +52,6 @@ theories/Param.v theories/Model.v -theories/reduction/Injectivity.v theories/Admissible.v diff --git a/theories/Admissible.v b/theories/Admissible.v index 21e14c8..2b6c7bf 100644 --- a/theories/Admissible.v +++ b/theories/Admissible.v @@ -12,6 +12,7 @@ From GhostTT.autosubst Require Import CCAST GAST core unscoped. From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl CScoping Scoping CTyping TermMode Typing BasicMetaTheory CCMetaTheory Erasure Revival Param Model. +From GhostTT.reduction Require Import Injectivity Model. From Coq Require Import Setoid Morphisms Relation_Definitions. Import ListNotations. @@ -22,14 +23,13 @@ Set Equations Transparent. Section Admissible. - Context Γ (hΓ : wf Γ). + Context {Γ : context} (hΓ : wf Γ). - Lemma wf_cons : - ∀ m i A, + Lemma wf_cons {m : mode} {i : level} {A : term} : Γ ⊢ A : Sort m i → wf (Γ ,, (m, A)). Proof. - intros m i A h. + intro h. econstructor. all: eauto. eapply mode_coherence. all: eauto. constructor. @@ -39,8 +39,7 @@ Section Admissible. econstructor ; eauto ; try eapply mode_coherence ; eauto using wf_cons ; try solve [ constructor ]. - Lemma type_pi : - ∀ i j mx m A B, + Lemma type_pi {i j : level} {mx m : mode} {A B : term } : Γ ⊢ A : Sort mx i → Γ ,, (mx, A) ⊢ B : Sort m j → Γ ⊢ Pi i j m mx A B : Sort m (umax mx m i j). @@ -48,8 +47,7 @@ Section Admissible. intros. adm. Qed. - Lemma type_lam : - ∀ mx m i j A B t, + Lemma type_lam {mx m : mode} {i j : level} {A B t : term} : Γ ⊢ A : Sort mx i → Γ ,, (mx, A) ⊢ B : Sort m j → Γ ,, (mx, A) ⊢ t : B → @@ -58,29 +56,26 @@ Section Admissible. intros. adm. Qed. - Lemma type_app : - ∀ i j mx m A B t u, + Lemma type_app {i j : level} {mx m : mode} {A B t u : term} : Γ ⊢ t : Pi i j m mx A B → Γ ⊢ u : A → Γ ⊢ app t u : B <[ u .. ]. Proof. - intros i j mx m A B t u ht hu. + intros ht hu. eapply validity in ht as hP. 2: assumption. destruct hP as [_ [k hP]]. ttinv hP hP'. intuition subst. adm. eapply type_pi. all: eauto. Qed. - Lemma type_erased : - ∀ i A, + Lemma type_erased {i : level} {A : term} : Γ ⊢ A : Sort mType i → Γ ⊢ Erased A : Sort mGhost i. Proof. intros. adm. Qed. - Lemma type_hide : - ∀ i A t, + Lemma type_hide {i : level} {A t : term} : Γ ⊢ A : Sort mType i → Γ ⊢ t : A → Γ ⊢ hide t : Erased A. @@ -88,15 +83,14 @@ Section Admissible. intros. adm. Qed. - Lemma type_reveal : - ∀ i j m A t P p, + Lemma type_reveal {i j : level} {m : mode} {A t P p : term} : In m [ mProp ; mGhost ] → Γ ⊢ t : Erased A → Γ ⊢ P : Erased A ⇒[ i | usup m j / mGhost | mKind ] Sort m j → Γ ⊢ p : Pi i j m mType A (app (S ⋅ P) (hide (var 0))) → Γ ⊢ reveal t P p : app P t. Proof. - intros i j m A t P p hm ht hP hp. + intros hm ht hP hp. eapply validity in ht as hE. 2: assumption. destruct hE as [_ [l hE]]. ttinv hE hA. destruct hA as [k ?]. intuition idtac. @@ -106,26 +100,25 @@ Section Admissible. eapply validity in hp as hp'. 2: assumption. destruct hp' as [_ [ll hp']]. ttinv hp' hp''. destruct hp'' as [? [? [? [? hc]]]]. - cbn in hc. apply sort_mode_inj in hc. subst. + cbn in hc. apply injectivity_Sort in hc. subst. adm. adm. Qed. - Lemma type_Reveal : - ∀ i A t p, + Lemma type_Reveal {i : level} {A t p : term} : Γ ⊢ t : Erased A → Γ ⊢ p : A ⇒[ i | 0 / mType | mKind ] Sort mProp 0 → Γ ⊢ Reveal t p : Sort mProp 0. Proof. - intros i A t p ht hp. + intros ht hp. eapply validity in ht as hE. 2: assumption. destruct hE as [_ [j hE]]. ttinv hE hA. destruct hA as [k [? [? hc]]]. - set (mt := mdc Γ t) in *. clearbody mt. apply sort_mode_inj in hc. subst. + set (mt := mdc Γ t) in *. clearbody mt. apply injectivity_Sort in hc. subst. eapply validity in hp as hp'. 2: assumption. destruct hp' as [_ [l hp']]. ttinv hp' hp''. destruct hp'' as [? [? [? [? hc]]]]. set (mp := mdc Γ p) in *. clearbody mp. cbn in hc. - apply sort_mode_inj in hc. subst. + apply injectivity_Sort in hc. subst. adm. Qed. @@ -141,7 +134,7 @@ Section Admissible. destruct hp' as [_ [l hp']]. ttinv hp' hp''. destruct hp'' as [? [? [? [? hc]]]]. set (mp := mdc Γ p) in *. clearbody mp. cbn in hc. - apply sort_mode_inj in hc. subst. + apply injectivity_Sort in hc. subst. adm. eapply meta_conv. - eapply type_app. all: eassumption. @@ -160,7 +153,7 @@ Section Admissible. destruct hp' as [_ [l hp']]. ttinv hp' hp''. destruct hp'' as [? [? [? [? hc]]]]. set (mp := mdc Γ p) in *. clearbody mp. cbn in hc. - apply sort_mode_inj in hc. subst. + apply injectivity_Sort in hc. subst. adm. eapply type_Reveal. 2: eassumption. eapply type_hide. all: eassumption. Qed. @@ -197,12 +190,12 @@ Section Admissible. destruct he' as [_ [l he']]. ttinv he' he''. destruct he'' as [? [? [? [? [? [? [? hc]]]]]]]. set (me := mdc Γ e) in *. clearbody me. cbn in hc. - apply sort_mode_inj in hc. subst. + apply injectivity_Sort in hc. subst. eapply validity in hP as hP'. 2: assumption. destruct hP' as [_ [lP hP']]. ttinv hP' hP''. destruct hP'' as [? [? [? [? hc]]]]. set (mp := mdc Γ P) in *. clearbody mp. cbn in hc. - apply sort_mode_inj in hc. subst. + apply injectivity_Sort in hc. subst. adm. eapply meta_conv. - eapply type_app. all: eauto. - cbn. reflexivity. @@ -230,8 +223,7 @@ Section Admissible. + reflexivity. Qed. - Lemma type_bot_elim : - ∀ i m A p, + Lemma type_bot_elim {i : level} {m : mode} {A p : term} : Γ ⊢ A : Sort m i → Γ ⊢ p : bot → Γ ⊢ bot_elim m A p : A. @@ -239,30 +231,28 @@ Section Admissible. intros. adm. Qed. - Lemma type_conv : - ∀ i m A B t, + Lemma type_conv {i : level} {m : mode} {A B t : term} : cscoping Γ t m → Γ ⊢ t : A → Γ ⊢ A ε≡ B → Γ ⊢ B : Sort m i → Γ ⊢ t : B. Proof. - intros i m A B t hst ht hc hB. + intros hst ht hc hB. eapply validity in ht as hE. 2: assumption. destruct hE as [_ [j hE]]. econstructor. all: eauto. all: eapply mode_coherence. all: eauto. all: constructor. Qed. - Lemma type_conv_alt : - ∀ i j m A B t, + Lemma type_conv_alt {i j : level} {m : mode} {A B t: term} : Γ ⊢ t : A → Γ ⊢ A ε≡ B → Γ ⊢ A : Sort m i → Γ ⊢ B : Sort m j → Γ ⊢ t : B. Proof. - intros i j m A B t ht hc hA hB. + intros ht hc hA hB. eapply type_conv. all: eauto. eapply mode_coherence. 3: eassumption. all: eassumption. Qed. diff --git a/theories/Model.v b/theories/Model.v index 53214ff..f518a3c 100644 --- a/theories/Model.v +++ b/theories/Model.v @@ -6,6 +6,7 @@ From GhostTT.autosubst Require Import CCAST GAST core unscoped. From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl CScoping Scoping CTyping TermMode Typing BasicMetaTheory CCMetaTheory Erasure Revival Param. +From GhostTT.reduction Require Export Injectivity Model. From Coq Require Import Setoid Morphisms Relation_Definitions. Import ListNotations. @@ -30,20 +31,6 @@ Axiom ctyval_inj : Γ ⊢ᶜ A ≡ A' ∧ Γ ⊢ᶜ a ≡ a'. -(** Deduced injectivity in GTT **) - -(** Injectivity of sort modes **) - -Lemma sort_mode_inj : - ∀ Γ m m' i i', - Γ ⊢ Sort m i ≡ Sort m' i' → - m = m'. -Proof. - intros Γ m m' i i' h. - eapply erase_conv in h as he. cbn in he. - destruct m, m'. all: try reflexivity. all: exfalso. all: cbn in he. - all: solve [ eapply ctyval_inj in he ; intuition discriminate ]. -Qed. (** Relative consistency **) @@ -64,259 +51,8 @@ Proof. eapply validity in ht as h. 2: constructor. destruct h as [hs [i h]]. ttinv h h'. cbn in h'. - eapply sort_mode_inj in h'. + eapply injectivity_Sort in h'. remember (md [] t) as m eqn:em in *. clear em. subst. cbn in htp. eexists. eassumption. Qed. - -(** Uniqueness of typing - - We show a restricted form of uniqueness ignoring universe levels. - This way we do not rely on the absence of cumulativity. - - In order to do this, we build a function which puts all universes to 0. - -**) - -Fixpoint urm (t : term) : term := - match t with - | var x => var x - | Sort m i => Sort m 0 - | Pi i j m mx A B => Pi 0 0 m mx (urm A) (urm B) - | lam mx A t => lam mx (urm A) (urm t) - | app u v => app (urm u) (urm v) - | Erased A => Erased (urm A) - | hide t => hide (urm t) - | reveal t P p => reveal (urm t) (urm P) (urm p) - | Reveal t P => Reveal (urm t) (urm P) - | toRev t p u => toRev (urm t) (urm p) (urm u) - | fromRev t p u => fromRev (urm t) (urm p) (urm u) - | gheq A u v => gheq (urm A) (urm u) (urm v) - | ghrefl A u => ghrefl (urm A) (urm u) - | ghcast A u v e P t => ghcast (urm A) (urm u) (urm v) (urm e) (urm P) (urm t) - | tbool => tbool - | ttrue => ttrue - | tfalse => tfalse - | tif m b P t f => tif m (urm b) (urm P) (urm t) (urm f) - | tnat => tnat - | tzero => tzero - | tsucc n => tsucc (urm n) - | tnat_elim m n P z s => tnat_elim m (urm n) (urm P) (urm z) (urm s) - | tvec A n => tvec (urm A) (urm n) - | tvnil A => tvnil (urm A) - | tvcons a n v => tvcons (urm a) (urm n) (urm v) - | tvec_elim m A n v P z s => tvec_elim m (urm A) (urm n) (urm v) (urm P) (urm z) (urm s) - | bot => bot - | bot_elim m A p => bot_elim m (urm A) (urm p) - end. - -Lemma urm_ren : - ∀ t ρ, - urm (ρ ⋅ t) = ρ ⋅ (urm t). -Proof. - intros t ρ. - induction t in ρ |- *. - all: solve [ cbn ; f_equal ; eauto ]. -Qed. - -Lemma urm_subst : - ∀ t σ, - urm (t <[ σ ]) = (urm t) <[ σ >> urm ]. -Proof. - intros t σ. - induction t in σ |- *. - all: try reflexivity. - all: try solve [ cbn ; f_equal ; eauto ]. - - cbn. f_equal. 1: eauto. - rewrite IHt2. apply ext_term. - intros []. 1: reflexivity. - cbn. ssimpl. rewrite urm_ren. reflexivity. - - cbn. f_equal. 1:eauto. - rewrite IHt2. apply ext_term. - intros []. 1: reflexivity. - cbn. ssimpl. rewrite urm_ren. reflexivity. -Qed. - -Lemma urm_scoping : - ∀ Γ t m, - scoping Γ t m → - scoping Γ (urm t) m. -Proof. - intros Γ t m h. - induction h. all: solve [ econstructor ; eauto ]. -Qed. - -Definition urm_ctx (Γ : context) := - map (λ '(m, A), (m, urm A)) Γ. - -Lemma sc_urm_ctx : - ∀ Γ, - sc (urm_ctx Γ) = sc Γ. -Proof. - intros Γ. - unfold sc, urm_ctx. rewrite map_map. - apply map_ext. intros [m A]. reflexivity. -Qed. - -Lemma urm_cscoping : - ∀ Γ t m, - cscoping Γ t m → - cscoping (urm_ctx Γ) (urm t) m. -Proof. - intros Γ t m h. rewrite sc_urm_ctx. - apply urm_scoping. assumption. -Qed. - -Lemma conv_urm : - ∀ Γ u v, - Γ ⊢ u ≡ v → - urm_ctx Γ ⊢ urm u ≡ urm v. -Proof. - intros Γ u v h. - induction h. - all: try solve [ cbn ; econstructor ; eauto using urm_cscoping ]. - - cbn. rewrite urm_subst. eapply conv_trans. - 1:{ - eapply conv_beta. - all: try eapply urm_cscoping ; eauto. - all: eapply urm_scoping. all: rewrite sc_urm_ctx. all: eassumption. - } - ssimpl. apply conv_refl. - - cbn. rewrite !urm_ren. constructor. - all: try eapply urm_cscoping ; eauto. - - cbn. constructor. all: eauto. - all: unfold ueq. all: eauto. -Qed. - -Notation "Γ ⊢ u ≈ v" := - (urm_ctx Γ ⊢ urm ε| u | ≡ urm ε| v |) - (at level 80, u, v at next level, format "Γ ⊢ u ≈ v"). - -Lemma urm_conv_aux : - ∀ Γ A A' B B', - Γ ⊢ A' ε≡ A → - Γ ⊢ B' ε≡ B → - Γ ⊢ A' ≈ B' → - Γ ⊢ A ≈ B. -Proof. - intros Γ A A' B B' hA hB h. - eapply conv_trans. - - apply conv_sym. eapply conv_urm. eassumption. - - eapply conv_trans. - 2:{ eapply conv_urm. eassumption. } - assumption. -Qed. - -Lemma conv_meta_refl : - ∀ Γ u v, - u = v → - Γ ⊢ u ≡ v. -Proof. - intros Γ u ? ->. - apply conv_refl. -Qed. - -Ltac unitac h1 h2 := - let h1' := fresh h1 in - let h2' := fresh h2 in - ttinv h1 h1' ; ttinv h2 h2' ; - destruct_exists h1' ; - destruct_exists h2' ; - intuition subst ; - eapply urm_conv_aux ; [ - eassumption .. - | idtac - ]. - -(** Injectivity of Π types in the source - - We assume injectivity in the source for ε-conversion (ε≡). - We argue that this should hold given it is essentially CC conversion. - Sadly our model doesn't allow us to obtain it for free and we would need to - resort to other methods to obtain it, typically the same used to get it for - CC (such as confluence of an underlying rewriting system). - - We only state it for the codomain - -**) - -Axiom pi_inj : - ∀ Γ i j m mx A B A' B', - Γ ⊢ Pi i j m mx A B ≡ Pi i j m mx A' B' → - Γ ,, (mx, A) ⊢ B ≡ B'. - -Lemma sscoping_urm : - ∀ Γ σ Δ, - sscoping Γ σ Δ → - sscoping Γ (σ >> urm) Δ. -Proof. - intros Γ σ Δ h. - induction h. - - constructor. - - constructor. - + assumption. - + ssimpl. eapply urm_scoping. assumption. -Qed. - -(* Conversion only requires the scope not the full context *) -Lemma conv_upto : - ∀ Γ Δ u v, - Γ ⊢ u ≡ v → - sc Γ = sc Δ → - Δ ⊢ u ≡ v. -Proof. - intros Γ Δ u v h e. - induction h in Δ, e |- *. - all: try solve [ cbn ; econstructor ; rewrite <- ?e ; eauto ]. - - constructor. all: eauto. - eapply IHh2. cbn. f_equal. assumption. - - constructor. all: eauto. - eapply IHh2. cbn. f_equal. assumption. -Qed. - -Lemma type_unique : - ∀ Γ t A B, - Γ ⊢ t : A → - Γ ⊢ t : B → - Γ ⊢ A ≈ B. -Proof. - intros Γ t A B hA hB. - induction t in Γ, A, B, hA, hB |- *. - all: try unitac hA hB. all: try apply conv_refl. - - apply conv_meta_refl. congruence. - - cbn. repeat scoping_fun. - eapply IHt2 in H10. 2: exact H9. - cbn in H10. - constructor. 1: apply conv_refl. 2,3: compute ; auto. - eapply conv_upto. 1: eassumption. - cbn. reflexivity. - - repeat scoping_fun. - eapply IHt1 in H8. 2: exact H7. - cbn in H8. eapply pi_inj in H8. - rewrite !castrm_subst. - rewrite !urm_subst. - eapply conv_subst. 2: eassumption. - apply sscoping_urm. apply sscoping_castrm. cbn. eapply sscoping_one. - rewrite sc_urm_ctx. eassumption. - - cbn. econstructor. eauto. - - cbn. constructor. 1: eauto. - gconv. -Qed. - -(** Modes are coherent with sorts **) - -Lemma mode_coherence : - ∀ Γ t A m i, - wf Γ → - Γ ⊢ A : Sort m i → - Γ ⊢ t : A → - cscoping Γ t m. -Proof. - intros Γ t A m i hΓ hA h. - eapply validity in h. 2: assumption. - destruct h as [hs [j hA']]. - eapply type_unique in hA'. 2: eapply hA. - cbn in hA'. eapply sort_mode_inj in hA'. subst. - assumption. -Qed. diff --git a/theories/reduction/Injectivity.v b/theories/reduction/Injectivity.v index 5a6e211..4aff5bc 100644 --- a/theories/reduction/Injectivity.v +++ b/theories/reduction/Injectivity.v @@ -3,7 +3,6 @@ From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. From GhostTT.reduction.multisteps Require Import ReductionToCongruence. -From GhostTT Require Import Model. Import ListNotations. @@ -34,7 +33,7 @@ Proof. Qed. Lemma red_castrm {Γ : scope} {t t' : term} : - Γ ⊨t ↣ t' → Γ ⊨ t ε↣ t'. + Γ ⊨t ⇶ t' → Γ ⊨ t ε⇶ t'. Proof. intro red_t. induction red_t. @@ -54,8 +53,8 @@ Qed. Theorem injectivity_lam {Γ : context} {m md_t md_t': mode} {A A' t t': term} : md_t ≠ ℙ → - (sc Γ) ⊢ lam m A t∷md_t → - (sc Γ) ⊢ lam m A' t'∷md_t → + Γ ⊢ lam m A t∷md_t → + Γ ⊢ lam m A' t'∷md_t → Γ ⊢ lam m A t ≡ lam m A' t' → Γ ⊢ A ε≡ A' ∧ (m,A)::Γ ⊢ t ε≡ t'. Proof. @@ -87,10 +86,10 @@ Proof. Qed. Theorem injectivity_Pi {Γ : context} {i i' j j': level} {m m' mx mx': mode} {A A' B B': term} : - sc Γ ⊢ Pi i j m mx A B ∷ 𝕂 → - sc Γ ⊢ Pi i' j' m' mx' A' B'∷ 𝕂 → + Γ ⊢ Pi i j m mx A B ∷ 𝕂 → + Γ ⊢ Pi i' j' m' mx' A' B'∷ 𝕂 → Γ ⊢ Pi i j m mx A B ≡ Pi i' j' m' mx' A' B' → - Γ ⊢ A ε≡ A' ∧ (mx,A)::Γ ⊢ B ε≡ B'. + m= m' ∧ mx = mx' ∧ Γ ⊢ A ε≡ A' ∧ (mx,A)::Γ ⊢ B ε≡ B'. Proof. intros scope_Pi scope_Pi' H. inversion scope_Pi; inversion scope_Pi'; subst. @@ -98,23 +97,26 @@ Proof. destruct H as [w [red1 red2]]. inversion red1. - inversion red2 as [e|]; subst. - * inversion e. split; gconv. + * inversion e. repeat split; gconv. * apply reds_Pi_inv in red2 as [* [* [* [* [e [ ]]]]]]. inversion e. - split; apply conv_sym; subst. + repeat split; auto; apply conv_sym; subst. all: eapply reductions_to_conversion; cbn; eauto. - inversion red2; subst. * apply reds_Pi_inv in red1 as [* [* [* [* [e [ ]]]]]]. inversion e. + repeat split; eauto using reductions_to_conversion. * apply reds_Pi_inv in red1 as [* [* [* [* [e [ ]]]]]]. apply reds_Pi_inv in red2 as [* [* [* [* [e'[ ]]]]]]. subst; inversion e'; subst. - split; eapply conv_trans. + repeat split; auto. + all: eapply conv_trans. 2,4: apply conv_sym. all: eapply reductions_to_conversion; eauto. Qed. +(* true but uses conv_upto so needs models => create another file Corollary injectivity_Pi_castrm {Γ : context} {i i' j j': level} {m m' mx mx': mode} {A A' B B': term} : sc Γ ⊢ Pi i j m mx A B ∷ 𝕂 → sc Γ ⊢ Pi i' j' m' mx' A' B'∷ 𝕂 → @@ -133,7 +135,7 @@ Proof. gconv; eauto using scoping_castrm. - inversion scope_Pi'. gconv; eauto using scoping_castrm. -Qed. + Qed. *) Theorem injectivity_Sort {Γ : context} {i i': level} {m m' : mode} : Γ ⊢ Sort m i ≡ Sort m' i' → m' = m. @@ -154,3 +156,56 @@ Proof. subst; inversion e'. reflexivity. Qed. + +Theorem injectivity_Erased {Γ : context} {t t': term} : + Γ ⊢ Erased t ∷ 𝕂 → Γ ⊢ Erased t' ∷ 𝕂 → + Γ ⊢ Erased t ≡ Erased t' → Γ ⊢ t ε≡ t'. +Proof. + intros scope_Erased scope_Erased' H. + inversion scope_Erased; inversion scope_Erased'; subst. + apply conversion_to_reduction in H. + destruct H as [w [red1 red2]]. + inversion red1. + - inversion red2 as [e|]; subst. + * inversion e. exact (conv_refl Γ ε|t|). + * apply reds_Erased_inv in red2 as [* [e red_t']]. + inversion e. apply conv_sym. + eapply reductions_to_conversion; eauto. + - inversion red2; subst. + * apply reds_Erased_inv in red1 as [* [e red_t']]. + inversion e. + eapply reductions_to_conversion; eauto. + * apply reds_Erased_inv in red2 as [* [e red_t']]. + apply reds_Erased_inv in red1 as [* [e' red_t'']]. + subst; inversion e; subst. + eapply conv_trans. + 2: apply conv_sym. + all: eapply reductions_to_conversion; eauto. +Qed. + + +Theorem injectivity_vec {Γ : context} {A A' n n': term} : + Γ ⊢ tvec A n ∷ 𝕂 → Γ ⊢ tvec A' n' ∷ 𝕂 → + Γ ⊢ tvec A n ≡ tvec A' n' → Γ ⊢ A ε≡ A' ∧ Γ ⊢ n ε≡ n'. +Proof. + intros scope_vec scope_vec' H. + inversion scope_vec; inversion scope_vec'; subst. + apply conversion_to_reduction in H. + destruct H as [w [red1 red2]]. + inversion red1. + - inversion red2 as [e|]; subst. + * inversion e. split; apply (conv_refl Γ). + * apply reds_vec_inv in red2 as [A0 [ n0 [e [red_A' red_n']]]]. + inversion e. split; apply conv_sym. + all: eapply reductions_to_conversion; eauto. + - inversion red2; subst. + * apply reds_vec_inv in red1 as [A0 [ n0 [e [red_A' red_n']]]]. + inversion e. + split; eapply reductions_to_conversion; eauto. + * apply reds_vec_inv in red2 as [A0 [ n0 [e [red_A' red_n']]]]. + apply reds_vec_inv in red1 as [A1 [ n1 [e' [red_A'' red_n'']]]]. + subst; inversion e; subst. + split; eapply conv_trans. + 2,4: apply conv_sym. + all: eapply reductions_to_conversion; eauto. +Qed. diff --git a/theories/reduction/Model.v b/theories/reduction/Model.v new file mode 100644 index 0000000..5dcb89b --- /dev/null +++ b/theories/reduction/Model.v @@ -0,0 +1,290 @@ +(*** Consequences of the model ***) + +From Coq Require Import Utf8 List Bool Lia. +From Equations Require Import Equations. +From GhostTT.autosubst Require Import GAST core unscoped. +From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl + Scoping TermMode Typing BasicMetaTheory. +From Coq Require Import Setoid Morphisms Relation_Definitions. +From GhostTT.reduction Require Import Notations Injectivity. + +Import ListNotations. +Import CombineNotations. + +Set Default Goal Selector "!". +Set Equations Transparent. + +(** Uniqueness of typing + + We show a restricted form of uniqueness ignoring universe levels. + This way we do not rely on the absence of cumulativity. + + In order to do this, we build a function which puts all universes to 0. + +**) + +Fixpoint urm (t : term) : term := + match t with + | var x => var x + | Sort m i => Sort m 0 + | Pi i j m mx A B => Pi 0 0 m mx (urm A) (urm B) + | lam mx A t => lam mx (urm A) (urm t) + | app u v => app (urm u) (urm v) + | Erased A => Erased (urm A) + | hide t => hide (urm t) + | reveal t P p => reveal (urm t) (urm P) (urm p) + | Reveal t P => Reveal (urm t) (urm P) + | toRev t p u => toRev (urm t) (urm p) (urm u) + | fromRev t p u => fromRev (urm t) (urm p) (urm u) + | gheq A u v => gheq (urm A) (urm u) (urm v) + | ghrefl A u => ghrefl (urm A) (urm u) + | ghcast A u v e P t => ghcast (urm A) (urm u) (urm v) (urm e) (urm P) (urm t) + | tbool => tbool + | ttrue => ttrue + | tfalse => tfalse + | tif m b P t f => tif m (urm b) (urm P) (urm t) (urm f) + | tnat => tnat + | tzero => tzero + | tsucc n => tsucc (urm n) + | tnat_elim m n P z s => tnat_elim m (urm n) (urm P) (urm z) (urm s) + | tvec A n => tvec (urm A) (urm n) + | tvnil A => tvnil (urm A) + | tvcons a n v => tvcons (urm a) (urm n) (urm v) + | tvec_elim m A n v P z s => tvec_elim m (urm A) (urm n) (urm v) (urm P) (urm z) (urm s) + | bot => bot + | bot_elim m A p => bot_elim m (urm A) (urm p) + end. + +Lemma urm_ren : + ∀ t ρ, + urm (ρ ⋅ t) = ρ ⋅ (urm t). +Proof. + intros t ρ. + induction t in ρ |- *. + all: solve [ cbn ; f_equal ; eauto ]. +Qed. + +Lemma urm_subst : + ∀ t σ, + urm (t <[ σ ]) = (urm t) <[ σ >> urm ]. +Proof. + intros t σ. + induction t in σ |- *. + all: try reflexivity. + all: try solve [ cbn ; f_equal ; eauto ]. + - cbn. f_equal. 1: eauto. + rewrite IHt2. apply ext_term. + intros []. 1: reflexivity. + cbn. ssimpl. rewrite urm_ren. reflexivity. + - cbn. f_equal. 1:eauto. + rewrite IHt2. apply ext_term. + intros []. 1: reflexivity. + cbn. ssimpl. rewrite urm_ren. reflexivity. +Qed. + +Lemma urm_scoping : + ∀ Γ t m, + scoping Γ t m → + scoping Γ (urm t) m. +Proof. + intros Γ t m h. + induction h. all: solve [ econstructor ; eauto ]. +Qed. + +Definition urm_ctx (Γ : context) := + map (λ '(m, A), (m, urm A)) Γ. + +Lemma sc_urm_ctx : + ∀ Γ, + sc (urm_ctx Γ) = sc Γ. +Proof. + intros Γ. + unfold sc, urm_ctx. rewrite map_map. + apply map_ext. intros [m A]. reflexivity. +Qed. + +Lemma urm_cscoping : + ∀ Γ t m, + cscoping Γ t m → + cscoping (urm_ctx Γ) (urm t) m. +Proof. + intros Γ t m h. rewrite sc_urm_ctx. + apply urm_scoping. assumption. +Qed. + +Lemma conv_urm : + ∀ Γ u v, + Γ ⊢ u ≡ v → + urm_ctx Γ ⊢ urm u ≡ urm v. +Proof. + intros Γ u v h. + induction h. + all: try solve [ cbn ; econstructor ; eauto using urm_cscoping ]. + - cbn. rewrite urm_subst. eapply conv_trans. + 1:{ + eapply conv_beta. + all: try eapply urm_cscoping ; eauto. + all: eapply urm_scoping. all: rewrite sc_urm_ctx. all: eassumption. + } + ssimpl. apply conv_refl. + - cbn. rewrite !urm_ren. constructor. + all: try eapply urm_cscoping ; eauto. + - cbn. constructor. all: eauto. + all: unfold ueq. all: eauto. +Qed. + +Notation "Γ ⊢ u ≈ v" := + (urm_ctx Γ ⊢ urm ε| u | ≡ urm ε| v |) + (at level 80, u, v at next level, format "Γ ⊢ u ≈ v"). + +Lemma urm_conv_aux : + ∀ Γ A A' B B', + Γ ⊢ A' ε≡ A → + Γ ⊢ B' ε≡ B → + Γ ⊢ A' ≈ B' → + Γ ⊢ A ≈ B. +Proof. + intros Γ A A' B B' hA hB h. + eapply conv_trans. + - apply conv_sym. eapply conv_urm. eassumption. + - eapply conv_trans. + 2:{ eapply conv_urm. eassumption. } + assumption. +Qed. + +Lemma conv_meta_refl : + ∀ Γ u v, + u = v → + Γ ⊢ u ≡ v. +Proof. + intros Γ u ? ->. + apply conv_refl. +Qed. + +Ltac unitac h1 h2 := + let h1' := fresh h1 in + let h2' := fresh h2 in + ttinv h1 h1' ; ttinv h2 h2' ; + destruct_exists h1' ; + destruct_exists h2' ; + intuition subst ; + eapply urm_conv_aux ; [ + eassumption .. + | idtac + ]. + +Lemma sscoping_urm : + ∀ Γ σ Δ, + sscoping Γ σ Δ → + sscoping Γ (σ >> urm) Δ. +Proof. + intros Γ σ Δ h. + induction h. + - constructor. + - constructor. + + assumption. + + ssimpl. eapply urm_scoping. assumption. +Qed. + +(* Conversion only requires the scope not the full context *) +Lemma conv_upto : + ∀ Γ Δ u v, + Γ ⊢ u ≡ v → + sc Γ = sc Δ → + Δ ⊢ u ≡ v. +Proof. + intros Γ Δ u v h e. + induction h in Δ, e |- *. + all: try solve [ cbn ; econstructor ; rewrite <- ?e ; eauto ]. + - constructor. all: eauto. + eapply IHh2. cbn. f_equal. assumption. + - constructor. all: eauto. + eapply IHh2. cbn. f_equal. assumption. +Qed. + +Lemma castrm_urm (t : term): + ε|urm t| = urm ε|t|. +Proof. + induction t; cbn; auto. + all: f_equal; auto. +Qed. + +Lemma type_unique {Γ : context} {t A B : term} : + Γ ⊢ t : A → + Γ ⊢ t : B → + Γ ⊢ A ≈ B. +Proof. + intros hA hB. + induction t in Γ, A, B, hA, hB |- *. + all: try unitac hA hB. all: try apply conv_refl. + - apply conv_meta_refl. congruence. + - cbn. repeat scoping_fun. + eapply IHt2 in H10. 2: exact H9. + cbn in H10. + constructor. 1: apply conv_refl. 2,3: compute ; auto. + eapply conv_upto. 1: eassumption. + cbn. reflexivity. + - repeat scoping_fun. + eapply IHt1 in H8. 2: exact H7. + cbn in H8. eapply injectivity_Pi in H8 as [_ [_ [_ H8]]]. + * rewrite !castrm_urm in H8. + rewrite !castrm_castrm in H8. + rewrite !castrm_subst. + rewrite !urm_subst. + eapply conv_subst. + 2: eassumption. + apply sscoping_urm. apply sscoping_castrm. cbn. eapply sscoping_one. + rewrite sc_urm_ctx. eassumption. + * gscope. + + apply urm_cscoping. + apply scoping_castrm. + assumption. + + rewrite sc_urm_ctx. + change ((Γ,, (x, x3)) ⊢ urm ε| x4 |∷𝕂). + rewrite <- sc_urm_ctx. + apply urm_cscoping. + apply scoping_castrm. + assumption. + * gscope. + + apply urm_cscoping. + apply scoping_castrm. + assumption. + + rewrite sc_urm_ctx. + change ((Γ,, (x, x9)) ⊢ urm ε| x10 |∷𝕂). + rewrite <- sc_urm_ctx. + apply urm_cscoping. + apply scoping_castrm. + assumption. + - cbn. econstructor. eauto. + - cbn. constructor. 1: eauto. + gconv. +Qed. + +(** Modes are coherent with sorts **) + +Lemma mode_coherence : + ∀ Γ t A m i, + wf Γ → + Γ ⊢ A : Sort m i → + Γ ⊢ t : A → + Γ ⊢ t ∷ m. +Proof. + intros Γ t A m i hΓ hA h. + eapply validity in h. 2: assumption. + destruct h as [hs [j hA']]. + eapply type_unique in hA'. 2: eapply hA. + cbn in hA'. eapply injectivity_Sort in hA'. subst. + assumption. +Qed. + +Lemma scoping_type {Γ: context} {A t: term}: + wf Γ → Γ ⊢ t:A → Γ ⊢ A∷𝕂. +Proof. + intros wfΓ type_t. + specialize (validity Γ _ _ wfΓ type_t) as [scope_t [i type_A]]. + specialize (validity Γ _ _ wfΓ type_A) as [scope_A [j type_scope]]. + ttinv type_scope type_scope. + apply injectivity_Sort in type_scope. + rewrite type_scope in *. + assumption. +Qed. diff --git a/theories/reduction/Notations.v b/theories/reduction/Notations.v index 24d6ff6..1797470 100644 --- a/theories/reduction/Notations.v +++ b/theories/reduction/Notations.v @@ -1,15 +1,18 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. -From GhostTT Require Import BasicAST Scoping TermNotations. - - +From GhostTT Require Import BasicAST ContextDecl Scoping TermNotations. (** General notations **) + (* Substitution of var 0 *) Notation "s '··'" := (scons s ids) (at level 1, format "s ··") : subst_scope. + +(* scope and context *) (* A is of scope m *) -Notation "Γ ⊢ A ∷ m" := (scoping Γ A m) - (at level 80, A, m at next level, format "Γ ⊢ A ∷ m"). +Notation "Γ ⊢ A ∷ m" := (scoping (sc Γ) A m) + (at level 80, Γ constr, A, m at next level, format "Γ ⊢ A ∷ m"). +Notation "Γ ⊨ A ∷ m" := (scoping Γ A m) + (at level 80, A, m at next level, format "Γ ⊨ A ∷ m"). (*Mode abreviations*) Notation ℙ := mProp. @@ -17,21 +20,21 @@ Notation 𝔾 := mGhost. Notation 𝕋 := mType. Notation 𝕂 := mKind. -Notation "⊤" := top. Notation "⊥" := bot. +Notation "⊤" := top. +Notation "⋆" := (lam ℙ ⊥ (var 0)). (** Notation for reductions **) -(* Step by step reduction *) -Reserved Notation "Γ ⊨ u ↣ v" - (at level 80, u, v at next level, format "Γ ⊨ u ↣ v"). -Reserved Notation "Γ ⊨ u ↣* v" - (at level 80, u, v at next level, format "Γ ⊨ u ↣* v"). -Reserved Notation "Γ ⊨ u ε↣ v" - (at level 80, u, v at next level, format "Γ ⊨ u ε↣ v"). - - (* Multi-step reduction *) Reserved Notation "Γ ⊨ u ⇶ v" (at level 80, u, v at next level, format "Γ ⊨ u ⇶ v"). +Reserved Notation "Γ ⊨ u ⇶* v" + (at level 80, u, v at next level, format "Γ ⊨ u ⇶* v"). +Reserved Notation "Γ ⊨ u ε⇶ v" + (at level 80, u, v at next level, format "Γ ⊨ u ε⇶ v"). +(* Step by step reduction *) +Reserved Notation "u ↣ v" + (at level 80, v at next level, format "u ↣ v"). + diff --git a/theories/reduction/SubjectReduction.v b/theories/reduction/SubjectReduction.v deleted file mode 100644 index f713e7f..0000000 --- a/theories/reduction/SubjectReduction.v +++ /dev/null @@ -1,275 +0,0 @@ -From Coq Require Import Utf8 List. -From GhostTT.autosubst Require Import GAST unscoped. -From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval - TermMode Scoping BasicMetaTheory. -From GhostTT Require Import Model. -From GhostTT.reduction Require Import ReductionProperties ReductionToCongruence Injectivity. -From GhostTT.reduction Require Export Reduction ReductionAndTransitivity. - -Import ListNotations. - -Set Default Goal Selector "!". - -Lemma scoping_type {Γ: context} {A t: term}: - wf Γ → Γ ⊢ t:A → sc Γ⊢ A∷𝕂. -Proof. - intros wfΓ type_t. - specialize (validity Γ _ _ wfΓ type_t) as [scope_t [i type_A]]. - specialize (validity Γ _ _ wfΓ type_A) as [scope_A [j type_scope]]. - ttinv type_scope type_scope. - apply injectivity_Sort in type_scope. - rewrite type_scope in *. - assumption. -Qed. - -Lemma type_lam_Pi_inv {Γ : context} {m mx: mode} {i j : level} {f A B : term}: - wf Γ → Γ ⊢ f : Pi i j m mx A B → (∃ A t, f = lam mx A t) ∨ (∃ n, f = var n). -Proof. - intros wfΓ type_f. - induction f. - 2-3: ttinv type_f Hconv; destruct_exists Hconv; repeat destruct Hconv as [_ Hconv]; admit. - - eauto. - - ttinv type_f Hconv; destruct_exists Hconv. - repeat destruct Hconv as [_ Hconv]. - specialize (scoping_type wfΓ type_f) as scope_Pi. - apply scope_pi_inv in scope_Pi. - destruct scope_Pi as [scope_A [scope_B _]]. - cbn in Hconv. - apply injectivity_Pi in Hconv; eauto using scoping_castrm. - * left. do 2 eexists. admit. -Abort. - -Ltac ttinv_destruct h HN:= - ttinv h HN; destruct_exists HN; - let rec destruct_conj HN := - lazymatch type of HN with - | _ ∧ _ => let H := fresh "H" in - destruct HN as [H HN]; destruct_conj HN - |_ => idtac end - in destruct_conj HN. - - - -Theorem subjet_reduction (Γ: context) (A t t': term) - (wfΓ : wf Γ) ( red_t : sc Γ ⊨ t ↣ t') (type_t : Γ ⊢ t:A) - no_red_irr red_t → Γ⊢ t':A. -Proof. - intros wfΓ red_t type_t not_Prop. - remember (sc Γ) as Γ0 eqn:e0. - induction red_t in Γ, Γ0, e0, wfΓ, A, red_t, type_t, not_Prop. - all: subst; cbn in *. - (** Easy cases **) - (* refl *) - 27: auto. - (* t of mode ℙ *) - 27-30: contradiction. - - (** Extra hypothesis **) - all: specialize (validity Γ _ _ wfΓ type_t) as [scope_t [i_u type_A]]. - all: specialize (scoping_type wfΓ type_t) as scope_A. - all: ttinv_destruct type_t conv_A. - - (** Computations **) - (* Beta_red *) - 1: admit. - (* reveal_hide *) - 1: admit. - (* Sort *) - 7: {econstructor. 4: gtype. all: gtype. } - (* Pi *) - 7: admit. - (* recursive computation *) - (* end case : if_true if_false nat_elim_zero *) - 1-3: solve [ refine (type_conv Γ i_u _ _ A _ _ scope_A _ _ conv_A type_A); - [gscope |eauto using red_scope | gtype; reflexivity]]. - (* end_case : vec_elim_nil *) - 2: admit. - (* recursion case *) - 1-2: admit. - - (* contexts *) - (* conv_A correct *) - 4,5,10: solve [refine (type_conv Γ i_u _ _ A _ _ scope_A _ _ conv_A type_A); - [gscope | gscope; eauto using red_scope | - gtype; eauto using red_scope; - erewrite scoping_md; [ |eassumption]; - intro HC; inversion HC]]. - 1: { refine (type_conv Γ i_u _ _ A _ _ (scoping_type wfΓ type_t) _ _ conv_A type_A); - [gscope | gscope; eauto using red_scope | ]. - gtype; eauto using red_scope. - all: try (erewrite scoping_md;[|eassumption]; intro HC; inversion HC). - + econstructor; eauto using red_scope. - apply IHred_t1; eauto. - all: try (erewrite scoping_md;[|eassumption]; intro HC; inversion HC). - + admit. (* Γ ≡ Δ → Γ ⊢ B:A → Δ ⊢ B:A *) - } - (* change type via type_conv and cons_A*) - 1: assert (Γ ⊢ Pi x x0 x1 mx A' x2 ε≡ A) as conv_A'. - 1:{ eapply conv_trans. - 2: eassumption. - cbn; gconv. - 2-3: right; reflexivity. - apply conv_sym. - eapply reduction_to_conversion in red_t1; eauto. } - 2: assert (Γ ⊢ x4 <[ v'··] ε≡ A) as conv_A'. - 2:{ eapply conv_trans. - 2: eassumption. - do 2 rewrite castrm_subst. - eapply (conv_subst_r Γ (x::sc Γ)). - + apply sscoping_castrm. - eauto using sscoping_one, red_scope. - + apply sscoping_castrm. - eauto using sscoping_one, red_scope. - + apply scoping_castrm. eassumption. - + intro n; destruct n; cbn; gconv. - apply conv_sym. - eapply reduction_to_conversion in red_t2; eauto. } - 2: { refine (type_conv Γ i_u _ _ A _ _ scope_A _ _ conv_A' type_A); gscope; eauto using red_scope. - + admit. - + admit. - + eapply type_app. - ++ eassumption. - ++ eauto using red_scope. - ++ eauto using red_scope. - ++ gscope. - ++ gtype. reflexivity. - ++ admit. - ++ assumption. - ++ assumption. } - - - refine (type_conv Γ i_u _ _ A _ _ scope_A _ _ conv_A' type_A). - 1,2: cbn; gscope; eauto using red_scope. - * eapply red_scope; eauto. - erewrite scoping_md; eauto. - * eapply type_lam; gscope; eauto using red_scope. - + rewrite (scoping_md _ A0 𝕂); eauto; intro HC; inversion HC. - +admit. - + econstructor; eauto using red_scope. - apply IHred_t1; eauto. - rewrite (scoping_md _ A0 𝕂); eauto; intro HC; inversion HC. - + admit. - - refine (type_conv Γ i_u _ _ A _ _ (scoping_type wfΓ type_t) _ _ conv_A type_A); [gscope | gscope; eauto using red_scope | ]. - * cbn; destruct (mdc Γ p) in *; eauto. - * admit. - * apply type_reveal. gtype; eauto using red_scope. - erewrite scoping_md; [intro HC| eassumption]. inversion HC. - - refine (type_conv Γ i_u _ _ A _ _ (scoping_type wfΓ type_t) _ _ conv_A type_A); [gscope | gscope; eauto using red_scope | ]. - gtype; eauto using red_scope. - erewrite scoping_md; eauto; intro HC; inversion HC. - - - - * - eapply type_lam. - - econstructor; gscope; eauto using red_scope. - * gtype. - * gscope. - * cbn in *; inversion scope_t; gscope. - all: eauto using red_scope. - * admit. - - admit. - - admit. - - specialize (validity Γ _ _ wfΓ type_t) as [_ [i' type_A]]. - ttinv type_t type_t'. - destruct_exists type_t'; destruct_conj type_t'. - eapply type_conv. - 4: apply type_sort. - * constructor. - * eauto using scoping_type. - * constructor. - * assumption. - * eassumption. - - specialize (validity Γ _ _ wfΓ type_t) as [_ [i' type_A]]. - ttinv type_t type_t'. - destruct_exists type_t'; destruct_conj type_t'. - eapply type_conv. - 4: {constructor; subst; eauto using red_scope. - + apply IHred_t1; eauto. - ++ admit. - ++ admit. - + apply IHred_t2; eauto. - ++ eapply wf_cons; eauto; admit. - ++ admit. - ++ admit. - } - * constructor. - * eauto using scoping_type. - * constructor; subst;eauto using red_scope. - * cbn in *. - assert (umax mx m (red_lvl mx i) (red_lvl m j) = umax mx m i j) as e. - { destruct m, mx; cbn. - all: reflexivity. } - rewrite e; assumption. - * eassumption. - - admit. - - specialize (validity Γ _ _ wfΓ type_t) as [_ [i type_A]]. - ttinv type_t type_t'. - destruct_exists type_t'; destruct_conj type_t'. - eapply type_conv. - 4: { gtype; eauto using red_scope. - + erewrite scoping_md; [ | eassumption]. - intro HC; inversion HC. - + admit. - + admit. - + admit. } - * gscope. eauto using red_scope. - * eauto using scoping_type. - * gscope; eauto using red_scope. - * eapply conv_trans. - + cbn. apply cong_Pi. - ++ apply conv_sym. eauto using reduction_to_conversion. - ++ gconv. - ++ right. reflexivity. - ++ right. reflexivity. - + eassumption. - * cbn in type_A. - erewrite scoping_md in *; eauto. - gscope. - - specialize (validity Γ _ _ wfΓ type_t) as [_ [i type_A]]. - ttinv type_t type_t'. - destruct_exists type_t'; destruct_conj type_t'. - admit. - - specialize (validity Γ _ _ wfΓ type_t) as [_ [i type_A]]. - ttinv type_t type_t'. - destruct_exists type_t'; destruct_conj type_t'. - eapply type_conv. - 4: { gtype; eauto using red_scope. - erewrite scoping_md; eauto. } - all: try eassumption. - * gscope. - * eauto using scoping_type. - * gscope; eauto using red_scope. - - specialize (validity Γ _ _ wfΓ type_t) as [_ [i type_A]]. - ttinv type_t type_t'. - destruct_exists type_t'; destruct_conj type_t'. - eapply type_conv. - 4: { gtype; eauto using red_scope. - erewrite scoping_md; [|eauto]. - intro HC; inversion HC. } - all: try eassumption. - * gscope. - * eauto using scoping_type. - * gscope; eauto using red_scope. - - specialize (validity Γ _ _ wfΓ type_t) as [_ [i type_A]]. - ttinv type_t type_t'. - destruct_exists type_t'; destruct_conj type_t'. - eapply type_conv. - 4: { gtype; eauto using red_scope; admit. } - all: try eassumption. - * gscope; eauto using red_scope. - * eauto using scoping_type. - * gscope; eauto using red_scope; admit. - * admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. -Admitted. diff --git a/theories/reduction/Utils.v b/theories/reduction/Utils.v new file mode 100644 index 0000000..bd8e061 --- /dev/null +++ b/theories/reduction/Utils.v @@ -0,0 +1,78 @@ +(* Definition of multistep reduction rules which corresponds to the congruence relation *) +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl + TermMode Scoping BasicMetaTheory Univ TermNotations Typing. +From GhostTT.reduction Require Import Notations. + +Set Default Goal Selector "!". + +Definition mode_eq : ∀ (x y : mode), { x = y } + { x ≠ y }. +Proof. + decide equality. +Defined. + +Definition red_lvl (m : mode) (i : level) : level := + if mode_eq m ℙ then 0 else i. + +Derive NoConfusion Subterm for term. + +Lemma md_ren' {Γ Δ :scope} {t: term} {ρ: nat → nat} (e : ∀ n, nth (ρ n) Γ 𝕋 = nth n Δ 𝕋) : + md Δ t = md Γ (ρ ⋅ t). +Proof. + induction t in Γ, Δ, t, ρ, e|- *. + all: cbn; eauto. + - cbn. match goal with H: ∀ _ _, _ → _ |- _ => + eapply H; eauto end. + intro n; destruct n; cbn; auto. + - match goal with H: ∀ _ _, _ → _ |- _ => + erewrite H; eauto end. +Qed. + +Lemma md_up_term {Γ : scope} {m: mode} {σ : nat → term} {n : nat} : + md (m::Γ) (up_term σ (S n)) = md Γ (σ n). +Proof. + asimpl; ssimpl. + unfold shift. + symmetry. + apply md_ren'. + induction n; eauto. +Qed. + +Lemma md_subst' {Γ Δ :scope} {t: term} {σ: nat → term} (e : ∀ n, md Γ (σ n) = nth n Δ 𝕋) : + md Δ t = md Γ (t<[σ]). +Proof. + induction t in Γ, Δ, t, σ, e|- *. + all: cbn; eauto. + - match goal with H: ∀ _ _, _ → _ |- _ => + erewrite H; eauto end. + intro n; destruct n; eauto. + erewrite md_up_term. auto. + - match goal with H: ∀ _ _, _ → _ |- _ => + erewrite H; eauto end. +Qed. + + +Lemma glenght_red_subst (A n v : term) (σ : nat → term) : + (glength A n v)<[σ] = glength (A<[σ]) (n<[σ]) (v<[σ]). +Proof. + change (tvec_elim 𝔾 (A <[ σ]) (n <[ σ]) (v <[ σ]) + (lam 𝔾 (Erased tnat) + (lam 𝕋 ((tvec (S ⋅ A) (var 0))<[up_term σ]) (Erased tnat)) + ) + (hide tzero) + (lam 𝕋 (A<[σ]) + (lam 𝔾 (Erased tnat) + (lam 𝕋 (tvec (S ⋅ S ⋅ A) (var 0) <[up_term (up_term σ)]) + (lam 𝔾 (Erased tnat) + (gS (var 0)) + <[(up_term (up_term (up_term σ)))]) + ) + ) + ) + = glength (A<[σ]) (n<[σ]) (v<[σ])). + unfold glength. + repeat f_equal. + all: asimpl; reflexivity. +Qed. + diff --git a/theories/reduction/multisteps/Confluence.v b/theories/reduction/multisteps/Confluence.v index 1c96c50..8c85861 100644 --- a/theories/reduction/multisteps/Confluence.v +++ b/theories/reduction/multisteps/Confluence.v @@ -11,8 +11,8 @@ Import ListNotations. Set Default Goal Selector "!". Theorem reduction_diamond (Γ : scope) (t u v: term) : - Γ ⊨ t ↣ u → Γ ⊨ t ↣ v → - ∃ w, Γ ⊨ u ↣ w ∧ Γ ⊨ v ↣ w. + Γ ⊨ t ⇶ u → Γ ⊨ t ⇶ v → + ∃ w, Γ ⊨ u ⇶ w ∧ Γ ⊨ v ⇶ w. Proof. intros red_tu red_tv. exists (ρ Γ t). @@ -20,8 +20,8 @@ Proof. Qed. Theorem reduction_local_confluence (Γ : scope) (t u v: term) : - Γ ⊨ t ↣ u → Γ ⊨ t ↣ v → - ∃ w, Γ ⊨ u ↣* w ∧ Γ ⊨ v ↣* w. + Γ ⊨ t ⇶ u → Γ ⊨ t ⇶ v → + ∃ w, Γ ⊨ u ⇶* w ∧ Γ ⊨ v ⇶* w. Proof. intros red_t0 red_t1. exists (ρ Γ t). @@ -29,8 +29,8 @@ Proof. Qed. Theorem reduction_confluence (Γ : scope) (t u v: term) : - Γ ⊨ t ↣* u → Γ ⊨ t ↣* v → - ∃ w, Γ ⊨ u ↣* w ∧ Γ ⊨ v ↣* w. + Γ ⊨ t ⇶* u → Γ ⊨ t ⇶* v → + ∃ w, Γ ⊨ u ⇶* w ∧ Γ ⊨ v ⇶* w. Proof. intros red_tu red_tv. induction red_tu as [t u red_tu | t u1 u0 red_tu red_u0 IH] in u, v, red_tu, red_tv |- *. @@ -47,7 +47,7 @@ Proof. split. + constructor; reflexivity. + apply (Trans Γ v u1 u0); assumption. - * assert (∃ w, (Γ⊨(ρ Γ t)↣*w) ∧ Γ⊨v1↣*w) as H. + * assert (∃ w, (Γ⊨(ρ Γ t)⇶*w) ∧ Γ⊨v1⇶*w) as H. { eapply IH'; eauto. + eauto using reduction_triangle; eauto. + constructor; reflexivity. diff --git a/theories/reduction/multisteps/Properties.v b/theories/reduction/multisteps/Properties.v index 83d6a76..cf2be11 100644 --- a/theories/reduction/multisteps/Properties.v +++ b/theories/reduction/multisteps/Properties.v @@ -9,43 +9,8 @@ From GhostTT.reduction.multisteps Require Export Reduction. Import ListNotations. Set Default Goal Selector "!". -Lemma md_ren' {Γ Δ :scope} {t: term} {ρ: nat → nat} (e : ∀ n, nth (ρ n) Γ 𝕋 = nth n Δ 𝕋) : - md Δ t = md Γ (ρ ⋅ t). -Proof. - induction t in Γ, Δ, t, ρ, e|- *. - all: cbn; eauto. - - cbn. match goal with H: ∀ _ _, _ → _ |- _ => - eapply H; eauto end. - intro n; destruct n; cbn; auto. - - match goal with H: ∀ _ _, _ → _ |- _ => - erewrite H; eauto end. -Qed. - -Lemma md_up_term {Γ : scope} {m: mode} {σ : nat → term} {n : nat} : - md (m::Γ) (up_term σ (S n)) = md Γ (σ n). -Proof. - asimpl; ssimpl. - unfold shift. - symmetry. - apply md_ren'. - induction n; eauto. -Qed. - -Lemma md_subst' {Γ Δ :scope} {t: term} {σ: nat → term} (e : ∀ n, md Γ (σ n) = nth n Δ 𝕋) : - md Δ t = md Γ (t<[σ]). -Proof. - induction t in Γ, Δ, t, σ, e|- *. - all: cbn; eauto. - - match goal with H: ∀ _ _, _ → _ |- _ => - erewrite H; eauto end. - intro n; destruct n; eauto. - erewrite md_up_term. auto. - - match goal with H: ∀ _ _, _ → _ |- _ => - erewrite H; eauto end. -Qed. - Lemma red_md {Γ : scope} {t t' : term} : - Γ ⊨ t ↣ t' → md Γ t = md Γ t'. + Γ ⊨ t ⇶ t' → md Γ t = md Γ t'. Proof. intro red_t. induction red_t in red_t |- *. @@ -64,27 +29,27 @@ Proof. Qed. Lemma red_scope {Γ : scope} {m : mode} {t t' : term} : - Γ ⊨ t ↣ t' → Γ ⊢ t∷m → Γ ⊢ t'∷m. + Γ ⊨ t ⇶ t' → Γ ⊨ t∷m → Γ ⊨ t'∷m. Proof. intros red_t scope_t. induction red_t in Γ, m, t, t', red_t, scope_t |- *. all: try solve [inversion scope_t; gscope]. - inversion scope_t. - match goal with H : _ ⊢ lam _ _ _∷_ |- _ => + match goal with H : _ ⊨ lam _ _ _∷_ |- _ => inversion H; subst end. eapply scoping_subst; eauto. eapply sscoping_one. erewrite scoping_md; eauto. - inversion scope_t. - match goal with H : _ ⊢ hide _∷_ |- _ => + match goal with H : _ ⊨ hide _∷_ |- _ => inversion H; subst end. gscope. - inversion scope_t. - match goal with H : _ ⊢ tsucc _∷_ |- _ => + match goal with H : _ ⊨ tsucc _∷_ |- _ => inversion H; subst end. gscope. - inversion scope_t. - match goal with H : _ ⊢ tvcons _ _ _∷_ |- _ => + match goal with H : _ ⊨ tvcons _ _ _∷_ |- _ => inversion H; subst end. gscope; eauto. * intro H; inversion H. @@ -96,32 +61,10 @@ Proof. subst. apply scope_star. Qed. -Lemma glenght_red_subst (A n v : term) (σ : nat → term) : - (glength A n v)<[σ] = glength (A<[σ]) (n<[σ]) (v<[σ]). -Proof. - change (tvec_elim 𝔾 (A <[ σ]) (n <[ σ]) (v <[ σ]) - (lam 𝔾 (Erased tnat) - (lam 𝕋 ((tvec (S ⋅ A) (var 0))<[up_term σ]) (Erased tnat)) - ) - (hide tzero) - (lam 𝕋 (A<[σ]) - (lam 𝔾 (Erased tnat) - (lam 𝕋 (tvec (S ⋅ S ⋅ A) (var 0) <[up_term (up_term σ)]) - (lam 𝔾 (Erased tnat) - (gS (var 0)) - <[(up_term (up_term (up_term σ)))]) - ) - ) - ) - = glength (A<[σ]) (n<[σ]) (v<[σ])). - unfold glength. - repeat f_equal. - all: asimpl; reflexivity. -Qed. Lemma red_lam_inv {Γ : scope} {mx : mode} {A t u: term} : - (Γ⊨lam mx A t↣ u ) → md Γ (lam mx A t) ≠ ℙ → - ( ∃ A' t', u = lam mx A' t' ∧ Γ⊨A↣A' ∧ mx :: Γ⊨t↣t'). + (Γ⊨lam mx A t⇶ u ) → md Γ (lam mx A t) ≠ ℙ → + ( ∃ A' t', u = lam mx A' t' ∧ Γ⊨A⇶A' ∧ mx :: Γ⊨t⇶t'). Proof. intros red_lam not_Prop. remember (lam mx A t) as lam_t eqn:e0. @@ -137,8 +80,8 @@ Qed. Lemma red_Pi_inv {Γ : scope} {i j: level} {m mx : mode} {A B t: term} : - Γ⊨Pi i j m mx A B↣ t → - (∃ A' B' i' j', t = Pi i' j' m mx A' B' ∧ Γ ⊨ A ↣ A' ∧ mx::Γ ⊨ B ↣ B'). + Γ⊨Pi i j m mx A B⇶ t → + (∃ A' B' i' j', t = Pi i' j' m mx A' B' ∧ Γ ⊨ A ⇶ A' ∧ mx::Γ ⊨ B ⇶ B'). Proof. intro red_Pi. inversion red_Pi; subst. @@ -150,7 +93,7 @@ Proof. Qed. Lemma red_Sort_inv {Γ: scope} {i: level} {m: mode} {t: term} : - Γ ⊨ Sort m i ↣ t → ∃ i', t = Sort m i'. + Γ ⊨ Sort m i ⇶ t → ∃ i', t = Sort m i'. Proof. intro red_sort. inversion red_sort. @@ -160,8 +103,30 @@ Proof. match goal with | HC : 𝕂 = ℙ |- _ => inversion HC end. Qed. +Lemma red_Erased_inv {Γ: scope} {t0 t': term} : + Γ ⊨ Erased t0 ⇶ t' → ∃ t0', t' = Erased t0' ∧ Γ ⊨ t0 ⇶ t0'. +Proof. + intro red1. + inversion red1. + - eauto. + - eexists; split; [reflexivity | gred]. + - cbn in *. + match goal with | HC : 𝕂 = ℙ |- _ => inversion HC end. +Qed. + +Lemma red_vec_inv {Γ: scope} {A0 n0 t': term} : + Γ ⊨ tvec A0 n0 ⇶ t' → ∃ A1 n1, t' = tvec A1 n1 ∧ Γ ⊨ A0 ⇶ A1 ∧ Γ ⊨ n0 ⇶ n1. +Proof. + intro red1. + inversion red1. + - eauto. + - repeat eexists; gred. + - cbn in *. + match goal with | HC : 𝕂 = ℙ |- _ => inversion HC end. +Qed. + Lemma red_hide_inv {Γ : scope} {t0 t' : term} : - Γ⊨hide t0 ↣t' → ∃ t0', t' = hide t0' ∧ Γ ⊨ t0 ↣ t0'. + Γ⊨hide t0 ⇶t' → ∃ t0', t' = hide t0' ∧ Γ ⊨ t0 ⇶ t0'. Proof. intro red_hide. inversion red_hide; subst. @@ -171,7 +136,7 @@ Proof. match goal with | HC : 𝔾 = ℙ |- _ => inversion HC end. Qed. -Lemma red_succ_inv (Γ : scope) (n t' : term) (red_succ : Γ⊨tsucc n ↣t' ) : ∃ n', t' = tsucc n' ∧ Γ ⊨ n ↣ n'. +Lemma red_succ_inv (Γ : scope) (n t' : term) (red_succ : Γ⊨tsucc n ⇶t' ) : ∃ n', t' = tsucc n' ∧ Γ ⊨ n ⇶ n'. Proof. inversion red_succ; subst. - eauto. @@ -180,8 +145,8 @@ Proof. match goal with | HC : 𝕋 = ℙ |- _ => inversion HC end. Qed. -Lemma red_nil_inv (Γ : scope) (A t' : term) (red_nil : Γ ⊨ tvnil A ↣ t' ) : - ∃ A', t' = tvnil A' ∧ Γ ⊨ A ↣ A'. +Lemma red_nil_inv (Γ : scope) (A t' : term) (red_nil : Γ ⊨ tvnil A ⇶ t' ) : + ∃ A', t' = tvnil A' ∧ Γ ⊨ A ⇶ A'. Proof. inversion red_nil; subst. - eauto. @@ -190,8 +155,8 @@ Proof. match goal with | HC : 𝕋 = ℙ |- _ => inversion HC end. Qed. -Lemma red_cons_inv (Γ : scope) (a n v t' : term) (red_cons : Γ ⊨ tvcons a n v ↣ t' ) : - ∃ a' n' v', t' = tvcons a' n' v' ∧ Γ ⊨ a ↣ a' ∧ Γ ⊨ n ↣ n' ∧ Γ ⊨ v ↣ v'. +Lemma red_cons_inv (Γ : scope) (a n v t' : term) (red_cons : Γ ⊨ tvcons a n v ⇶ t' ) : + ∃ a' n' v', t' = tvcons a' n' v' ∧ Γ ⊨ a ⇶ a' ∧ Γ ⊨ n ⇶ n' ∧ Γ ⊨ v ⇶ v'. Proof. inversion red_cons; subst. - do 3 eexists; eauto. @@ -202,7 +167,7 @@ Qed. Lemma red_ren (Γ Δ : scope) (ρ: nat → nat) (t t': term) : (∀ n, nth (ρ n) Γ 𝕋 = nth n Δ 𝕋) → - Δ ⊨ t ↣ t' → Γ ⊨ ρ ⋅ t ↣ ρ ⋅ t'. + Δ ⊨ t ⇶ t' → Γ ⊨ ρ ⋅ t ⇶ ρ ⋅ t'. Proof. intros Hscope red_t. induction red_t in Γ, Δ, ρ, Hscope, t, t', red_t |- *. @@ -224,8 +189,8 @@ Proof. Qed. Lemma up_subst_red (Γ : scope) (m : mode) (σ σ' : nat → term) : - (∀ n, Γ ⊨ σ n ↣ σ' n) → - (∀ n, m::Γ ⊨ up_term σ n ↣ up_term σ' n). + (∀ n, Γ ⊨ σ n ⇶ σ' n) → + (∀ n, m::Γ ⊨ up_term σ n ⇶ up_term σ' n). Proof. intros Hyp n. destruct n. @@ -236,8 +201,8 @@ Proof. Qed. Lemma red_subst_r (Γ : scope) (t : term) (σ σ' : nat → term) : - (∀ n, Γ ⊨ σ n ↣ σ' n) → - Γ ⊨ t <[σ] ↣ t <[σ']. + (∀ n, Γ ⊨ σ n ⇶ σ' n) → + Γ ⊨ t <[σ] ⇶ t <[σ']. Proof. intro red_σ. induction t in Γ, σ, σ', red_σ |- *. @@ -247,9 +212,9 @@ Qed. Lemma red_subst (Γ Δ : scope) (t t' : term) (σ σ' : nat → term) : (∀ n, md Γ (σ n) = nth n Δ 𝕋) → - (∀ n, Γ ⊨ σ n ↣ σ' n) → - Δ ⊨ t ↣ t' → - Γ ⊨ t <[σ] ↣ t' <[σ']. + (∀ n, Γ ⊨ σ n ⇶ σ' n) → + Δ ⊨ t ⇶ t' → + Γ ⊨ t <[σ] ⇶ t' <[σ']. Proof. intros Hscope red_σ red_t. remember Δ as Δ0 eqn:e. @@ -276,7 +241,7 @@ Qed. Ltac red_lam_inv_auto A' t' e red_A' red_t':= match goal with - | red_lam : ?Γ⊨lam ?m ?A ?t ↣?u |- _ => + | red_lam : ?Γ⊨lam ?m ?A ?t ⇶?u |- _ => eapply red_lam_inv in red_lam; eauto; destruct red_lam as [A' [t' [e [red_A' red_t']]]]; try subst u @@ -284,7 +249,7 @@ Ltac red_lam_inv_auto A' t' e red_A' red_t':= Ltac red_Pi_inv_auto A' B' i' j' e red_A' red_B':= match goal with - | red_Pi : ?Γ⊨Pi ?i ?j ?m ?mx ?A ?B ↣?u |- _ => + | red_Pi : ?Γ⊨Pi ?i ?j ?m ?mx ?A ?B ⇶?u |- _ => eapply red_Pi_inv in red_Pi; eauto; destruct red_lam as [A' [B' [i' [j' [e [red_A' red_B']]]]]]; try subst u @@ -292,7 +257,7 @@ Ltac red_Pi_inv_auto A' B' i' j' e red_A' red_B':= Ltac red_hide_inv_auto t0' e:= match goal with - | red_hide : ?Γ⊨hide ?t0 ↣?t' |- _ => + | red_hide : ?Γ⊨hide ?t0 ⇶?t' |- _ => apply red_hide_inv in red_hide; destruct red_hide as [t0' [e red_hide]]; try subst t' @@ -300,7 +265,7 @@ Ltac red_hide_inv_auto t0' e:= Ltac red_succ_inv_auto n' e:= match goal with - | red_succ : ?Γ⊨tsucc ?t0 ↣?t' |- _ => + | red_succ : ?Γ⊨tsucc ?t0 ⇶?t' |- _ => apply red_succ_inv in red_succ; destruct red_succ as [n' [e red_succ]]; try subst t' @@ -308,7 +273,7 @@ Ltac red_succ_inv_auto n' e:= Ltac red_nil_inv_auto A' e:= match goal with - | red_nil : ?Γ⊨tvnil ?A ↣?t' |- _ => + | red_nil : ?Γ⊨tvnil ?A ⇶?t' |- _ => apply red_nil_inv in red_nil; destruct red_nil as [A' [e red_nil]]; try subst t' @@ -316,7 +281,7 @@ Ltac red_nil_inv_auto A' e:= Ltac red_conv_inv_auto a' n' v' e red_a' red_n' red_v':= match goal with - | red_cons : ?Γ⊨tvcons ?a ?n ?v ↣?t' |- _ => + | red_cons : ?Γ⊨tvcons ?a ?n ?v ⇶?t' |- _ => apply red_cons_inv in red_cons; destruct red_cons as [a' [n' [v' [e [red_a' [red_n' red_v']]]]]]; try subst t' @@ -325,11 +290,11 @@ Ltac red_conv_inv_auto a' n' v' e red_a' red_n' red_v':= Ltac red_basic_inv := let e := fresh "e" in match goal with - | H : ?Γ ⊨ tzero ↣ ?t |- _ => + | H : ?Γ ⊨ tzero ⇶ ?t |- _ => inversion H - | H : ?Γ ⊨ ttrue ↣ ?t |- _ => + | H : ?Γ ⊨ ttrue ⇶ ?t |- _ => inversion H - | H : ?Γ ⊨ tfalse ↣ ?t |- _ => + | H : ?Γ ⊨ tfalse ⇶ ?t |- _ => inversion H end; subst. diff --git a/theories/reduction/multisteps/Reduction.v b/theories/reduction/multisteps/Reduction.v index f851217..ccc5e80 100644 --- a/theories/reduction/multisteps/Reduction.v +++ b/theories/reduction/multisteps/Reduction.v @@ -3,274 +3,235 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl TermMode Scoping BasicMetaTheory. From GhostTT Require Export Univ TermNotations Typing. -From GhostTT.reduction Require Export Notations. +From GhostTT.reduction Require Export Notations Utils. Import ListNotations. Set Default Goal Selector "!". (* ------------------------------------------------------------------------- *) -Section Properties_Star. - - Lemma type_star : ∀ Γ, Γ ⊢ ⋆ : ⊤. - Proof. - intro Γ. - apply type_lam. - all: eauto using scope_bot, type_bot. - - apply scope_var; reflexivity. - - eapply type_var; reflexivity. - Qed. - - Lemma scope_star : ∀ Γ, Γ ⊢ ⋆ ∷ ℙ. - Proof. - intro. - apply scope_lam. - - apply scope_bot. - - apply scope_var; reflexivity. - Qed. - -End Properties_Star. - -(* ------------------------------------------------------------------------- *) -Section Rewriting. - (** rewriting rules **) - - Definition mode_eq : ∀ (x y : mode), { x = y } + { x ≠ y }. - Proof. - decide equality. - Defined. - - Definition red_lvl (m : mode) (i : level) : level := - if mode_eq m ℙ then 0 else i. - - - Inductive reduction (Γ : scope) : term → term → Prop := - - (* Computation rules *) - - | red_beta : - ∀ mx A t t' u u', - (mx::Γ) ⊨ t ↣ t' → - md Γ u = mx → - Γ ⊨ u ↣ u' → - Γ ⊨ app (lam mx A t) u ↣ t' <[ u' ·· ] - - | red_reveal_hide : - ∀ t P p t' p', - Γ ⊨ t ↣ t' → - Γ ⊨ p ↣ p' → - In (md Γ p) [ℙ;𝔾] → - Γ ⊨ reveal (hide t) P p ↣ app p' t' - - | red_if_true : - ∀ m P t f t', - Γ ⊨ t ↣ t' → - md Γ t = m → - Γ ⊨ tif m ttrue P t f ↣ t' - - | red_if_false : - ∀ m P t f f', - Γ ⊨ f ↣ f' → - md Γ f = m → - Γ ⊨ tif m tfalse P t f ↣ f' - - | red_nat_elim_zero : - ∀ m P z s z', - Γ ⊨ z ↣ z' → - md Γ z = m → - Γ ⊨ tnat_elim m tzero P z s ↣ z' - - | red_nat_elim_succ : - ∀ m P z s n P' z' s' n', - Γ ⊨ n ↣ n' → - Γ ⊨ P ↣ P' → - Γ ⊨ z ↣ z' → - Γ ⊨ s ↣ s' → - md Γ s = m → - Γ ⊨ tnat_elim m (tsucc n) P z s ↣ app (app s' n') (tnat_elim m n' P' z' s') - - | red_vec_elim_nil : - ∀ m A B P z s z', - Γ ⊨ z ↣ z' → - md Γ z = m → - Γ ⊨ tvec_elim m A (hide tzero) (tvnil B) P z s ↣ z' - - | red_vec_elim_cons : - ∀ m A a n n0 v P z s A' a' n' v' P' z' s', - Γ ⊨ A ↣ A' → - Γ ⊨ a ↣ a' → - Γ ⊨ n ↣ n' → - Γ ⊨ v ↣ v' → - Γ ⊨ P ↣ P' → - Γ ⊨ z ↣ z' → - Γ ⊨ s ↣ s' → - md Γ s = m → - Γ ⊨ tvec_elim m A n0 (tvcons a n v) P z s ↣ - app (app (app (app s' a') (glength A' n' v')) v') (tvec_elim m A' n' v' P' z' s') - - (* Congruence rules *) - - (* A rule to quotient away all levels of Prop, making it impredicatime *) - | red_Prop : - ∀ i, Γ ⊨ Sort ℙ i ↣ Sort ℙ 0 - - | red_Pi : - ∀ i j m mx A A' B B', - Γ ⊨ A ↣ A' → - (mx::Γ) ⊨ B ↣ B' → - Γ ⊨ Pi i j m mx A B ↣ Pi (red_lvl mx i) (red_lvl m j) m mx A' B' - - | red_Pi' : (* needed for red_subst *) - ∀ i j m mx A A' B B', - Γ ⊨ A ↣ A' → - (mx::Γ) ⊨ B ↣ B' → - Γ ⊨ Pi i j m mx A B ↣ Pi i j m mx A' B' - - | red_lam : - ∀ mx A A' t t', - Γ ⊨ A ↣ A' → - (mx::Γ) ⊨ t ↣ t' → - Γ ⊨ lam mx A t ↣ lam mx A' t' - - | red_app : - ∀ u u' v v', - Γ ⊨ u ↣ u' → - Γ ⊨ v ↣ v' → - Γ ⊨ app u v ↣ app u' v' - - | red_Erased : - ∀ A A', - Γ ⊨ A ↣ A' → - Γ ⊨ Erased A ↣ Erased A' - - | red_hide : - ∀ u u', - Γ ⊨ u ↣ u' → - Γ ⊨ hide u ↣ hide u' - - | red_reveal : - ∀ t t' P P' p p', - Γ ⊨ t ↣ t' → - Γ ⊨ P ↣ P' → - Γ ⊨ p ↣ p' → - Γ ⊨ reveal t P p ↣ reveal t' P' p' - - | red_Reveal : - ∀ t t' p p', - Γ ⊨ t ↣ t' → - Γ ⊨ p ↣ p' → - Γ ⊨ Reveal t p ↣ Reveal t' p' - - | red_gheq : - ∀ A A' u u' v v', - Γ ⊨ A ↣ A' → - Γ ⊨ u ↣ u' → - Γ ⊨ v ↣ v' → - Γ ⊨ gheq A u v ↣ gheq A' u' v' - - | red_if : - ∀ m b b' P P' t t' f f', - Γ ⊨ b ↣ b' → - Γ ⊨ P ↣ P' → - Γ ⊨ t ↣ t' → - Γ ⊨ f ↣ f' → - Γ ⊨ tif m b P t f ↣ tif m b' P' t' f' - - | red_succ : - ∀ n n', - Γ ⊨ n ↣ n' → - Γ ⊨ tsucc n ↣ tsucc n' - - | red_nat_elim : - ∀ m n n' P P' z z' s s', - Γ ⊨ n ↣ n' → - Γ ⊨ P ↣ P' → - Γ ⊨ z ↣ z' → - Γ ⊨ s ↣ s' → - Γ ⊨ tnat_elim m n P z s ↣ tnat_elim m n' P' z' s' - - | red_vec : - ∀ A A' n n', - Γ ⊨ A ↣ A' → - Γ ⊨ n ↣ n' → - Γ ⊨ tvec A n ↣ tvec A' n' - - | red_vnil : - ∀ A A', - Γ ⊨ A ↣ A' → - Γ ⊨ tvnil A ↣ tvnil A' - - | red_vcons : - ∀ a a' n n' v v', - Γ ⊨ a ↣ a' → - Γ ⊨ n ↣ n' → - Γ ⊨ v ↣ v' → - Γ ⊨ tvcons a n v ↣ tvcons a' n' v' - - | red_vec_elim : - ∀ m A A' n n' v v' P P' z z' s s', - Γ ⊨ A ↣ A' → - Γ ⊨ n ↣ n' → - Γ ⊨ v ↣ v' → - Γ ⊨ P ↣ P' → - Γ ⊨ z ↣ z' → - Γ ⊨ s ↣ s' → - Γ ⊨ tvec_elim m A n v P z s ↣ tvec_elim m A' n' v' P' z' s' - - | red_bot_elim : - ∀ m A A' p p', - Γ ⊨ A ↣ A' → - Γ ⊨ p ↣ p' → - Γ ⊨ bot_elim m A p ↣ bot_elim m A' p' - - (* Structural rule *) - - | red_refl : - ∀ u, - Γ ⊨ u ↣ u - - (* Proof irrelevance *) - - | red_irr : - ∀ p, - md Γ p = ℙ → - Γ ⊨ p ↣ ⋆ - - | red_toRev : - ∀ t t' p p' u u', - Γ ⊨ t ↣ t' → - Γ ⊨ p ↣ p' → - Γ ⊨ u ↣ u' → - Γ ⊨ toRev t p u ↣ toRev t' p' u' - - | red_fromRev : - ∀ t t' p p' u u', - Γ ⊨ t ↣ t' → - Γ ⊨ p ↣ p' → - Γ ⊨ u ↣ u' → - Γ ⊨ fromRev t p u ↣ fromRev t' p' u' - - | red_ghrefl : - ∀ A A' u u', - Γ ⊨ A ↣ A' → - Γ ⊨ u ↣ u' → - Γ ⊨ ghrefl A u ↣ ghrefl A' u' - - | red_ghcast : - ∀ A A' u u' v v' e e' P P' t t', - Γ ⊨ A ↣ A' → - Γ ⊨ u ↣ u' → - Γ ⊨ v ↣ v' → - Γ ⊨ e ↣ e' → - Γ ⊨ P ↣ P' → - Γ ⊨ t ↣ t' → - Γ ⊨ ghcast A u v e P t ↣ ghcast A' u' v' e' P' t' - - - where "Γ ⊨ u ↣ v" := (reduction Γ u v). - - -End Rewriting. -Notation "Γ ⊨ u ↣ v" := (reduction Γ u v). +Inductive reduction (Γ : scope) : term → term → Prop := + + (* Computation rules *) + + | red_beta : + ∀ mx A t t' u u', + (mx::Γ) ⊨ t ⇶ t' → + md Γ u = mx → + Γ ⊨ u ⇶ u' → + Γ ⊨ app (lam mx A t) u ⇶ t' <[ u' ·· ] + + | red_reveal_hide : + ∀ t P p t' p', + Γ ⊨ t ⇶ t' → + Γ ⊨ p ⇶ p' → + In (md Γ p) [ℙ;𝔾] → + Γ ⊨ reveal (hide t) P p ⇶ app p' t' + + | red_if_true : + ∀ m P t f t', + Γ ⊨ t ⇶ t' → + md Γ t = m → + Γ ⊨ tif m ttrue P t f ⇶ t' + + | red_if_false : + ∀ m P t f f', + Γ ⊨ f ⇶ f' → + md Γ f = m → + Γ ⊨ tif m tfalse P t f ⇶ f' + + | red_nat_elim_zero : + ∀ m P z s z', + Γ ⊨ z ⇶ z' → + md Γ z = m → + Γ ⊨ tnat_elim m tzero P z s ⇶ z' + + | red_nat_elim_succ : + ∀ m P z s n P' z' s' n', + Γ ⊨ n ⇶ n' → + Γ ⊨ P ⇶ P' → + Γ ⊨ z ⇶ z' → + Γ ⊨ s ⇶ s' → + md Γ s = m → + Γ ⊨ tnat_elim m (tsucc n) P z s ⇶ app (app s' n') (tnat_elim m n' P' z' s') + + | red_vec_elim_nil : + ∀ m A B P z s z', + Γ ⊨ z ⇶ z' → + md Γ z = m → + Γ ⊨ tvec_elim m A (hide tzero) (tvnil B) P z s ⇶ z' + + | red_vec_elim_cons : + ∀ m A a n n0 v P z s A' a' n' v' P' z' s', + Γ ⊨ A ⇶ A' → + Γ ⊨ a ⇶ a' → + Γ ⊨ n ⇶ n' → + Γ ⊨ v ⇶ v' → + Γ ⊨ P ⇶ P' → + Γ ⊨ z ⇶ z' → + Γ ⊨ s ⇶ s' → + md Γ s = m → + Γ ⊨ tvec_elim m A n0 (tvcons a n v) P z s ⇶ + app (app (app (app s' a') (glength A' n' v')) v') (tvec_elim m A' n' v' P' z' s') + + + (* A rule to quotient away all levels of Prop, making it impredicatime *) + | red_Prop : + ∀ i, Γ ⊨ Sort ℙ i ⇶ Sort ℙ 0 + + | red_Pi : + ∀ i j m mx A A' B B', + Γ ⊨ A ⇶ A' → + (mx::Γ) ⊨ B ⇶ B' → + Γ ⊨ Pi i j m mx A B ⇶ Pi (red_lvl mx i) (red_lvl m j) m mx A' B' + + (* Congruence rules *) + + | red_Pi' : (* needed for red_subst *) + ∀ i j m mx A A' B B', + Γ ⊨ A ⇶ A' → + (mx::Γ) ⊨ B ⇶ B' → + Γ ⊨ Pi i j m mx A B ⇶ Pi i j m mx A' B' + + | red_lam : + ∀ mx A A' t t', + Γ ⊨ A ⇶ A' → + (mx::Γ) ⊨ t ⇶ t' → + Γ ⊨ lam mx A t ⇶ lam mx A' t' + + | red_app : + ∀ u u' v v', + Γ ⊨ u ⇶ u' → + Γ ⊨ v ⇶ v' → + Γ ⊨ app u v ⇶ app u' v' + + | red_Erased : + ∀ A A', + Γ ⊨ A ⇶ A' → + Γ ⊨ Erased A ⇶ Erased A' + + | red_hide : + ∀ u u', + Γ ⊨ u ⇶ u' → + Γ ⊨ hide u ⇶ hide u' + + | red_reveal : + ∀ t t' P P' p p', + Γ ⊨ t ⇶ t' → + Γ ⊨ P ⇶ P' → + Γ ⊨ p ⇶ p' → + Γ ⊨ reveal t P p ⇶ reveal t' P' p' + + | red_Reveal : + ∀ t t' p p', + Γ ⊨ t ⇶ t' → + Γ ⊨ p ⇶ p' → + Γ ⊨ Reveal t p ⇶ Reveal t' p' + + | red_gheq : + ∀ A A' u u' v v', + Γ ⊨ A ⇶ A' → + Γ ⊨ u ⇶ u' → + Γ ⊨ v ⇶ v' → + Γ ⊨ gheq A u v ⇶ gheq A' u' v' + + | red_if : + ∀ m b b' P P' t t' f f', + Γ ⊨ b ⇶ b' → + Γ ⊨ P ⇶ P' → + Γ ⊨ t ⇶ t' → + Γ ⊨ f ⇶ f' → + Γ ⊨ tif m b P t f ⇶ tif m b' P' t' f' + + | red_succ : + ∀ n n', + Γ ⊨ n ⇶ n' → + Γ ⊨ tsucc n ⇶ tsucc n' + + | red_nat_elim : + ∀ m n n' P P' z z' s s', + Γ ⊨ n ⇶ n' → + Γ ⊨ P ⇶ P' → + Γ ⊨ z ⇶ z' → + Γ ⊨ s ⇶ s' → + Γ ⊨ tnat_elim m n P z s ⇶ tnat_elim m n' P' z' s' + + | red_vec : + ∀ A A' n n', + Γ ⊨ A ⇶ A' → + Γ ⊨ n ⇶ n' → + Γ ⊨ tvec A n ⇶ tvec A' n' + + | red_vnil : + ∀ A A', + Γ ⊨ A ⇶ A' → + Γ ⊨ tvnil A ⇶ tvnil A' + + | red_vcons : + ∀ a a' n n' v v', + Γ ⊨ a ⇶ a' → + Γ ⊨ n ⇶ n' → + Γ ⊨ v ⇶ v' → + Γ ⊨ tvcons a n v ⇶ tvcons a' n' v' + + | red_vec_elim : + ∀ m A A' n n' v v' P P' z z' s s', + Γ ⊨ A ⇶ A' → + Γ ⊨ n ⇶ n' → + Γ ⊨ v ⇶ v' → + Γ ⊨ P ⇶ P' → + Γ ⊨ z ⇶ z' → + Γ ⊨ s ⇶ s' → + Γ ⊨ tvec_elim m A n v P z s ⇶ tvec_elim m A' n' v' P' z' s' + + | red_bot_elim : + ∀ m A A' p p', + Γ ⊨ A ⇶ A' → + Γ ⊨ p ⇶ p' → + Γ ⊨ bot_elim m A p ⇶ bot_elim m A' p' + + | red_toRev : + ∀ t t' p p' u u', + Γ ⊨ t ⇶ t' → + Γ ⊨ p ⇶ p' → + Γ ⊨ u ⇶ u' → + Γ ⊨ toRev t p u ⇶ toRev t' p' u' + + | red_fromRev : + ∀ t t' p p' u u', + Γ ⊨ t ⇶ t' → + Γ ⊨ p ⇶ p' → + Γ ⊨ u ⇶ u' → + Γ ⊨ fromRev t p u ⇶ fromRev t' p' u' + + | red_ghrefl : + ∀ A A' u u', + Γ ⊨ A ⇶ A' → + Γ ⊨ u ⇶ u' → + Γ ⊨ ghrefl A u ⇶ ghrefl A' u' + + | red_ghcast : + ∀ A A' u u' v v' e e' P P' t t', + Γ ⊨ A ⇶ A' → + Γ ⊨ u ⇶ u' → + Γ ⊨ v ⇶ v' → + Γ ⊨ e ⇶ e' → + Γ ⊨ P ⇶ P' → + Γ ⊨ t ⇶ t' → + Γ ⊨ ghcast A u v e P t ⇶ ghcast A' u' v' e' P' t' + + (* Structural rule *) + | red_refl : + ∀ u, + Γ ⊨ u ⇶ u + + (* Proof irrelevance *) + | red_irr : + ∀ p, + md Γ p = ℙ → + Γ ⊨ p ⇶ ⋆ + + + where "Γ ⊨ u ⇶ v" := (reduction Γ u v). (* ------------------------------------------------------------------------- *) (* multi-step rewriting automation *) @@ -286,4 +247,5 @@ Hint Resolve red_beta red_reveal_hide red_if_true red_if_false red_nat_elim_zero Ltac gred := unshelve typeclasses eauto with gtt_scope gtt_red shelvedb ; shelve_unifiable. -(** end multi-step rewriting automation **) + +(*** end of file ***) diff --git a/theories/reduction/multisteps/ReductionToCongruence.v b/theories/reduction/multisteps/ReductionToCongruence.v index fd7b163..ffb21a9 100644 --- a/theories/reduction/multisteps/ReductionToCongruence.v +++ b/theories/reduction/multisteps/ReductionToCongruence.v @@ -7,13 +7,13 @@ From GhostTT.reduction.multisteps Require Export Reduction Transitivity. Import ListNotations. -Notation "Γ ⊨ u ε↣ v" := (Γ ⊨ ε|u| ↣ ε|v|). +Notation "Γ ⊨ u ε⇶ v" := (Γ ⊨ ε|u| ⇶ ε|v|). Set Default Goal Selector "!". Lemma conv_subst_r (Γ : context) (Δ : scope) (m : mode) (σ σ' : nat → term) (t : term) : sscoping (sc Γ) σ Δ → sscoping (sc Γ) σ' Δ → - Δ ⊢ ε|t|∷m → (∀ n, Γ ⊢ σ n ≡ σ' n) → Γ ⊢ ε|t| <[σ] ≡ ε|t| <[σ']. + Δ ⊨ ε|t|∷m → (∀ n, Γ ⊢ σ n ≡ σ' n) → Γ ⊢ ε|t| <[σ] ≡ ε|t| <[σ']. Proof. intros Hscope Hscope' scope_t conv_σ. induction t in Γ, Δ, Hscope, Hscope', m, scope_t, σ, σ', conv_σ |- *. @@ -27,15 +27,15 @@ Qed. Theorem reduction_to_conversion : - ∀ Γ m t t', (sc Γ) ⊢ t∷m → (sc Γ) ⊨ t ↣ t' → Γ ⊢ t ε≡ t'. + ∀ Γ m t t', Γ ⊢ t∷m → (sc Γ) ⊨ t ⇶ t' → Γ ⊢ t ε≡ t'. Proof. intros Γ m t t' scope_t red_t. remember (sc Γ) as Γ0 eqn:e. induction red_t in Γ, Γ0, e, t, m, scope_t, t', red_t |- *. - 28:{ apply conv_irr; [| apply scope_star]. + 32:{ apply conv_irr; [| apply scope_star]. apply scoping_castrm. erewrite scoping_md in *; eauto. congruence. } - 9, 27: solve [gconv]. + 9, 31: solve [gconv]. all: inversion scope_t; subst. 26-28: solve [apply conv_irr; apply scoping_castrm; subst; gscope; eauto using red_scope]. 11-26: solve [gconv; cbn; f_equal; assumption]. @@ -76,7 +76,7 @@ Proof. Qed. Corollary reductions_to_conversion : - ∀ Γ m t t', (sc Γ) ⊢ t∷m → (sc Γ) ⊨ t ↣* t' → Γ ⊢ t ε≡ t'. + ∀ Γ m t t', Γ ⊢ t∷m → (sc Γ) ⊨ t ⇶* t' → Γ ⊢ t ε≡ t'. Proof. intros Γ m t t' scope_t red_t. induction red_t. @@ -90,24 +90,24 @@ Qed. Local Ltac conversion_to_reduction_exists := match goal with | H : ∃ _, _ ∧ _ |- - ∃ w, ?Γ ⊨ ?c _ ↣* w ∧ ?Γ ⊨ ?c _ ↣* w => + ∃ w, ?Γ ⊨ ?c _ ⇶* w ∧ ?Γ ⊨ ?c _ ⇶* w => let w := fresh "w" in destruct H as [ w [ ]]; exists (c w) | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _ |- - ∃ w, ?Γ ⊨ ?c _ _ ↣* w ∧ ?Γ ⊨ ?c _ _↣* w => + ∃ w, ?Γ ⊨ ?c _ _ ⇶* w ∧ ?Γ ⊨ ?c _ _⇶* w => let w0 := fresh "w0" in let w1 := fresh "w1" in destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; exists (c w0 w1) | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _, H2 : ∃ _, _ ∧ _ |- - ∃ w, ?Γ ⊨ ?c _ _ _ ↣* w ∧ ?Γ ⊨ ?c _ _ _ ↣* w => + ∃ w, ?Γ ⊨ ?c _ _ _ ⇶* w ∧ ?Γ ⊨ ?c _ _ _ ⇶* w => let w0 := fresh "w0" in let w1 := fresh "w1" in let w2 := fresh "w2" in destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; destruct H2 as [ w2 [ ]]; exists (c w0 w1 w2) | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _, H2 : ∃ _, _ ∧ _, H3 : ∃ _, _ ∧ _ |- - ∃ w, ?Γ ⊨ ?c _ _ _ _ ↣* w ∧ ?Γ ⊨ ?c _ _ _ _ ↣* w => + ∃ w, ?Γ ⊨ ?c _ _ _ _ ⇶* w ∧ ?Γ ⊨ ?c _ _ _ _ ⇶* w => let w0 := fresh "w0" in let w1 := fresh "w1" in let w2 := fresh "w2" in let w3 := fresh "w3" in destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]]; @@ -115,7 +115,7 @@ Local Ltac conversion_to_reduction_exists := exists (c w0 w1 w2 w3) | H0 : ∃ _, _ ∧ _, H1 : ∃ _, _ ∧ _, H2 : ∃ _, _ ∧ _, H3 : ∃ _, _ ∧ _ , H4 : ∃ _, _ ∧ _, H5 : ∃ _, _ ∧ _ |- - ∃ w, ?Γ ⊨ tvec_elim ?m _ _ _ _ _ _ ↣* w ∧ ?Γ ⊨ tvec_elim ?m _ _ _ _ _ _↣* w => + ∃ w, ?Γ ⊨ tvec_elim ?m _ _ _ _ _ _ ⇶* w ∧ ?Γ ⊨ tvec_elim ?m _ _ _ _ _ _⇶* w => let w0 := fresh "w0" in let w1 := fresh "w1" in let w2 := fresh "w2" in let w3 := fresh "w3" in let w4 := fresh "w4" in let w5 := fresh "w5" in @@ -126,13 +126,13 @@ Local Ltac conversion_to_reduction_exists := end. Theorem conversion_to_reduction : - ∀ Γ t t', Γ ⊢ t ≡ t' → ∃ u, (sc Γ) ⊨ t ↣* u ∧ (sc Γ) ⊨ t' ↣* u. + ∀ Γ t t', Γ ⊢ t ≡ t' → ∃ u, (sc Γ) ⊨ t ⇶* u ∧ (sc Γ) ⊨ t' ⇶* u. Proof. intros Γ t t' conv_t. induction conv_t in conv_t |- *. - all: try solve [match goal with | |- ∃ _, _ ∧ (_ ⊨ ?c ↣* _) => exists c end; split; apply red_trans_direct; gred; eauto using scoping_md]. + all: try solve [match goal with | |- ∃ _, _ ∧ (_ ⊨ ?c ⇶* _) => exists c end; split; apply red_trans_direct; gred; eauto using scoping_md]. all: try solve [conversion_to_reduction_exists; split; greds]. - - match goal with | |- ∃ _, _ ∧ (_ ⊨ ?c ↣* _) => exists c end; split; greds. + - match goal with | |- ∃ _, _ ∧ (_ ⊨ ?c ⇶* _) => exists c end; split; greds. all: try (apply red_trans_direct; gred). erewrite scoping_md; eauto. - exists (Sort ℙ 0); split; greds. @@ -153,7 +153,7 @@ Proof. destruct H0 as [ w0 [ ]]; destruct H1 as [ w1 [ ]] end. - assert (∃ w: term, (sc Γ⊨w0↣*w) ∧ sc Γ⊨w1↣*w) as Hw. + assert (∃ w: term, (sc Γ⊨w0⇶*w) ∧ sc Γ⊨w1⇶*w) as Hw. { eauto using reduction_confluence. } destruct Hw as [wf []]. exists wf. diff --git a/theories/reduction/multisteps/Rho.v b/theories/reduction/multisteps/Rho.v index 9d73f13..be59b0c 100644 --- a/theories/reduction/multisteps/Rho.v +++ b/theories/reduction/multisteps/Rho.v @@ -264,7 +264,7 @@ Proof. all: case (mode_eq _ _); [reflexivity |contradiction ]. Qed. -Theorem rho_one_step : ∀ Γ t, Γ ⊨ t ↣ ρ Γ t. +Theorem rho_one_step : ∀ Γ t, Γ ⊨ t ⇶ ρ Γ t. Proof. intros Γ t. induction t in Γ |- *. @@ -274,7 +274,7 @@ Proof. - unfold red_lvl. case (mode_eq _ _) as [ e | ne ]; subst; gred. - case term_view_app as [m A t | ]; simp ρ. - 1: assert ( Γ⊨lam m A t↣ρ Γ (lam m A t)) as red_lam; eauto. + 1: assert ( Γ⊨lam m A t⇶ρ Γ (lam m A t)) as red_lam; eauto. 1: simp ρ in red_lam. all: case (mode_eq _ _) as [ e | ne ]; gred. all: case (mode_eq _ _) as [ e' | ne' ]; gred. @@ -283,7 +283,7 @@ Proof. assumption. - case term_view_reveal as [t P p | ]; simp ρ. all: case (mode_eq _ _) as [ e | ne ]; gred. - * assert (Γ⊨hide t↣ρ Γ (hide t)) as H; eauto. + * assert (Γ⊨hide t⇶ρ Γ (hide t)) as H; eauto. simp ρ in H. red_hide_inv_auto t' e. noconf e. @@ -299,7 +299,7 @@ Proof. all: repeat case (mode_eq _ _) as [ | ]; cbn. 3: eapply red_nat_elim; gred. all: gred. - assert (Γ⊨tsucc n↣ρ Γ (tsucc n)) as H; eauto. + assert (Γ⊨tsucc n⇶ρ Γ (tsucc n)) as H; eauto. simp ρ in H. red_succ_inv_auto n' e'. noconf e'. @@ -308,7 +308,7 @@ Proof. all:repeat case (mode_eq _ _) as [ | ]; cbn. 3: eapply red_vec_elim; gred. all: try solve [gred]. - assert (Γ⊨tvcons a n0 v↣ρ Γ (tvcons a n0 v)) as H; eauto. + assert (Γ⊨tvcons a n0 v⇶ρ Γ (tvcons a n0 v)) as H; eauto. simp ρ in H. red_conv_inv_auto a' n' v' e' red_a' red_n' red_v'. cbn; gred. @@ -316,7 +316,7 @@ Proof. Qed. Lemma reduction_triangle (Γ : scope) (t u : term) : - Γ ⊨ t ↣ u → Γ ⊨ u ↣ (ρ Γ t). + Γ ⊨ t ⇶ u → Γ ⊨ u ⇶ (ρ Γ t). Proof. intros red_t. induction red_t in t, Γ, u, red_t |- *. @@ -355,7 +355,7 @@ Proof. all: repeat case (mode_eq _ _) as [ | ]; gred. * cbn. erewrite <- red_md; eauto. cbn. assumption. * red_lam_inv_auto A' t' e' red_A red_u. - match goal with H: _ ⊨ lam _ _ _ ↣ _ |- _ => + match goal with H: _ ⊨ lam _ _ _ ⇶ _ |- _ => eapply red_lam_inv in H; eauto; [ destruct H as [A'' [t'' [ Hu'' [red_A' red_u']]]] | ] end. + apply sym_eq in Hu''; inversion Hu''; subst. diff --git a/theories/reduction/multisteps/Transitivity.v b/theories/reduction/multisteps/Transitivity.v index 90335f1..e994627 100644 --- a/theories/reduction/multisteps/Transitivity.v +++ b/theories/reduction/multisteps/Transitivity.v @@ -11,19 +11,19 @@ Set Default Goal Selector "!". (* Definition *) Inductive reduction_trans (Γ : scope) (u v: term) : Prop := | Refl: u = v → reduction_trans Γ u v - | Trans w : Γ ⊨ u ↣ w → reduction_trans Γ w v → reduction_trans Γ u v. + | Trans w : Γ ⊨ u ⇶ w → reduction_trans Γ w v → reduction_trans Γ u v. -Notation "Γ ⊨ u ↣* v" := (reduction_trans Γ u v). +Notation "Γ ⊨ u ⇶* v" := (reduction_trans Γ u v). (* Usefull properties *) -Lemma red_trans_direct {Γ : scope } {u v: term} : Γ ⊨ u ↣ v → Γ ⊨ u ↣* v. +Lemma red_trans_direct {Γ : scope } {u v: term} : Γ ⊨ u ⇶ v → Γ ⊨ u ⇶* v. Proof. refine ( λ H, Trans Γ u v v H (Refl Γ v v _)). reflexivity. Qed. Lemma red_trans_trans {Γ : scope} {u v: term} : - ∀ w, Γ ⊨ u ↣* w → Γ ⊨ w ↣* v → Γ ⊨ u ↣* v. + ∀ w, Γ ⊨ u ⇶* w → Γ ⊨ w ⇶* v → Γ ⊨ u ⇶* v. Proof. intros w red_u red_w. induction red_u as [ u | u w w' red_u red_w' IHw]. @@ -32,7 +32,7 @@ Proof. Qed. Corollary reds_scope (Γ : scope) (m: mode) (t t': term) : - Γ ⊨ t ↣* t' → Γ⊢t∷m → Γ⊢t'∷m. + Γ ⊨ t ⇶* t' → Γ⊨t∷m → Γ⊨t'∷m. Proof. intros reds_t scope_t. induction reds_t. @@ -49,7 +49,7 @@ Local Ltac end_things H:= eapply Trans; [ gred | eassumption]]. Lemma reds_beta (Γ : scope) (mx : mode) (A t t' u u' : term) : - mx :: Γ⊨t↣*t'→ md Γ u = mx → Γ⊨u↣*u' → Γ⊨app (lam mx A t) u↣*t' <[u'··]. + mx :: Γ⊨t⇶*t'→ md Γ u = mx → Γ⊨u⇶*u' → Γ⊨app (lam mx A t) u⇶*t' <[u'··]. Proof. intros red_t scope_u red_u. induction red_u. @@ -61,8 +61,8 @@ Proof. Qed. Lemma reds_reveal_hide (Γ : scope) (mp : mode) (t P p t' p' : term): - Γ⊨t↣*t' → Γ⊨p↣*p' → In (md Γ p) [ℙ;𝔾] → - Γ⊨reveal (hide t) P p↣*app p' t'. + Γ⊨t⇶*t' → Γ⊨p⇶*p' → In (md Γ p) [ℙ;𝔾] → + Γ⊨reveal (hide t) P p⇶*app p' t'. Proof. intros red_t red_p Hscope. induction red_t. @@ -81,13 +81,13 @@ Qed. (* Lemma reds_vec_elim_nil *) (* Lemma reds_vec_elim_cons *) -Lemma reds_Prop (Γ : scope) (i : level): Γ⊨Sort ℙ i↣*Sort ℙ 0. +Lemma reds_Prop (Γ : scope) (i : level): Γ⊨Sort ℙ i⇶*Sort ℙ 0. Proof. apply red_trans_direct. gred. Qed. Lemma reds_Pi (Γ : scope) (i j : level) (m mx : mode) (A A' B B' : term) : - Γ⊨A↣*A' → mx :: Γ⊨B↣*B' → Γ⊨Pi i j m mx A B↣*Pi (red_lvl mx i) (red_lvl m j) m mx A' B'. + Γ⊨A⇶*A' → mx :: Γ⊨B⇶*B' → Γ⊨Pi i j m mx A B⇶*Pi (red_lvl mx i) (red_lvl m j) m mx A' B'. Proof. intros red_A red_B. induction red_A. @@ -98,7 +98,7 @@ Proof. Qed. Lemma reds_lam (Γ : scope) (mx : mode) (A A' t t' : term) : - Γ⊨A↣*A' → mx :: Γ⊨t↣*t' → Γ⊨lam mx A t↣*lam mx A' t'. + Γ⊨A⇶*A' → mx :: Γ⊨t⇶*t' → Γ⊨lam mx A t⇶*lam mx A' t'. Proof. intros red_A red_t. induction red_A. @@ -107,7 +107,7 @@ Proof. Qed. Lemma reds_app (Γ : scope) (u u' v v' : term) : - Γ⊨u↣*u' → Γ⊨v↣*v' → Γ⊨app u v↣*app u' v'. + Γ⊨u⇶*u' → Γ⊨v⇶*v' → Γ⊨app u v⇶*app u' v'. Proof. intros red_u red_v. induction red_u. @@ -116,19 +116,19 @@ Proof. Qed. Lemma reds_Erased (Γ : scope) (A A' : term) : - Γ⊨A↣*A' → Γ⊨Erased A↣*Erased A'. + Γ⊨A⇶*A' → Γ⊨Erased A⇶*Erased A'. Proof. intro red_A; end_things red_A. Qed. Lemma reds_hide (Γ : scope) (A A' : term) : - Γ⊨A↣*A' → Γ⊨hide A↣*hide A'. + Γ⊨A⇶*A' → Γ⊨hide A⇶*hide A'. Proof. intro red_A; end_things red_A. Qed. Lemma reds_reveal (Γ : scope) (t t' P P' p p' : term) : - Γ⊨t↣*t' → Γ⊨P↣*P' → Γ⊨p↣*p' → Γ⊨reveal t P p↣*reveal t' P' p'. + Γ⊨t⇶*t' → Γ⊨P⇶*P' → Γ⊨p⇶*p' → Γ⊨reveal t P p⇶*reveal t' P' p'. Proof. intros red_t red_P red_p. induction red_t. @@ -139,7 +139,7 @@ Proof. Qed. Lemma reds_Reveal (Γ : scope) (t t' p p' : term) : - Γ⊨t↣*t' → Γ⊨p↣*p' → Γ⊨Reveal t p↣*Reveal t' p'. + Γ⊨t⇶*t' → Γ⊨p⇶*p' → Γ⊨Reveal t p⇶*Reveal t' p'. Proof. intros red_t red_p. induction red_t. @@ -148,7 +148,7 @@ Proof. Qed. Lemma reds_gheq (Γ : scope) (A A' u u' v v' : term): - Γ⊨A↣*A' → Γ⊨u↣*u' → Γ⊨v↣*v' → Γ⊨gheq A u v↣*gheq A' u' v'. + Γ⊨A⇶*A' → Γ⊨u⇶*u' → Γ⊨v⇶*v' → Γ⊨gheq A u v⇶*gheq A' u' v'. Proof. intros red_A red_u red_v. induction red_A. @@ -159,7 +159,7 @@ Proof. Qed. Lemma reds_if (Γ : scope) (m : mode) (b b' P P' t t' f f' : term) : - Γ⊨b↣*b' → Γ⊨P↣*P' → Γ⊨t↣*t' → Γ⊨f↣*f' → Γ⊨tif m b P t f↣*tif m b' P' t' f'. + Γ⊨b⇶*b' → Γ⊨P⇶*P' → Γ⊨t⇶*t' → Γ⊨f⇶*f' → Γ⊨tif m b P t f⇶*tif m b' P' t' f'. Proof. intros red_b red_P red_t red_f. induction red_b. @@ -173,13 +173,13 @@ Qed. Lemma reds_succ (Γ : scope) (n n' : term): - Γ⊨n↣*n' → Γ⊨tsucc n↣*tsucc n'. + Γ⊨n⇶*n' → Γ⊨tsucc n⇶*tsucc n'. Proof. intro red_A; end_things red_A. Qed. Lemma reds_nat_elim (Γ : scope) (m : mode) (n n' P P' z z' s s' : term) : - Γ⊨n↣*n' → Γ⊨P↣*P' → Γ⊨z↣*z' → Γ⊨s↣*s' → Γ⊨tnat_elim m n P z s↣*tnat_elim m n' P' z' s'. + Γ⊨n⇶*n' → Γ⊨P⇶*P' → Γ⊨z⇶*z' → Γ⊨s⇶*s' → Γ⊨tnat_elim m n P z s⇶*tnat_elim m n' P' z' s'. Proof. intros red_n red_P red_z red_s. induction red_n. @@ -192,7 +192,7 @@ Proof. Qed. Lemma reds_vec (Γ : scope) (A A' n n' : term) : - Γ⊨A↣*A' → Γ⊨n↣*n' → Γ⊨tvec A n↣*tvec A' n'. + Γ⊨A⇶*A' → Γ⊨n⇶*n' → Γ⊨tvec A n⇶*tvec A' n'. Proof. intros red_A red_n. induction red_A. @@ -201,13 +201,13 @@ Proof. Qed. Lemma reds_vnil (Γ : scope) (A A' : term) : - Γ⊨A↣*A' → Γ⊨tvnil A↣*tvnil A'. + Γ⊨A⇶*A' → Γ⊨tvnil A⇶*tvnil A'. Proof. intro red_A; end_things red_A. Qed. Lemma reds_vcons (Γ : scope) (a a' n n' v v' : term) : - Γ⊨a↣*a' → Γ⊨n↣*n' → Γ⊨v↣*v' → Γ⊨tvcons a n v↣*tvcons a' n' v'. + Γ⊨a⇶*a' → Γ⊨n⇶*n' → Γ⊨v⇶*v' → Γ⊨tvcons a n v⇶*tvcons a' n' v'. Proof. intros red_a red_n red_v. induction red_a. @@ -218,8 +218,8 @@ Proof. Qed. Lemma reds_vec_elim (Γ : scope) (m : mode) (A A' n n' v v' P P' z z' s s' : term): - Γ⊨A↣*A' → Γ⊨n↣*n' → Γ⊨v↣*v' → Γ⊨P↣*P' → Γ⊨z↣*z' → Γ⊨s↣*s' - → Γ⊨tvec_elim m A n v P z s↣*tvec_elim m A' n' v' P' z' s'. + Γ⊨A⇶*A' → Γ⊨n⇶*n' → Γ⊨v⇶*v' → Γ⊨P⇶*P' → Γ⊨z⇶*z' → Γ⊨s⇶*s' + → Γ⊨tvec_elim m A n v P z s⇶*tvec_elim m A' n' v' P' z' s'. Proof. intros red_A red_n red_v red_P red_z red_s. induction red_A. @@ -236,7 +236,7 @@ Proof. Qed. Lemma reds_bot_elim (Γ : scope) (m : mode) (A A' p p' : term) : - Γ⊨A↣*A' → Γ⊨p↣*p' → Γ⊨bot_elim m A p↣*bot_elim m A' p'. + Γ⊨A⇶*A' → Γ⊨p⇶*p' → Γ⊨bot_elim m A p⇶*bot_elim m A' p'. Proof. intros red_A red_p. induction red_A. @@ -264,8 +264,8 @@ Ltac greds := (* reds inversions *) Lemma reds_lam_inv {Γ : scope} {m : mode} {A t u: term} : - Γ⊨lam m A t↣* u → md Γ (lam m A t) ≠ ℙ → - (∃ A' t', u = lam m A' t' ∧ Γ ⊨ A ↣* A' ∧ m::Γ ⊨ t ↣* t'). + Γ⊨lam m A t⇶* u → md Γ (lam m A t) ≠ ℙ → + (∃ A' t', u = lam m A' t' ∧ Γ ⊨ A ⇶* A' ∧ m::Γ ⊨ t ⇶* t'). Proof. intros red_lam not_Prop. remember (lam m A t) as t0 eqn:e0. @@ -286,8 +286,8 @@ Proof. Qed. Lemma reds_Pi_inv {Γ : scope} {i j: level} {m mx : mode} {A B t: term} : - Γ⊨Pi i j m mx A B↣* t → - (∃ A' B' i' j', t = Pi i' j' m mx A' B' ∧ Γ ⊨ A ↣* A' ∧ mx::Γ ⊨ B ↣* B'). + Γ⊨Pi i j m mx A B⇶* t → + (∃ A' B' i' j', t = Pi i' j' m mx A' B' ∧ Γ ⊨ A ⇶* A' ∧ mx::Γ ⊨ B ⇶* B'). Proof. intro red_Pi. remember (Pi i j m mx A B) as t0 eqn:e0. @@ -306,7 +306,7 @@ Proof. Qed. Lemma reds_Sort_inv {Γ : scope} {i: level} {m : mode} {t: term} : - Γ⊨Sort m i ↣* t → ∃ i', t= Sort m i'. + Γ⊨Sort m i ⇶* t → ∃ i', t= Sort m i'. Proof. intro red_sort. remember (Sort m i) as t0 eqn:e0. @@ -317,3 +317,33 @@ Proof. destruct H as [i' e]. eauto. Qed. + +Lemma reds_Erased_inv {Γ : scope} {u0 v: term} : + Γ⊨ Erased u0 ⇶* v → ∃ v0, v = Erased v0 ∧ Γ ⊨ u0 ⇶* v0. +Proof. + intro reds. + remember (Erased u0) as u eqn:eu. + induction reds as [|u w v H red_v IH] in v, u0, eu, reds. + - subst; eexists; split; [reflexivity | constructor; reflexivity]. + - subst. + apply red_Erased_inv in H. + destruct H as [v0 [ev red_u0]]. + subst. + specialize (IH v0 eq_refl) as [w0 [e red_v0]]. + exists w0; split; [exact e | econstructor; eassumption]. +Qed. + +Lemma reds_vec_inv {Γ : scope} {A0 n0 v: term} : + Γ⊨ tvec A0 n0 ⇶* v → ∃ A1 n1, v = tvec A1 n1 ∧ Γ ⊨ A0 ⇶* A1 ∧ Γ ⊨ n0 ⇶* n1. +Proof. + intro reds. + remember (tvec A0 n0) as u eqn:eu. + induction reds as [|u w v H red_v IH] in v, A0, n0, eu, reds. + - subst; repeat eexists; constructor; reflexivity. + - subst. + apply red_vec_inv in H. + destruct H as [A1 [n1 [ev [red_A0 red_n0]]]]. + subst. + specialize (IH A1 n1 eq_refl) as [A2 [n2 [e [red_A1 red_n1]]]]. + exists A2, n2; split; [exact e | split; econstructor; eassumption]. +Qed. diff --git a/theories/reduction/onestep/Properties.v b/theories/reduction/onestep/Properties.v new file mode 100644 index 0000000..9939999 --- /dev/null +++ b/theories/reduction/onestep/Properties.v @@ -0,0 +1,59 @@ +(* Definition of reduction rules which corresponds to the congruence relation *) +(* and proof that the system is confluent *) +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. +From GhostTT.reduction Require Import Notations Injectivity. +From GhostTT.reduction.wrapping Require Import Properties. +From GhostTT.reduction.onestep Require Export Reduction. + +Import ListNotations. +Set Default Goal Selector "!". + +Lemma red_scope {Γ : context} {Ω t t': term} : + wf Γ → Γ ⊢ t:Ω → t ↣ t' → Γ ⊢ t'∷mdc Γ t. +Proof. + intros wfΓ type_t red_t. + specialize (validity Γ t Ω wfΓ type_t) as [scope_t _]. + remember (sc Γ) as Γ0 eqn:eΓ. + induction red_t in Γ, Γ0, eΓ, Ω, t, t', red_t, type_t, scope_t. + all: try solve [inversion scope_t; gscope]. + - ttinv type_t type_t. + destruct type_t as [mx' [m [i [j [A0 [B0 type_t]]]]]]. + destruct type_t as [scope_B0 [scope_lam [scope_u [scope_A0 type_t]]]]. + destruct type_t as [type_lam [type_u [type_A0 [type_B0 conv_Ω]]]]. + ttinv type_lam H. + destruct H as [i0 [j0 [m0 [B H]]]]. + destruct H as [scope_A [scope_B [scope_t0 H]]]. + destruct H as [type_A [type_B [type_t0 conv_pi]]]. + cbn in conv_pi. + apply injectivity_Pi in conv_pi; gscope; eauto using scoping_castrm. + destruct conv_pi as [e0 [e1 [conv_A conv_B]]]. + subst. cbn. + erewrite scoping_md; eauto. + eapply scoping_subst; eauto. + exact (sscoping_one (sc Γ) _ _ scope_u). + - inversion scope_t. + match goal with H : _ ⊨ hide _∷_ |- _ => + inversion H; subst end. + gscope. + - inversion scope_t. + match goal with H : _ ⊨ tsucc _∷_ |- _ => + inversion H; subst end. + gscope. + - inversion scope_t. + match goal with H : _ ⊨ tvcons _ _ _∷_ |- _ => + inversion H; subst end. + gscope; eauto. + * intro H; inversion H. + * eapply scoping_ren; eauto using rscoping_S. + * eapply scoping_ren; eauto using rscoping_S. + eapply scoping_ren; eauto using rscoping_S. + * right; left. reflexivity. + - eapply scoping_change_box; eauto. + apply typing_box in type_t as [B type_u]. + apply scoping_box in scope_t. + subst; cbn in *. + eapply IHred_t; eauto. + apply sc_scope_box_term. +Qed. diff --git a/theories/reduction/onestep/Reduction.v b/theories/reduction/onestep/Reduction.v index c0571ba..068dbdd 100644 --- a/theories/reduction/onestep/Reduction.v +++ b/theories/reduction/onestep/Reduction.v @@ -1,17 +1,15 @@ (* Definition of reduction rules which corresponds to the congruence relation *) (* and proof that the system is confluent *) From Coq Require Import Utf8 List. -From GhostTT.autosubst Require Import GAST unscoped. -From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. -From GhostTT Require Export Univ TermNotations Typing. +From GhostTT.autosubst Require Import GAST. +From GhostTT Require Import SubstNotations. From GhostTT.reduction Require Export Notations Utils. +From GhostTT.reduction.wrapping Require Export Core. Import ListNotations. Set Default Goal Selector "!". (* ------------------------------------------------------------------------- *) -Reserved Notation "t ↣ t'" (at level 70). - Inductive reduction : term → term → Prop := (* Computation rules *) @@ -55,21 +53,8 @@ Inductive reduction : term → term → Prop := | red_subterm : ∀ u u' C, u ↣ u' → - C <[u··] ↣ C <[u'··] - - (* (* Structural rule *) - - | red_refl : - ∀ u, - Γ ⊨ u ↣ u *) - - (* (* Proof irrelevance *) - - | red_irr : - ∀ p, - md Γ p = ℙ → - Γ ⊨ p ↣ ⋆ *) + C □ u ↣ C □ u' where "u ↣ v" := (reduction u v). - (* ------------------------------------------------------------------------- *) +(* ------------------------------------------------------------------------- *) diff --git a/theories/reduction/onestep/SubjectReduction.v b/theories/reduction/onestep/SubjectReduction.v new file mode 100644 index 0000000..2960f44 --- /dev/null +++ b/theories/reduction/onestep/SubjectReduction.v @@ -0,0 +1,341 @@ +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval + TermMode Scoping BasicMetaTheory. +From GhostTT.reduction Require Import Injectivity Model. +From GhostTT.reduction.wrapping Require Import Core Properties. +From GhostTT.reduction.onestep Require Export Reduction. +From GhostTT Require Import Admissible. + +Import ListNotations. + +Set Default Goal Selector "!". + +Section subjet_reduction. + + Ltac ttinv_destruct h := + destruct_exists h; + let rec destruct_conj h := + lazymatch type of h with + | _ ∧ _ => let H := fresh "H" in + destruct h as [H h]; destruct_conj h + |_ => idtac end + in destruct_conj h. + + Local Notation 𝝢 := (Erased tnat). + Local Notation "P ⁺¹" := (S ⋅ P) (at level 1). + Local Notation "P ⁺²" := (S ⋅ S ⋅ P) (at level 1). + Local Notation "P ⁺³" := (S ⋅ S ⋅ S ⋅ P) (at level 1). + Local Notation "P ⁺⁴" := (S ⋅ S ⋅ S ⋅ S ⋅ P) (at level 1). + + Lemma subst_S1 {t u : term} : + (S ⋅ t) <[ u··] = t. + Proof. + intros; asimpl; reflexivity. + Qed. + + Lemma subst_Sn {t : term} {σ : nat → term} : + (S ⋅ t) <[up_term σ] = S ⋅ t <[σ]. + Proof. + intros; asimpl; reflexivity. + Qed. + + + Theorem subjet_reduction (Γ: context) (Ω t t': term) + (wfΓ : wf Γ) (red_t : t ↣ t') (type_t : Γ ⊢ t:Ω) : + Γ⊢ t':Ω. + Proof. + destruct red_t + as [m A t0 u|t0 P p| m P t f|m P t f|m P z s|m P z s n|m A B P z s|m A a n0 n v P z s | i | i j m mx A B | ] + in Γ, wfΓ, t, t', Ω, red_t, type_t. + all: specialize (validity Γ _ _ wfΓ type_t) as [scope_t [i_u type_Ω]]. + all: specialize (scoping_type wfΓ type_t) as scope_Ω. + + (** Congruences **) + 11: {specialize (typing_box type_t) as [B type_u]. + admit. } + + + (** Computations **) + all: ttinv type_t type_t. + (* if_true if_false nat_elim_zero *) + 3-5: ttinv_destruct type_t; eauto using type_conv. + + + (* Beta_red *) + - (* destruction *) + destruct type_t as [mx_u [m_t0 [i [j [A0 [B0 type_t]]]]]]. + destruct type_t as [scope_B0 [scope_lam [scope_u [scope_A0 type_t]]]]. + destruct type_t as [type_lam [type_u [type_A0 [type_B0 conv_Ω]]]]. + ttinv type_lam H. + destruct H as [i0 [j0 [m0 [B H]]]]. + destruct H as [scope_A [scope_B [scope_t0 H]]]. + destruct H as [type_A [type_B [type_t0 conv_Pi]]]. + cbn in *; apply injectivity_Pi in conv_Pi as [e0 [e1 [conv_A conv_B]]]; subst. + 2,3: gscope; apply scoping_castrm; assumption. + apply castrm_castrm_conv in conv_A, conv_B. + + (* conversion *) + assert (Γ ⊢ B <[ u··] ε≡ Ω) as conv_Ω'. + { eapply conv_trans; [ | exact conv_Ω]. + rewrite !castrm_subst. + eapply conv_subst; eauto. + fsimpl; cbn. + eapply sscoping_one. + exact (scoping_castrm (sc Γ) u mx_u scope_u). } + + (* typing *) + refine (type_conv wfΓ _ _ conv_Ω' type_Ω). + * erewrite scoping_md; eauto. + eapply scoping_subst; eauto using sscoping_one, sscoping_comp_one. + * eapply typing_subst; eauto. + eapply styping_one; eauto. + apply conv_sym in conv_A. + exact (type_conv wfΓ scope_u type_u conv_A type_A). + + + (* reveal_hide *) + - (* destruction *) + destruct type_t as [i [j [m [A0 type_t]]]]. + destruct type_t as [scope_p [scope_hide [scope_P [scope_A0 type_t]]]]. + destruct type_t as [Hm [type_hide [type_P [type_p conv_Ω]]]]. + ttinv type_hide H. + destruct H as [i0 [A H]]. + destruct H as [scope_A [scope_t0 H]]. + destruct H as [type_A [type_t0 conv_Erased]]. + specialize (validity Γ _ _ wfΓ type_hide) as [_ [i_h type_Erased]]. + ttinv type_Erased H. + destruct H as [i_e [_ [type_A0 _]]]. + cbn in conv_Erased; apply injectivity_Erased in conv_Erased as conv_A; subst. + 2,3: gscope; apply scoping_castrm; assumption. + apply castrm_castrm_conv in conv_A. + clear type_Erased i_h conv_Erased. + + (* conversion *) + assert (Γ ⊢ app (S ⋅ P) (hide (var 0)) <[t0··] ε≡ Ω) as conv_Ω'. + { eapply conv_trans; [ | exact conv_Ω]. + asimpl. apply conv_refl. } + + (* typing *) + refine (type_conv wfΓ _ _ conv_Ω' type_Ω). + * do 2 (erewrite scoping_md; gscope). + * apply (type_app wfΓ type_p). + exact (type_conv wfΓ scope_t0 type_t0 conv_A type_A0). + + + (* nat_elim_succ *) + - (* destruction *) + destruct type_t as [i [Hm type_t]]. + destruct type_t as [scope_succ [scope_P [scope_z [scope_s type_t]]]]. + destruct type_t as [type_succ [type_P [type_z [type_s conv_Ω]]]]. + ttinv type_succ H. + destruct H as [i0 [type_n _]]. + + (* conversion *) + assert ( Γ ⊢ app (S ⋅ P) (tsucc (S ⋅ n)) <[(tnat_elim m n P z s)··] ε≡ Ω) as conv_Ω'. + { eapply conv_trans; [ | exact conv_Ω]. + asimpl. apply conv_refl. } + + (* sub-typing *) + assert (Γ ⊢ app s n : Pi i i m m (app P n) (app (S ⋅ P) (tsucc (S ⋅ n)))) as type_s_n. + { specialize (type_app wfΓ type_s type_n) as type_s_n. + cbn in type_s_n. + rewrite subst_Sn, subst_S1 in type_s_n. + exact type_s_n. } + + (* typing *) + refine (type_conv wfΓ _ _ conv_Ω' type_Ω). + * do 2 (erewrite scoping_md; gscope). + * refine (type_app wfΓ type_s_n _). + gtype. + + + (* vec_elim_nil *) + - (* destruction *) + destruct type_t as [i [j [Hm type_t]]]. + destruct type_t as [scope_nil [scope_P [scope_z [scope_s type_t]]]]. + destruct type_t as [type_nil [type_P [type_z [type_s type_t]]]]. + destruct type_t as [scope_hide [scope_A [type_A [type_hide conv_Ω]]]]. + ttinv type_nil H. + destruct H as [i0 H]. + destruct H as [scope_B [type_B conv_vec]]. + cbn in conv_vec; apply injectivity_vec in conv_vec as [conv_B _]; subst. + 2,3: gscope; apply scoping_castrm; assumption. + apply castrm_castrm_conv in conv_B. + + (* conversion *) + assert (Γ ⊢ app (app P (hide tzero)) (tvnil A) ε≡ Ω) as conv_Ω'. + { eapply conv_trans; [ | exact conv_Ω]. + cbn; gconv. exact (conv_sym Γ ε|B| ε|A| conv_B). } + + (* typing *) + exact (type_conv wfΓ scope_z type_z conv_Ω' type_Ω). + + + (* vec_elim_cons *) (* TLDR : just don't read. *) + - (* destruction *) + destruct type_t as [i [j [Hm type_t]]]. + destruct type_t as [scope_cons [scope_P [scope_z [scope_s type_t]]]]. + destruct type_t as [type_cons [type_P [type_z [type_s type_t]]]]. + destruct type_t as [scope_n [scope_A [type_A [type_n conv_Ω]]]]. + ttinv type_cons H. + destruct H as [i0 [A0 H]]. + destruct H as [scope_a [scope_n0 [scope_v H]]]. + destruct H as [type_a [type_n0 [type_v [type_A0 [scope_A0 conv_vec]]]]]. + cbn in conv_vec; apply injectivity_vec in conv_vec as [conv_A0 conv_gS_n0]; subst. + 2,3: gscope; eauto using scoping_castrm. + 2: right; left; reflexivity. + change (Γ ⊢ ε|gS n0| ε≡ ε|n|) in conv_gS_n0. + apply castrm_castrm_conv in conv_A0, conv_gS_n0. + specialize (cong_vec Γ _ _ _ _ conv_A0 (conv_refl Γ ε|n0|)) as conv_vec. + change (Γ ⊢ tvec A0 n0 ε≡ tvec A n0) in conv_vec. + specialize (type_vec Γ A n0 _ scope_A scope_n0 type_A type_n0) as type_vec. + specialize (type_conv wfΓ scope_v type_v conv_vec type_vec) as type_v'. + + + (* work on type_s and type_P *) + (* unfold umax in type_s, type_P. + cbn in type_s, type_P. + assert ((if isProp m then 0 else if isProp m then j else Nat.max j j) = if isProp m then 0 else j) as e. + { case (isProp m); [reflexivity| apply PeanoNat.Nat.max_id]. } + rewrite e in *; clear e. + assert (∀ j, Nat.max i (if isProp m then 0 else j) = if isProp m then i else Nat.max i j) as e. + { intro. case (isProp m); [apply PeanoNat.Nat.max_0_r| reflexivity]. } + rewrite e in *; clear e. + assert (∀ i j, (if isProp m then 0 else if isProp m then i else j) = if isProp m then 0 else j ) as e. + { intros. case (isProp m); reflexivity. } + rewrite !e in *; clear e. + remember (if isProp m then 0 else Nat.max i j) as k eqn:ek. + remember (if isProp m then 0 else j) as l eqn:el. + change (Γ ⊢ s : + Pi i k m 𝕋 A + (Pi 0 k m 𝔾 𝝢 + (Pi i l m 𝕋 (tvec A⁺² (var 0)) + (Pi j j m m (app (app P⁺³ (var 1)) (var 0)) + (app (app P⁺⁴ (gS (var 2))) + (tvcons (var 3) (var 2) (var 1))))))) + in type_s. *) + remember (glength A n0 v) as gl eqn:e. + remember (umax m m j j) as k2 eqn:ek2. + (* + remember (umax 𝔾 m 0 (umax 𝕋 m i k2)) as k0 eqn:ek0. + remember (umax 𝕋 m i k2) as k1 eqn:ek1. + change (Γ ⊢ s : + Pi i k0 m 𝕋 A + (Pi 0 k1 m 𝔾 𝝢 + (Pi i k2 m 𝕋 (tvec A⁺² (var 0)) + (Pi j j m m (app (app P⁺³ (var 1)) (var 0)) + (app (app P⁺⁴ (gS (var 2))) + (tvcons (var 3) (var 2) (var 1))))))) + in type_s. *) + + (* sub-sub-sub-typing *) + (* assert (Γ ⊢ app s a : + Pi 0 k1 m 𝔾 𝝢 + (Pi i k2 m 𝕋 (tvec A⁺¹ (var 0)) + (Pi j j m m (app (app P⁺² (var 1)) (var 0)) + (app (app P⁺³ (gS (var 2))) + (tvcons a⁺³ (var 2) (var 1)))))) + as type_s_a. + { *) + specialize (type_conv wfΓ scope_a type_a conv_A0 type_A) as type_a'. + specialize (type_app wfΓ type_s type_a') as type_s_a. + cbn in type_s_a. + rewrite !subst_Sn, !subst_S1 in type_s_a. + (* exact type_s_a. } *) + + (* TODO *) + (* assumption hard to prove *) + assert (Γ ⊢ gl ε≡ n0) as conv_gl. { admit. } + assert (Γ ⊢ gl : 𝝢) as type_gl. { admit. } + assert (Γ ⊢ gl ∷ 𝕋) as scope_gl. { admit. } + + (* sub-sub-typing *) + assert (Γ ⊢ app (app s a) gl : + Pi i k2 m 𝕋 (tvec A n0) + (Pi j j m m (app (app P⁺¹ n0⁺¹) (var 0)) + (app (app P⁺² (n⁺²)) + (tvcons a⁺² n0⁺² (var 1))))) + as type_s_a_gl. + { assert (Γ ⊢ + Pi i k2 m 𝕋 (tvec A gl) + (Pi j j m m (app (app P⁺¹ gl⁺¹) (var 0)) + (app (app P⁺² (gS gl⁺²)) (tvcons a⁺² gl⁺² (var 1)))) + ε≡ + Pi i k2 m 𝕋 (tvec A n0) + (Pi j j m m (app (app P⁺¹ n0⁺¹) (var 0)) + (app (app P⁺² (n⁺²)) (tvcons a⁺² n0⁺² (var 1))))) + as conv_pi. + { cbn; gconv. + 4-7: right; reflexivity. + all: rewrite !castrm_ren. + 2: eapply (conv_trans _ _ (ε|gS gl|)⁺² ε|n|⁺²); [apply conv_refl | ]. + all: repeat (eapply conv_ren; eauto using rtyping_S). + refine (conv_trans Γ ε|gS gl| ε|gS n0| ε|n| _ conv_gS_n0). + gconv. + } + specialize (type_app wfΓ type_s_a type_gl) as type_s_a_gl. + cbn in type_s_a_gl. + rewrite !subst_Sn, !subst_S1 in type_s_a_gl. + change (Γ ⊢ app (app s a) gl : + Pi i k2 m 𝕋 (tvec A gl) + (Pi j j m m (app (app P⁺¹ gl⁺¹) (var 0)) + (app (app P⁺² (gS gl⁺²)) + (tvcons a⁺² gl⁺² (var 1))))) + in type_s_a_gl. + refine (type_conv wfΓ _ type_s_a_gl conv_pi _). + * gscope. + * apply (type_pi wfΓ type_vec). + subst k2. + refine (type_pi (wf_cons wfΓ type_vec) _ _). + all:admit. (* some app to type... *) + } + + (* sub-typing *) + (* assert (Γ ⊢ app (app (app s a) gl) v : + Pi j j m m (app (app P n0) v) + (app (app P⁺¹ n⁺¹) + (tvcons a⁺¹ n0⁺¹ v⁺¹))) + as type_s_a_gl_v. + { *) + specialize (type_app wfΓ type_s_a_gl type_v') as type_s_a_gl_v. + cbn in type_s_a_gl_v. + rewrite !subst_Sn, !subst_S1 in type_s_a_gl_v. + (* exact type_s_a_gl_v. + } *) + + (* conversion *) + assert ( Γ ⊢ app (app P⁺¹ n⁺¹) (tvcons a⁺¹ n0⁺¹ v⁺¹) <[(tvec_elim m A n0 v P z s)··] ε≡ Ω) as conv_Ω'. + { eapply conv_trans; [ | exact conv_Ω]. + asimpl. apply conv_refl. } + + (* typing *) + refine (type_conv wfΓ _ _ conv_Ω' type_Ω). + * subst; repeat (erewrite scoping_md; gscope); auto. + (* + intro H; inversion H. + + eapply scoping_ren; eauto using rscoping_S. + + do 2 (eapply scoping_ren; eauto using rscoping_S). + + right; left; reflexivity. *) + * refine (type_app wfΓ type_s_a_gl_v _). + subst; gtype. + + (* Sort *) + - exact (type_conv wfΓ (scope_sort (sc Γ) ℙ 0) (type_sort Γ ℙ 0) type_t type_Ω). + + + (* Pi *) + - (* destruction *) + destruct type_t as [scope_A [scope_B type_t]]. + destruct type_t as [type_A [type_B conv_Ω]]. + + (* typing *) + refine (type_conv wfΓ _ _ conv_Ω type_Ω). + * gscope. + * case m, mx; cbn; gtype. + all: refine (type_conv _ _ _ _ _); eauto. + all: try apply cong_Prop. + all: try apply type_sort. + all: eauto using wf_cons. + Admitted. + +End subjet_reduction. diff --git a/theories/reduction/wrapping/Core.v b/theories/reduction/wrapping/Core.v new file mode 100644 index 0000000..f3282cb --- /dev/null +++ b/theories/reduction/wrapping/Core.v @@ -0,0 +1,161 @@ +(* Definition of reduction rules which corresponds to the congruence relation *) +(* and proof that the system is confluent *) +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping. +From GhostTT Require Export Univ TermNotations Typing. +From GhostTT.reduction Require Export Notations Utils. + +Import ListNotations. +Set Default Goal Selector "!". + + +Inductive box_term_1 : Set := + | Box_Pi_1 : level → level → mode → mode → term → box_term_1 + | Box_lam_1 : mode → term → box_term_1 + | Box_app_1 : term → box_term_1 + | Box_app_2 : term → box_term_1 + | Box_Erased : box_term_1 + | Box_hide : box_term_1 + | Box_reveal_1 : term → term → box_term_1 + | Box_reveal_2 : term → term → box_term_1 + | Box_reveal_3 : term → term → box_term_1 + | Box_Reveal_1 : term → box_term_1 + | Box_Reveal_2 : term → box_term_1 + | Box_toRev_1 : term → term → box_term_1 + | Box_toRev_2 : term → term → box_term_1 + | Box_toRev_3 : term → term → box_term_1 + | Box_fromRev_1 : term → term → box_term_1 + | Box_fromRev_2 : term → term → box_term_1 + | Box_fromRev_3 : term → term → box_term_1 + | Box_gheq_1 : term → term → box_term_1 + | Box_gheq_2 : term → term → box_term_1 + | Box_gheq_3 : term → term → box_term_1 + | Box_ghrefl_1 : term → box_term_1 + | Box_ghrefl_2 : term → box_term_1 + | Box_ghcast_1 : term → term → term → term → term → box_term_1 + | Box_ghcast_2 : term → term → term → term → term → box_term_1 + | Box_ghcast_3 : term → term → term → term → term → box_term_1 + | Box_ghcast_4 : term → term → term → term → term → box_term_1 + | Box_ghcast_5 : term → term → term → term → term → box_term_1 + | Box_ghcast_6 : term → term → term → term → term → box_term_1 + | Box_tif_1 : mode → term → term → term → box_term_1 + | Box_tif_2 : mode → term → term → term → box_term_1 + | Box_tif_3 : mode → term → term → term → box_term_1 + | Box_tif_4 : mode → term → term → term → box_term_1 + | Box_tsucc : box_term_1 + | Box_tnat_elim_1 : mode → term → term → term → box_term_1 + | Box_tnat_elim_2 : mode → term → term → term → box_term_1 + | Box_tnat_elim_3 : mode → term → term → term → box_term_1 + | Box_tnat_elim_4 : mode → term → term → term → box_term_1 + | Box_tvec_1 : term → box_term_1 + | Box_tvec_2 : term → box_term_1 + | Box_tvnil : box_term_1 + | Box_tvcons_1 : term → term → box_term_1 + | Box_tvcons_2 : term → term → box_term_1 + | Box_tvcons_3 : term → term → box_term_1 + | Box_tvec_elim_1 : mode → term → term → term → term → term → box_term_1 + | Box_tvec_elim_2 : mode → term → term → term → term → term → box_term_1 + | Box_tvec_elim_3 : mode → term → term → term → term → term → box_term_1 + | Box_tvec_elim_4 : mode → term → term → term → term → term → box_term_1 + | Box_tvec_elim_5 : mode → term → term → term → term → term → box_term_1 + | Box_tvec_elim_6 : mode → term → term → term → term → term → box_term_1 + | Box_bot_elim_1 : mode → term → box_term_1 + | Box_bot_elim_2 : mode → term → box_term_1. + +Inductive box_term_2 : Set := + | Box_Pi_2 : level → level → mode → mode → term → box_term_2 + | Box_lam_2 : mode → term → box_term_2. + +Inductive box_term : Set := + | Box_1 : box_term_1 → box_term + | Box_2 : box_term_2 → box_term. + +Definition eval_box_term (C : box_term) (t0: term) : term := + match C with + | Box_1 (Box_Pi_1 i j m mx B) => Pi i j m mx t0 B + | Box_2 (Box_Pi_2 i j m mx A) => Pi i j m mx A t0 + | Box_1 (Box_lam_1 m t) => lam m t0 t + | Box_2 (Box_lam_2 m A) => lam m A t0 + | Box_1 (Box_app_1 v) => app t0 v + | Box_1 (Box_app_2 u) => app u t0 + | Box_1 (Box_Erased) => Erased t0 + | Box_1 (Box_hide) => hide t0 + | Box_1 (Box_reveal_1 P p) => reveal t0 P p + | Box_1 (Box_reveal_2 t p) => reveal t t0 p + | Box_1 (Box_reveal_3 t P) => reveal t P t0 + | Box_1 (Box_Reveal_1 p) => Reveal t0 p + | Box_1 (Box_Reveal_2 t) => Reveal t t0 + | Box_1 (Box_toRev_1 p u) => toRev t0 p u + | Box_1 (Box_toRev_2 t u) => toRev t t0 u + | Box_1 (Box_toRev_3 t p) => toRev t p t0 + | Box_1 (Box_fromRev_1 p u) => fromRev t0 p u + | Box_1 (Box_fromRev_2 t u) => fromRev t t0 u + | Box_1 (Box_fromRev_3 t p) => fromRev t p t0 + | Box_1 (Box_gheq_1 u V) => gheq t0 u V + | Box_1 (Box_gheq_2 a V) => gheq a t0 V + | Box_1 (Box_gheq_3 a u) => gheq a u t0 + | Box_1 (Box_ghrefl_1 u) => ghrefl t0 u + | Box_1 (Box_ghrefl_2 A) => ghrefl A t0 + | Box_1 (Box_ghcast_1 u v e P t) => ghcast t0 u v e P t + | Box_1 (Box_ghcast_2 A v e P t) => ghcast A t0 v e P t + | Box_1 (Box_ghcast_3 A u e P t) => ghcast A u t0 e P t + | Box_1 (Box_ghcast_4 A u v P t) => ghcast A u v t0 P t + | Box_1 (Box_ghcast_5 A u v e t) => ghcast A u v e t0 t + | Box_1 (Box_ghcast_6 A u v e P) => ghcast A u v e P t0 + | Box_1 (Box_tif_1 m P t f) => tif m t0 P t f + | Box_1 (Box_tif_2 m b t f) => tif m b t0 t f + | Box_1 (Box_tif_3 m b P f) => tif m b P t0 f + | Box_1 (Box_tif_4 m b P t) => tif m b P t t0 + | Box_1 (Box_tsucc) => tsucc t0 + | Box_1 (Box_tnat_elim_1 m P z s) => tnat_elim m t0 P z s + | Box_1 (Box_tnat_elim_2 m n z s) => tnat_elim m n t0 z s + | Box_1 (Box_tnat_elim_3 m n P s) => tnat_elim m n P t0 s + | Box_1 (Box_tnat_elim_4 m n P z) => tnat_elim m n P z t0 + | Box_1 (Box_tvec_1 n) => tvec t0 n + | Box_1 (Box_tvec_2 A) => tvec A t0 + | Box_1 (Box_tvnil) => tvnil t0 + | Box_1 (Box_tvcons_1 n v) => tvcons t0 n v + | Box_1 (Box_tvcons_2 a v) => tvcons a t0 v + | Box_1 (Box_tvcons_3 a n) => tvcons a n t0 + | Box_1 (Box_tvec_elim_1 m n v P z s) => tvec_elim m t0 n v P z s + | Box_1 (Box_tvec_elim_2 m A v P z s) => tvec_elim m A t0 v P z s + | Box_1 (Box_tvec_elim_3 m A n P z s) => tvec_elim m A n t0 P z s + | Box_1 (Box_tvec_elim_4 m A n v z s) => tvec_elim m A n v t0 z s + | Box_1 (Box_tvec_elim_5 m A n v P s) => tvec_elim m A n v P t0 s + | Box_1 (Box_tvec_elim_6 m A n v P z) => tvec_elim m A n v P z t0 + | Box_1 (Box_bot_elim_1 m p) => bot_elim m t0 p + | Box_1 (Box_bot_elim_2 m A) => bot_elim m A t0 + end. + +Definition context_box_term (Γ: context) (C : box_term) : context := + match C with + | Box_1 _ => Γ + | Box_2 (Box_Pi_2 i j m mx A) => (Γ,, (mx,A)) + | Box_2 (Box_lam_2 m A) => (Γ,, (m,A)) + end. + +Definition scope_box_term (Γ: scope) (C : box_term) : scope := + match C with + | Box_1 _ => Γ + | Box_2 (Box_Pi_2 i j m mx A) => mx::Γ + | Box_2 (Box_lam_2 m A) => m::Γ + end. + +Notation "□¹_term" := box_term_1. +Notation "□²_term" := box_term_2. +Notation "□_term" := box_term. +Notation "C □ t" := (eval_box_term C t) (at level 70). +Notation "C □¹ t" := (eval_box_term (Box_1 C) t) (at level 70). +Notation "C □² t" := (eval_box_term (Box_2 C) t) (at level 70). + +Notation "[ Γ ] □ C" := (context_box_term Γ C) (at level 70). +Notation "[| Γ |] □ C" := (scope_box_term Γ C) (at level 70). + +Notation "Γ □ C ⊢ t : A" := (typing (context_box_term Γ C) t A) + (at level 80, t, A at next level, format "Γ □ C ⊢ t : A", only printing). +Notation "Γ □ C ⊢ t ∷ m" := (scoping (sc (context_box_term Γ C)) t m) + (at level 80, t, m at next level, format "Γ □ C ⊢ t ∷ m", only printing). +Notation "Γ □ C ⊨ t ∷ m" := (scoping (scope_box_term Γ C) t m) + (at level 80, t, m at next level, format "Γ □ C ⊨ t ∷ m", only printing). + diff --git a/theories/reduction/wrapping/Properties.v b/theories/reduction/wrapping/Properties.v new file mode 100644 index 0000000..8a470df --- /dev/null +++ b/theories/reduction/wrapping/Properties.v @@ -0,0 +1,103 @@ +(* Definition of reduction rules which corresponds to the congruence relation *) +(* and proof that the system is confluent *) +From Coq Require Import Utf8 List. +From GhostTT.autosubst Require Import GAST unscoped. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. +From GhostTT.reduction.wrapping Require Export Core. + +Import ListNotations. +Set Default Goal Selector "!". + +Ltac ttinv_destruct h HN:= + ttinv h HN; destruct_exists HN; + let rec destruct_conj HN := + lazymatch type of HN with + | _ ∧ _ => let H := fresh "H" in + destruct HN as [H HN]; destruct_conj HN + |_ => idtac end + in destruct_conj HN. + +Lemma scoping_box {Γ : scope} {m : mode } {u : term} {C: □_term} : + Γ ⊨ C □ u∷m → [|Γ|] □ C ⊨ u∷md ([|Γ|] □ C) u. +Proof. + intros scope_Cu. + destruct C as [C|C]. + all: destruct C; cbn in *. + all: inversion scope_Cu. + all: erewrite scoping_md. + all: eauto. +Qed. + +Lemma scoping_change_box {Γ : scope} {m : mode } {u u': term} {C: □_term} : + Γ ⊨ C □ u∷m → [|Γ|] □ C ⊨ u'∷md ([|Γ|] □ C) u → Γ ⊨ C □ u'∷m. +Proof. + intros scope_Cu scope_u'. + destruct C as [C|C]. + all: destruct C; cbn in *. + all: inversion scope_Cu. + all: try (gscope; erewrite <- scoping_md; eauto). +Qed. + +Lemma typing_box {Γ : context} {u A: term} {C: □_term} : + Γ ⊢ C □ u : A → ∃ B, [Γ] □ C ⊢ u : B. +Proof. + intro type_Cu. + destruct C as [C|C]. + all: destruct C; cbn in *. + all: ttinv_destruct type_Cu type_Cu. + all: eauto. +Qed. + +Lemma sc_scope_box_term {Γ : context} {C: □_term} : + [| sc Γ |] □ C = sc ([ Γ ] □ C). +Proof. + destruct C as [C|[C|C]]; reflexivity. +Qed. + +(* +Lemma type_conv_urm {Γ : context} {i : level} {m : mode} {A B t : term}: + Γ⊢A∷𝕂 → Γ⊢B∷𝕂 → Γ⊢t∷m → Γ ⊢ t : A → Γ ⊢ A ≈ B → Γ ⊢ B : Sort m i → Γ ⊢ t : B. +Proof. + intros scope_A scope_B scope_t type_A conv_urm_A type_B. + induction conv_urm_A. + type_conv. + +Lemma typing_change_box {Γ : context} {A B u u': term} {C: □_term} (wfΓ : wf Γ): + Γ ⊢ C □ u:A → [Γ] □ C ⊢ u :B → + [Γ] □ C ⊢ u'∷mdc ([Γ] □ C) u → + [Γ] □ C ⊢ u':B → Γ ⊢ C □ u':A. +Proof. + intros type_Cu type_u scope_u' type_u'. + specialize (validity Γ _ _ wfΓ type_Cu) as [scope_Cu [i_u type_A]]. + specialize (scoping_box scope_Cu) as scope_u. + remember (mdc Γ (C □ u)) as md_t eqn:e_t. + rewrite sc_scope_box_term in *. + remember ([Γ] □ C) as Δ eqn:e_Δ. + remember (mdc Δ u) as md_u eqn:e_u. + specialize (scoping_type wfΓ type_Cu) as scope_A. + assert (Γ⊢C □ u'∷md_t) as scope_Cu'. + { eapply scoping_change_box; eauto. + rewrite sc_scope_box_term. + subst; assumption. } + destruct C as [[]|[]]; cbn in *. + all: ttinv_destruct type_Cu conv_A. + all: refine (type_conv Γ i_u _ _ A _ _ scope_A _ _ conv_A type_A); [gscope| gscope | ]; cbn; eauto using scoping_subst. + 3, 5: admit. + all: gtype. + all: try (subst; erewrite scoping_md in scope_u'; eauto). + * specialize (type_unique type_u H1) as e. + * admit. + - gtype. + * gtype. + ttinv_destruct H2 H2. + * typing + all: inversion scope_Cu. + all: inversion scope_Cu'. + all: ttinv_destruct type_Cu conv_A. + all: refine (type_conv Γ i_u _ _ A _ _ scope_A _ _ conv_A type_A); gscope; gtype. + - gscope. + - gscope. + (* + all: gtype. + all: try (gscope; erewrite <- scoping_md; eauto). + Qed.*) Abort. *) From 253cca834b72dca5e582726dc675213f30f52436 Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Wed, 26 Jun 2024 14:32:54 +0200 Subject: [PATCH 09/15] fix missing code --- theories/BasicMetaTheory.v | 24 ++++++++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/theories/BasicMetaTheory.v b/theories/BasicMetaTheory.v index f24e4b3..d812bf3 100644 --- a/theories/BasicMetaTheory.v +++ b/theories/BasicMetaTheory.v @@ -393,6 +393,30 @@ Proof. intros Γ. constructor. all: constructor. Qed. +Section star. + + Notation "⋆" := (lam mProp bot (var 0)). + + Lemma type_star : ∀ Γ, Γ ⊢ ⋆ : top. + Proof. + intro Γ. + apply type_lam. + all: eauto using scope_bot, type_bot. + - apply scope_var; reflexivity. + - eapply type_var; reflexivity. + Qed. + + Lemma scope_star {Γ : scope} : + scoping Γ ⋆ mProp. + Proof. + apply scope_lam. + - apply scope_bot. + - apply scope_var; reflexivity. + Qed. + +End star. + + (** Conversion entails mode equality **) Definition rscoping_comp (Γ : scope) ρ (Δ : scope) := From 1d22326f57b39c10c62441000a3e40e9e2cb07ef Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Wed, 26 Jun 2024 14:46:19 +0200 Subject: [PATCH 10/15] Admissible doen't needs Param --- theories/Admissible.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Admissible.v b/theories/Admissible.v index 2b6c7bf..c05f70e 100644 --- a/theories/Admissible.v +++ b/theories/Admissible.v @@ -11,7 +11,7 @@ From Equations Require Import Equations. From GhostTT.autosubst Require Import CCAST GAST core unscoped. From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl CScoping Scoping CTyping TermMode Typing BasicMetaTheory CCMetaTheory Erasure - Revival Param Model. + Revival. From GhostTT.reduction Require Import Injectivity Model. From Coq Require Import Setoid Morphisms Relation_Definitions. From 3f7f3afcc0fa520b32a05f2118a3ead0ad0724e8 Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Thu, 27 Jun 2024 14:22:35 +0200 Subject: [PATCH 11/15] split param --- _CoqProject | 14 +- theories/Examples.v | 2 +- theories/FreeTheorem.v | 2 +- theories/Model.v | 2 +- theories/Potential.v | 2 +- theories/param/Cast.v | 175 ++ theories/param/Conversion.v | 711 +++++++ theories/param/Param.v | 9 + theories/param/Renaming.v | 463 +++++ theories/param/Scope.v | 424 +++++ theories/param/Substitution.v | 561 ++++++ theories/param/Term.v | 322 ++++ theories/{Param.v => param/Typing.v} | 2568 +------------------------- 13 files changed, 2682 insertions(+), 2573 deletions(-) create mode 100644 theories/param/Cast.v create mode 100644 theories/param/Conversion.v create mode 100644 theories/param/Param.v create mode 100644 theories/param/Renaming.v create mode 100644 theories/param/Scope.v create mode 100644 theories/param/Substitution.v create mode 100644 theories/param/Term.v rename theories/{Param.v => param/Typing.v} (70%) diff --git a/_CoqProject b/_CoqProject index 539dbb3..f9da9d3 100644 --- a/_CoqProject +++ b/_CoqProject @@ -34,6 +34,8 @@ theories/reduction/multisteps/ReductionToCongruence.v theories/reduction/Injectivity.v theories/reduction/Model.v +theories/Admissible.v + theories/reduction/wrapping/Core.v theories/reduction/wrapping/Properties.v @@ -48,12 +50,18 @@ theories/CCMetaTheory.v theories/Erasure.v theories/Revival.v -theories/Param.v -theories/Model.v +theories/param/Term.v +theories/param/Scope.v +theories/param/Renaming.v +theories/param/Substitution.v +theories/param/Conversion.v +theories/param/Cast.v +theories/param/Typing.v +theories/param/Param.v -theories/Admissible.v +theories/Model.v theories/FreeTheorem.v theories/Examples.v diff --git a/theories/Examples.v b/theories/Examples.v index efb5335..eac9e3d 100644 --- a/theories/Examples.v +++ b/theories/Examples.v @@ -10,7 +10,7 @@ From Equations Require Import Equations. From GhostTT.autosubst Require Import CCAST GAST core unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl Scoping TermMode CastRemoval Typing BasicMetaTheory CScoping CTyping - CCMetaTheory Admissible Erasure Revival Param Model. + CCMetaTheory Admissible Erasure Revival param.Param Model. From Coq Require Import Setoid Morphisms Relation_Definitions. Import ListNotations. diff --git a/theories/FreeTheorem.v b/theories/FreeTheorem.v index 617a5c0..1017e55 100644 --- a/theories/FreeTheorem.v +++ b/theories/FreeTheorem.v @@ -16,7 +16,7 @@ From Equations Require Import Equations. From GhostTT.autosubst Require Import CCAST GAST core unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl Scoping TermMode CastRemoval Typing BasicMetaTheory CScoping CTyping - CCMetaTheory Admissible Erasure Revival Param Model. + CCMetaTheory Admissible Erasure Revival param.Param Model. From Coq Require Import Setoid Morphisms Relation_Definitions. Import ListNotations. diff --git a/theories/Model.v b/theories/Model.v index f518a3c..b71c2a8 100644 --- a/theories/Model.v +++ b/theories/Model.v @@ -5,7 +5,7 @@ From Equations Require Import Equations. From GhostTT.autosubst Require Import CCAST GAST core unscoped. From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl CScoping Scoping CTyping TermMode Typing BasicMetaTheory CCMetaTheory Erasure - Revival Param. + Revival param.Param. From GhostTT.reduction Require Export Injectivity Model. From Coq Require Import Setoid Morphisms Relation_Definitions. diff --git a/theories/Potential.v b/theories/Potential.v index 8ac0cda..a7be4c7 100644 --- a/theories/Potential.v +++ b/theories/Potential.v @@ -4,7 +4,7 @@ From Coq Require Import Utf8 List Bool Lia. From Equations Require Import Equations. From GhostTT.autosubst Require Import CCAST GAST core unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl - Scoping TermMode CastRemoval Typing BasicMetaTheory Param RTyping Admissible. + Scoping TermMode CastRemoval Typing BasicMetaTheory param.Param RTyping Admissible. From Coq Require Import Setoid Morphisms Relation_Definitions. Import ListNotations. diff --git a/theories/param/Cast.v b/theories/param/Cast.v new file mode 100644 index 0000000..fcf4d15 --- /dev/null +++ b/theories/param/Cast.v @@ -0,0 +1,175 @@ +From Coq Require Import Utf8 List Bool Lia. +From Equations Require Import Equations. +From GhostTT.autosubst Require Import CCAST GAST core unscoped. +From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl + CScoping Scoping CTyping TermMode Typing BasicMetaTheory CCMetaTheory Erasure + Revival. +From Coq Require Import Setoid Morphisms Relation_Definitions. +From GhostTT.param Require Export Term Scope. +From GhostTT.param Require Import Conversion. + +Import ListNotations. +Import CombineNotations. + +Set Default Goal Selector "!". +Set Equations Transparent. +(** Parametricity ignores casts **) + +Lemma ccong_pmPi : + ∀ Γ mx m Te Ae Ap Bp Te' Ae' Ap' Bp', + Γ ⊢ᶜ Te ≡ Te' → + Γ ⊢ᶜ Ae ≡ Ae' → + Γ ⊢ᶜ Ap ≡ Ap' → + let Γ' := + if isProp mx then + None :: Some (cProp, Ap) :: Γ + else if isKind mx then + Some (cType, capp (S ⋅ Ap) (cvar 0)) :: Some (cType, Ae) :: Γ + else + Some (cProp, capp (S ⋅ Ap) (cvar 0)) :: Some (cType, Ae) :: Γ + in + Γ' ⊢ᶜ Bp ≡ Bp' → + Γ ⊢ᶜ pmPi mx m Te Ae Ap Bp ≡ pmPi mx m Te' Ae' Ap' Bp'. +Proof. + cbn zeta. + intros Γ mx m Te Ae Ap Bp Te' Ae' Ap' Bp' hTe hAe hAp hBp. + unfold pmPi. destruct (isProp m) eqn:ep. + - unfold pmPiP. destruct_ifs. all: econv. + all: try apply crtyping_S. + apply cstyping_one_none. + - unfold pmPiNP. econv. + destruct (isProp mx) eqn:epx. all: econv. + all: try apply crtyping_S. + + eapply crtyping_shift. apply crtyping_S. + + eapply cstyping_one_none. + + destruct (isKind mx) eqn:ekx. + * { + eapply crtyping_shift_eq. + - eapply crtyping_shift. apply crtyping_S. + - f_equal. f_equal. f_equal. cbn. ssimpl. reflexivity. + } + * { + eapply crtyping_shift_eq. + - eapply crtyping_shift. apply crtyping_S. + - f_equal. f_equal. f_equal. cbn. ssimpl. reflexivity. + } +Qed. + +Hint Opaque pmPi : cc_conv. +Hint Resolve ccong_pmPi : cc_conv. + +Lemma meta_ctx_conv_conv : + ∀ Γ Δ u v, + Γ ⊢ᶜ u ≡ v → + Γ = Δ → + Δ ⊢ᶜ u ≡ v. +Proof. + intros Γ ? u v h ->. + assumption. +Qed. + +Lemma meta_ctx_param_conv : + ∀ sΓ Γp sΔ Δp u v, + Δp ⊢ᶜ ⟦ sΔ | u ⟧p ≡ ⟦ sΔ | v ⟧p → + Γp = Δp → + sΓ = sΔ → + Γp ⊢ᶜ ⟦ sΓ | u ⟧p ≡ ⟦ sΓ | v ⟧p. +Proof. + intros sΓ Γp sΔ Δp u v h -> ->. + assumption. +Qed. + +Hint Opaque plam : cc_conv. + +Lemma meta_ccscoping_conv : + ∀ Γ t m m', + ccscoping Γ t m → + m = m' → + ccscoping Γ t m'. +Proof. + intros Γ t m m' h ->. + assumption. +Qed. + +Lemma param_castrm : + ∀ Γ t m, + cscoping Γ t m → + ⟦ Γ ⟧p ⊢ᶜ ⟦ sc Γ | (ε|t|) ⟧p ≡ ⟦ sc Γ | t ⟧p. +Proof. + intros Δ t m h. + remember (sc Δ) as Γ eqn:e in *. + induction h in Δ, e |- *. + all: try solve [ econv ]. + all: try solve [ + cbn ; rewrite <- ?md_castrm ; + rewrite ?erase_castrm, ?revive_castrm ; + destruct_ifs ; econv + ]. + - cbn. rewrite !erase_castrm. econv. + subst. + eapply meta_ctx_conv_conv. + + eapply IHh2 with (Δ := Δ ,, (mx, ε|A|)). + reflexivity. + + cbn. rewrite !erase_castrm. reflexivity. + - cbn. rewrite !erase_castrm. + destruct_ifs. + + econstructor. 1: eauto. + eapply cconv_close. + eapply meta_ctx_conv_conv. + * eapply IHh2 with (Δ := Δ ,, (mx, ε|A|)). + subst. reflexivity. + * cbn. subst. rewrite !erase_castrm. rewrite e0. reflexivity. + + econv. + eapply meta_ctx_conv_conv. + * eapply IHh2 with (Δ := Δ ,, (mx, ε|A|)). + subst. reflexivity. + * cbn. subst. rewrite !erase_castrm. rewrite e0,e1. reflexivity. + + econv. + eapply meta_ctx_conv_conv. + * eapply IHh2 with (Δ := Δ ,, (mx, ε|A|)). + subst. reflexivity. + * cbn. subst. rewrite !erase_castrm. rewrite e0,e1. reflexivity. + - eapply cconv_trans. 1:{ eapply IHh6. eassumption. } + destruct (isKind m) eqn:ek. 1:{ mode_eqs. contradiction. } + subst. + apply cconv_irr. + + eapply param_scoping in h6. rewrite ek in h6. + rewrite <- csc_param_ctx in h6. assumption. + + rewrite csc_param_ctx. eapply meta_ccscoping_conv. + * eapply param_scoping. constructor. all: eassumption. + * rewrite ek. reflexivity. + - cbn. rewrite !erase_castrm, !revive_castrm. + destruct m. + + eapply ccong_pmifK. all: eauto. all: econv. + + eapply ccong_pmif. all: eauto. all: econv. + + eapply ccong_pmif. all: eauto. all: econv. + + econv. + - cbn. rewrite !erase_castrm, !revive_castrm. + destruct m. + + contradiction. + + econv. + + econv. + + econv. + - cbn. rewrite !erase_castrm, !revive_castrm. + destruct m. + + contradiction. + + econv. + + econv. + + econv. +Qed. + +Lemma param_castrm_conv : + ∀ Γ u v mu mv, + cscoping Γ u mu → + cscoping Γ v mv → + Γ ⊢ u ε≡ v → + ⟦ Γ ⟧p ⊢ᶜ ⟦ sc Γ | u ⟧p ≡ ⟦ sc Γ | v ⟧p. +Proof. + intros Γ u v mu mv hu hv h. + eapply param_conv in h. + eapply cconv_trans. + - apply cconv_sym. eapply param_castrm. eassumption. + - eapply cconv_trans. 1: eassumption. + eapply param_castrm. eassumption. +Qed. + diff --git a/theories/param/Conversion.v b/theories/param/Conversion.v new file mode 100644 index 0000000..7f3e576 --- /dev/null +++ b/theories/param/Conversion.v @@ -0,0 +1,711 @@ +From Coq Require Import Utf8 List Bool Lia. +From Equations Require Import Equations. +From GhostTT.autosubst Require Import CCAST GAST core unscoped. +From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl + CScoping Scoping CTyping TermMode Typing BasicMetaTheory CCMetaTheory Erasure + Revival. +From Coq Require Import Setoid Morphisms Relation_Definitions. +From GhostTT.param Require Export Term Scope. +From GhostTT.param Require Import Renaming Substitution. + +Import ListNotations. +Import CombineNotations. + +Set Default Goal Selector "!". +Set Equations Transparent. + + +(** Parametricity preserves conversion **) + +Lemma vreg_vpar_dec : + ∀ n, { n = vpar (Nat.div2 n) } + { n = vreg (Nat.div2 n) }. +Proof. + intros n. + destruct (PeanoNat.Nat.Even_Odd_dec n). + - left. unfold vpar. + etransitivity. + + apply PeanoNat.Nat.Even_double. assumption. + + unfold Nat.double. lia. + - right. unfold vreg. + etransitivity. + + apply PeanoNat.Nat.Odd_double. assumption. + + unfold Nat.double. lia. +Qed. + +Lemma ccong_pPi : + ∀ Γ mx A B C A' B' C', + Γ ⊢ᶜ A ≡ A' → + Γ ⊢ᶜ B ≡ B' → + Some (mx, capp (S ⋅ B) (cvar 0)) :: Some (cType, A) :: Γ ⊢ᶜ C ≡ C' → + Γ ⊢ᶜ pPi mx A B C ≡ pPi mx A' B' C'. +Proof. + intros Γ mx A B C A' B' C' hA hB hC. + unfold pPi. + econv. + eapply cconv_ren. 2: eassumption. + apply crtyping_S. +Qed. + +Hint Resolve ccong_pPi : cc_conv. + +Lemma ccong_plam : + ∀ Γ mp A B t A' B' t', + Γ ⊢ᶜ A ≡ A' → + Γ ⊢ᶜ B ≡ B' → + Some (mp, capp (S ⋅ B) (cvar 0)) :: Some (cType, A) :: Γ ⊢ᶜ t ≡ t' → + Γ ⊢ᶜ plam mp A B t ≡ plam mp A' B' t'. +Proof. + intros Γ mp A B t A' B' t' hA hB ht. + econv. + eapply cconv_ren. 2: eassumption. + apply crtyping_S. +Qed. + +Hint Resolve ccong_plam : cc_conv. + +Transparent epm_lift rpm_lift. + +Lemma ccong_epm_lift : + ∀ Γ Γ' t t', + ⟦ Γ ⟧ε ⊢ᶜ t ≡ t' → + Γ' = ⟦ Γ ⟧p → + Γ' ⊢ᶜ epm_lift t ≡ epm_lift t'. +Proof. + intros Γ Γ' t t' ht ->. + unfold epm_lift. eapply cconv_ren. + - apply typing_er_sub_param. + - assumption. +Qed. + +Lemma ccong_rpm_lift : + ∀ Γ Γ' t t', + ⟦ Γ ⟧v ⊢ᶜ t ≡ t' → + Γ' = ⟦ Γ ⟧p → + Γ' ⊢ᶜ rpm_lift t ≡ rpm_lift t'. +Proof. + intros Γ Γ' t t' ht ->. + unfold rpm_lift. eapply cconv_ren. + - apply typing_rev_sub_param. + - assumption. +Qed. + +(* Need to have them opaque so that they can't be confused. *) +Hint Opaque epm_lift rpm_lift : cc_scope cc_conv cc_type. +Hint Resolve ccong_epm_lift ccong_rpm_lift : cc_conv. + +Opaque epm_lift rpm_lift. + +Hint Resolve cconv_ren cconv_subst : cc_conv. + +Lemma erase_conv_eq : + ∀ Γ Γ' Γ'' u v, + Γ ⊢ u ≡ v → + Γ' = ⟦ Γ ⟧ε → + Γ'' = sc Γ → + Γ' ⊢ᶜ ⟦ Γ'' | u ⟧ε ≡ ⟦ Γ'' | v ⟧ε. +Proof. + intros Γ Γ' Γ'' u v h -> ->. + apply erase_conv. assumption. +Qed. + +Hint Resolve erase_conv_eq : cc_conv. + +Lemma revive_conv_eq : + ∀ Γ Γ' Γ'' u v, + Γ ⊢ u ≡ v → + Γ' = ⟦ Γ ⟧v → + Γ'' = sc Γ → + Γ' ⊢ᶜ ⟦ Γ'' | u ⟧v ≡ ⟦ Γ'' | v ⟧v. +Proof. + intros Γ Γ' Γ'' u v h -> ->. + apply revive_conv. assumption. +Qed. + +Hint Resolve revive_conv_eq : cc_conv. + +Lemma crtyping_shift_eq : + ∀ Γ Δ Ξ mx A ρ, + crtyping Γ ρ Δ → + Ξ = Some (mx, ρ ⋅ A) :: Γ → + crtyping Ξ (up_ren ρ) (Some (mx, A) :: Δ). +Proof. + intros Γ Δ Ξ mx A ρ hρ ->. + apply crtyping_shift. assumption. +Qed. + +Lemma csc_param_ctx : + ∀ Γ, + csc ⟦ Γ ⟧p = param_sc (sc Γ). +Proof. + intros Γ. + induction Γ as [| [m A] Γ ih]. + - cbn. reflexivity. + - cbn. destruct_ifs. all: cbn. + + f_equal. f_equal. apply ih. + + f_equal. f_equal. apply ih. + + f_equal. f_equal. apply ih. +Qed. + +Lemma ccong_perif : + ∀ Γ be be' Pe Pe' te te' fe fe', + Γ ⊢ᶜ be ≡ be' → + Γ ⊢ᶜ Pe ≡ Pe' → + Γ ⊢ᶜ te ≡ te' → + Γ ⊢ᶜ fe ≡ fe' → + Γ ⊢ᶜ perif be Pe te fe ≡ perif be' Pe' te' fe'. +Proof. + intros Γ be be' Pe Pe' te te' fe fe' hb hP ht hf. + unfold perif. econv. apply crtyping_S. +Qed. + +Hint Opaque perif : cc_conv. +Hint Resolve ccong_perif : cc_conv. + +Lemma ccong_pmif : + ∀ Γ bP Pe PP te tP fe fP bP' Pe' PP' te' tP' fe' fP', + Γ ⊢ᶜ bP ≡ bP' → + Γ ⊢ᶜ Pe ≡ Pe' → + Γ ⊢ᶜ PP ≡ PP' → + Γ ⊢ᶜ te ≡ te' → + Γ ⊢ᶜ tP ≡ tP' → + Γ ⊢ᶜ fe ≡ fe' → + Γ ⊢ᶜ fP ≡ fP' → + Γ ⊢ᶜ pmif bP Pe PP te tP fe fP ≡ pmif bP' Pe' PP' te' tP' fe' fP'. +Proof. + intros Γ bP Pe PP te tP fe fP bP' Pe' PP' te' tP' fe' fP'. + intros hbP hPe hPP hte htP hfe hfP. + unfold pmif. econv. all: apply crtyping_S. +Qed. + +Hint Opaque pmif : cc_conv. +Hint Resolve ccong_pmif : cc_conv. + +Lemma ccong_pmifK : + ∀ Γ be bP Pe PP te tP fe fP be' bP' Pe' PP' te' tP' fe' fP', + Γ ⊢ᶜ be ≡ be' → + Γ ⊢ᶜ bP ≡ bP' → + Γ ⊢ᶜ Pe ≡ Pe' → + Γ ⊢ᶜ PP ≡ PP' → + Γ ⊢ᶜ te ≡ te' → + Γ ⊢ᶜ tP ≡ tP' → + Γ ⊢ᶜ fe ≡ fe' → + Γ ⊢ᶜ fP ≡ fP' → + Γ ⊢ᶜ pmifK be bP Pe PP te tP fe fP ≡ pmifK be' bP' Pe' PP' te' tP' fe' fP'. +Proof. + intros Γ be bP Pe PP te tP fe fP be' bP' Pe' PP' te' tP' fe' fP'. + intros hbe hbP hPe hPP hte htP hfe hfP. + unfold pmifK. econv. all: apply crtyping_S. +Qed. + +Hint Opaque pmifK : cc_conv. +Hint Resolve ccong_pmifK : cc_conv. + +Lemma pren_S : + ∀ n, pren S n = S (S n). +Proof. + intro n. + change (pren S n) with (pren (id >> S) n). + rewrite pren_comp_S. rewrite pren_id. reflexivity. +Qed. + +Lemma pren_S_pw : + pointwise_relation _ eq (pren S) (S >> S). +Proof. + intro n. apply pren_S. +Qed. + +Lemma param_conv : + ∀ Γ u v, + Γ ⊢ u ≡ v → + ⟦ Γ ⟧p ⊢ᶜ ⟦ sc Γ | u ⟧p ≡ ⟦ sc Γ | v ⟧p. +Proof. + intros Γ u v h. + induction h. + (* all: try solve [ cbn ; destruct_ifs ; eauto with cc_conv ]. *) + - cbn. + erewrite scoping_md. 2: eassumption. + destruct_ifs. all: mode_eqs. all: try discriminate. + + eapply cconv_trans. + 1:{ constructor. 2: apply cconv_refl. constructor. } + cbn. + eapply cconv_trans. 1: constructor. + ssimpl. apply ccmeta_refl. + erewrite param_subst. + 2:{ eapply sscoping_one. eassumption. } + 2: eapply sscoping_comp_one. + ssimpl. eapply ext_cterm_scoped. + 1:{ eapply param_scoping. eassumption. } + (* The following we do basically four times, but we don't know how + to factorise. + *) + intros [| []] hx. all: cbn. 1,2: reflexivity. + unfold inscope in hx. cbn in hx. + unfold psubst. rewrite div2_SS. cbn. + destruct (vreg_vpar_dec n) as [en | en]. + * rewrite en in hx. rewrite nth_error_param_vpar in hx. + destruct nth_error as [mx|] eqn:e1. 2: discriminate. + cbn in hx. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_succ. + destruct PeanoNat.Nat.odd eqn:eodd. + 1:{ rewrite en in eodd. rewrite odd_vpar in eodd. discriminate. } + destruct (isProp mx) eqn:e2. 1: discriminate. + destruct (isKind mx) eqn:e3. all: mode_eqs. + -- cbn. f_equal. assumption. + -- destruct mx. all: try discriminate. + ++ cbn. f_equal. assumption. + ++ cbn. f_equal. assumption. + * set (p := Nat.div2 n) in *. + rewrite en in hx. rewrite nth_error_param_vreg in hx. + destruct nth_error as [mx|] eqn:e1. 2: discriminate. + cbn in hx. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_succ. + destruct PeanoNat.Nat.odd eqn:eodd. + 2:{ rewrite en in eodd. rewrite odd_vreg in eodd. discriminate. } + destruct (isProp mx) eqn:e2. + -- mode_eqs. cbn. f_equal. assumption. + -- unfold relv, ghv. rewrite e1. + destruct_ifs. + ++ rewrite en. reflexivity. + ++ rewrite en. reflexivity. + ++ destruct mx. all: discriminate. + + destruct (isType mx) eqn:e2. 2:{ destruct mx. all: discriminate. } + mode_eqs. + eapply cconv_trans. + 1:{ constructor. 2: apply cconv_refl. constructor. } + cbn. + eapply cconv_trans. 1: constructor. + ssimpl. apply ccmeta_refl. + erewrite param_subst. + 2:{ eapply sscoping_one. eassumption. } + 2: eapply sscoping_comp_one. + ssimpl. eapply ext_cterm_scoped. + 1:{ eapply param_scoping. eassumption. } + (* Basically same as above, is there a nice lemma to state? *) + intros [| []] hx. all: cbn. 1,2: reflexivity. + unfold inscope in hx. cbn in hx. + unfold psubst. rewrite div2_SS. cbn. + destruct (vreg_vpar_dec n) as [en | en]. + * set (p := Nat.div2 n) in *. + rewrite en in hx. rewrite nth_error_param_vpar in hx. + destruct nth_error as [mx|] eqn:emx. 2: discriminate. + cbn in hx. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_succ. + destruct PeanoNat.Nat.odd eqn:eodd. + 1:{ rewrite en in eodd. rewrite odd_vpar in eodd. discriminate. } + destruct (isProp mx) eqn:e2. 1: discriminate. + destruct (isKind mx) eqn:e3. all: mode_eqs. + -- cbn. f_equal. assumption. + -- destruct mx. all: try discriminate. + ++ cbn. f_equal. assumption. + ++ cbn. f_equal. assumption. + * set (p := Nat.div2 n) in *. + rewrite en in hx. rewrite nth_error_param_vreg in hx. + destruct nth_error as [mx|] eqn:emx. 2: discriminate. + cbn in hx. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_succ. + destruct PeanoNat.Nat.odd eqn:eodd. + 2:{ rewrite en in eodd. rewrite odd_vreg in eodd. discriminate. } + destruct (isProp mx) eqn:e2. + -- mode_eqs. cbn. f_equal. assumption. + -- unfold relv, ghv. rewrite emx. + destruct_ifs. + ++ rewrite en. reflexivity. + ++ rewrite en. reflexivity. + ++ destruct mx. all: discriminate. + + eapply cconv_trans. + 1:{ constructor. 2: apply cconv_refl. constructor. } + cbn. + eapply cconv_trans. 1: constructor. + ssimpl. apply ccmeta_refl. + erewrite param_subst. + 2:{ eapply sscoping_one. eassumption. } + 2: eapply sscoping_comp_one. + ssimpl. eapply ext_cterm_scoped. + 1:{ eapply param_scoping. eassumption. } + (* Basically same as above, is there a nice lemma to state? *) + intros [| []] hx. all: cbn. 1,2: reflexivity. + unfold inscope in hx. cbn in hx. + unfold psubst. rewrite div2_SS. cbn. + destruct (vreg_vpar_dec n) as [en | en]. + * set (p := Nat.div2 n) in *. + rewrite en in hx. rewrite nth_error_param_vpar in hx. + destruct nth_error as [mx|] eqn:emx. 2: discriminate. + cbn in hx. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_succ. + destruct PeanoNat.Nat.odd eqn:eodd. + 1:{ rewrite en in eodd. rewrite odd_vpar in eodd. discriminate. } + destruct (isProp mx) eqn:e3. 1: discriminate. + destruct (isKind mx) eqn:e4. all: mode_eqs. + -- cbn. f_equal. assumption. + -- destruct mx. all: try discriminate. + ++ cbn. f_equal. assumption. + ++ cbn. f_equal. assumption. + * set (p := Nat.div2 n) in *. + rewrite en in hx. rewrite nth_error_param_vreg in hx. + destruct nth_error as [mx|] eqn:emx. 2: discriminate. + cbn in hx. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_succ. + destruct PeanoNat.Nat.odd eqn:eodd. + 2:{ rewrite en in eodd. rewrite odd_vreg in eodd. discriminate. } + destruct (isProp mx) eqn:e3. + -- mode_eqs. cbn. f_equal. assumption. + -- unfold relv, ghv. rewrite emx. + destruct_ifs. + ++ rewrite en. reflexivity. + ++ rewrite en. reflexivity. + ++ destruct mx. all: discriminate. + + eapply cconv_trans. 1: constructor. + unfold close. ssimpl. apply ccmeta_refl. + erewrite param_subst. + 2:{ eapply sscoping_one. eassumption. } + 2: eapply sscoping_comp_one. + ssimpl. eapply ext_cterm_scoped. + 1:{ eapply param_scoping. eassumption. } + (* Basically same as above, is there a nice lemma to state? *) + intros [| []] hx. all: cbn. 1,2: reflexivity. + unfold inscope in hx. cbn in hx. + unfold psubst. rewrite div2_SS. cbn. + destruct (vreg_vpar_dec n) as [en | en]. + * set (p := Nat.div2 n) in *. + rewrite en in hx. rewrite nth_error_param_vpar in hx. + destruct nth_error as [mx|] eqn:emx. 2: discriminate. + cbn in hx. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_succ. + destruct PeanoNat.Nat.odd eqn:eodd. + 1:{ rewrite en in eodd. rewrite odd_vpar in eodd. discriminate. } + destruct (isProp mx) eqn:e3. 1: discriminate. + destruct (isKind mx) eqn:e4. all: mode_eqs. + -- cbn. f_equal. assumption. + -- destruct mx. all: try discriminate. + ++ cbn. f_equal. assumption. + ++ cbn. f_equal. assumption. + * set (p := Nat.div2 n) in *. + rewrite en in hx. rewrite nth_error_param_vreg in hx. + destruct nth_error as [mx|] eqn:emx. 2: discriminate. + cbn in hx. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_succ. + destruct PeanoNat.Nat.odd eqn:eodd. + 2:{ rewrite en in eodd. rewrite odd_vreg in eodd. discriminate. } + destruct (isProp mx) eqn:e3. + -- mode_eqs. cbn. f_equal. assumption. + -- unfold relv, ghv. rewrite emx. + destruct_ifs. + ++ rewrite en. reflexivity. + ++ rewrite en. reflexivity. + ++ destruct mx. all: discriminate. + + destruct mx. all: discriminate. + - cbn. + erewrite scoping_md. 2: eassumption. + erewrite scoping_md. 2: eassumption. + destruct_if e. + 1:{ + destruct H2 as [emp | [ emp | ]]. 3: contradiction. + all: subst. all: discriminate. + } + cbn. apply cconv_refl. + - cbn. destruct m. + + unfold pmifK. change (epm_lift etrue) with etrue. + eapply cconv_trans. + 1:{ + econstructor. 2: econv. + constructor. + } + eapply cconv_trans. 1: constructor. + ssimpl. econv. + + unfold pmif. constructor. + + unfold pmif. constructor. + + constructor. + - cbn. destruct m. + + unfold pmifK. change (epm_lift etrue) with etrue. + eapply cconv_trans. + 1:{ + econstructor. 2: econv. + constructor. + } + eapply cconv_trans. 1: constructor. + ssimpl. econv. + + unfold pmif. constructor. + + unfold pmif. constructor. + + constructor. + - cbn. eapply param_scoping in H0, H1, H2. + rewrite <- csc_param_ctx in H0, H1, H2. + destruct m. + + contradiction. + + cbn in *. + eapply cconv_irr. + * escope. all: eauto using csc_param_ctx. + eapply scoping_epm_lift. 1: escope. apply csc_param_ctx. + * assumption. + + cbn in *. + eapply cconv_irr. + * escope. all: eauto using csc_param_ctx. + eapply scoping_epm_lift. 1: escope. apply csc_param_ctx. + * assumption. + + cbn in *. + eapply cconv_irr. + * escope. 2: eauto using csc_param_ctx. + eapply scoping_epm_lift. 1: escope. apply csc_param_ctx. + * assumption. + - cbn. + remd. + eapply param_scoping in H0, H1, H2, H3. + rewrite <- csc_param_ctx in H0, H1, H2, H3. + destruct m. + + contradiction. + + cbn in *. + eapply cconv_irr. + * escope. all: eauto using csc_param_ctx. + eapply scoping_epm_lift. 1: escope. 1: reflexivity. + apply csc_param_ctx. + * escope. all: eauto using csc_param_ctx. + eapply scoping_epm_lift. 1: escope. + all: try reflexivity. + apply csc_param_ctx. + + cbn in *. + eapply cconv_irr. + * escope. all: eauto using csc_param_ctx. + eapply scoping_epm_lift. 1: escope. 1: reflexivity. + apply csc_param_ctx. + * { + escope. all: eauto using csc_param_ctx. + eapply scoping_rpm_lift. 2: apply csc_param_ctx. + econstructor. + 1,2: eapply scoping_to_rev. + all: escope. all: reflexivity. + } + + cbn in *. + eapply cconv_irr. + * escope. 2: eauto using csc_param_ctx. + eapply scoping_epm_lift. 1: escope. 1: reflexivity. + apply csc_param_ctx. + * escope. all: eauto using csc_param_ctx. + - cbn. eapply param_scoping in H0, H1, H2, H3, H4. + rewrite <- csc_param_ctx in H0, H1, H2, H3, H4. + destruct m. + + contradiction. + + cbn in *. eapply cconv_irr. + * { + escope. all: eauto using csc_param_ctx. + - change (rpm_lift ?t) with t. escope. + - change (epm_lift ?t) with (vreg ⋅ t). cbn. + change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + escope. apply csc_param_ctx. + } + * auto. + + cbn in *. eapply cconv_irr. + * { + escope. all: eauto using csc_param_ctx. + - change (rpm_lift ?t) with t. escope. + - change (epm_lift ?t) with (vreg ⋅ t). cbn. + change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + escope. apply csc_param_ctx. + } + * auto. + + cbn in *. eapply cconv_irr. + * { + escope. all: eauto using csc_param_ctx. + - change (rpm_lift ?t) with t. escope. + - change (epm_lift ?t) with (vreg ⋅ t). cbn. + change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + escope. apply csc_param_ctx. + } + * auto. + - simpl. remd. + eapply param_scoping in H0, H1, H2, H3, H4, H5, H6, H7. + rewrite <- csc_param_ctx in H0, H1, H2, H3, H4, H5, H6, H7. + destruct m. + + contradiction. + + cbn in *. eapply cconv_irr. + * { + escope. all: eauto using csc_param_ctx. + change (epm_lift ?t) with (vreg ⋅ t). cbn. + change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + escope. all: apply csc_param_ctx. + } + * change (epm_lift ?t) with (vreg ⋅ t). cbn. + change (rpm_lift ?t) with (vreg ⋅ t). cbn. + erewrite !erase_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. + erewrite !param_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. + change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + change (vreg ⋅ ⟦ ?G | ?t ⟧v) with (⟦ G | t ⟧pv). + ssimpl. rewrite pren_S_pw. ssimpl. + rewrite <- !rinstInst'_cterm. + change (S >> vreg) with (vreg >> S >> S). + rewrite <- !funcomp_assoc. + change (S >> vreg) with (vreg >> S >> S). + rewrite !funcomp_assoc. + rewrite <- !renRen_cterm. + change (ren_cterm vreg ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + escope. all: cbn. all: eauto using csc_param_ctx. + + cbn in *. eapply cconv_irr. + * { + escope. all: eauto using csc_param_ctx. + change (epm_lift ?t) with (vreg ⋅ t). cbn. + change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + escope. all: apply csc_param_ctx. + } + * change (epm_lift ?t) with (vreg ⋅ t). cbn. + change (rpm_lift ?t) with (vreg ⋅ t). cbn. + erewrite !erase_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. + erewrite !param_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. + change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + change (vreg ⋅ ⟦ ?G | ?t ⟧v) with (⟦ G | t ⟧pv). + ssimpl. rewrite pren_S_pw. ssimpl. + rewrite <- !rinstInst'_cterm. + change (S >> vreg) with (vreg >> S >> S). + rewrite <- !funcomp_assoc. + change (S >> vreg) with (vreg >> S >> S). + rewrite !funcomp_assoc. + rewrite <- !renRen_cterm. + change (ren_cterm vreg ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + change (ren_cterm vreg ⟦ ?G | ?t ⟧v) with (⟦ G | t ⟧pv). + escope. all: cbn. all: eauto using csc_param_ctx. + + cbn in *. eapply cconv_irr. + * { + escope. all: eauto using csc_param_ctx. + change (epm_lift ?t) with (vreg ⋅ t). cbn. + change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + escope. all: apply csc_param_ctx. + } + * change (epm_lift ?t) with (vreg ⋅ t). cbn. + change (rpm_lift ?t) with (vreg ⋅ t). cbn. + erewrite !erase_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. + erewrite !param_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. + change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + change (vreg ⋅ ⟦ ?G | ?t ⟧v) with (⟦ G | t ⟧pv). + ssimpl. rewrite pren_S_pw. ssimpl. + rewrite <- !rinstInst'_cterm. + change (S >> vreg) with (vreg >> S >> S). + rewrite <- !funcomp_assoc. + change (S >> vreg) with (vreg >> S >> S). + rewrite !funcomp_assoc. + rewrite <- !renRen_cterm. + change (ren_cterm vreg ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). + change (ren_cterm vreg ⟦ ?G | ?t ⟧v) with (⟦ G | t ⟧pv). + escope. all: cbn. all: eauto using csc_param_ctx. + - cbn. apply cconv_refl. + - cbn. + destruct m, mx. all: simpl. + + econv. all: try reflexivity. + all: eauto using crtyping_S. + cbn. eapply crtyping_shift_eq. + * apply crtyping_shift. apply crtyping_S. + * cbn. f_equal. ssimpl. reflexivity. + + econv. all: try reflexivity. + all: eauto using crtyping_S. + cbn. eapply crtyping_shift_eq. + * apply crtyping_shift. apply crtyping_S. + * cbn. f_equal. ssimpl. reflexivity. + + econv. all: try reflexivity. + all: eauto using crtyping_S, cstyping_one_none. + cbn. eapply crtyping_shift_eq. + * apply crtyping_shift. apply crtyping_S. + * cbn. f_equal. ssimpl. reflexivity. + + econv. all: try reflexivity. + all: eauto using crtyping_S, cstyping_one_none. + * cbn. apply crtyping_shift. apply crtyping_S. + * eapply cstyping_one_none. + + econv. all: try reflexivity. + all: eauto using crtyping_S. + cbn. eapply crtyping_shift_eq. + * apply crtyping_shift. apply crtyping_S. + * cbn. f_equal. ssimpl. reflexivity. + + econv. all: try reflexivity. + all: eauto using crtyping_S. + cbn. eapply crtyping_shift_eq. + * apply crtyping_shift. apply crtyping_S. + * cbn. f_equal. ssimpl. reflexivity. + + econv. all: try reflexivity. + all: eauto using crtyping_S, cstyping_one_none. + cbn. eapply crtyping_shift_eq. + * apply crtyping_shift. apply crtyping_S. + * cbn. f_equal. ssimpl. reflexivity. + + econv. all: try reflexivity. + all: eauto using crtyping_S, cstyping_one_none. + * apply crtyping_shift. apply crtyping_S. + * apply cstyping_one_none. + + econv. all: try reflexivity. + all: eauto using crtyping_S. + cbn. eapply crtyping_shift_eq. + * apply crtyping_shift. apply crtyping_S. + * cbn. f_equal. ssimpl. reflexivity. + + econv. all: try reflexivity. + all: eauto using crtyping_S. + cbn. eapply crtyping_shift_eq. + * apply crtyping_shift. apply crtyping_S. + * cbn. f_equal. ssimpl. reflexivity. + + econv. all: try reflexivity. + all: eauto using crtyping_S, cstyping_one_none. + * eapply cstyping_nones. + * cbn. eapply crtyping_shift_eq. + -- apply crtyping_shift. apply crtyping_S. + -- cbn. f_equal. ssimpl. reflexivity. + + econv. all: try reflexivity. + all: eauto using crtyping_S, cstyping_one_none. + * apply crtyping_shift. apply crtyping_S. + * apply cstyping_one_none. + + econv. all: try reflexivity. + eauto using crtyping_S. + + econv. all: try reflexivity. + eauto using crtyping_S. + + econv. all: try reflexivity. + eauto using crtyping_S. + + econv. apply cstyping_one_none. + - cbn in *. destruct_ifs. + + econv. apply cstyping_one_none. + + econv. all: try reflexivity. + apply crtyping_S. + + econv. all: try reflexivity. + apply crtyping_S. + - cbn. + eapply conv_md in h2 as e2. rewrite <- e2. + destruct_ifs. + + econv. all: reflexivity. + + econv. all: reflexivity. + + econv. + - cbn. + eapply conv_md in h as e. rewrite <- e. + destruct_ifs. all: econv. + - cbn. + eapply conv_md in h as e. rewrite <- e. + destruct_ifs. all: econv. + - cbn. + eapply conv_md in h3 as e3. rewrite <- e3. + destruct_ifs. all: econv. all: reflexivity. + - cbn. + eapply conv_md in h2 as e2. rewrite <- e2. + destruct_ifs. all: econv. all: reflexivity. + - cbn. + econv. all: reflexivity. + - cbn. destruct m. + + econv. all: reflexivity. + + econv. all: reflexivity. + + econv. all: reflexivity. + + econv. + - cbn. econv. + - cbn. destruct m. + all: econv. all: reflexivity. + - cbn. econv. all: reflexivity. + - cbn. econv. + - cbn. econv. + - cbn. destruct m. + + econv. + + econv. all: reflexivity. + + econv. all: reflexivity. + + econv. all: reflexivity. + - cbn. + destruct_ifs. all: econv. all: reflexivity. + - econv. + - apply cconv_sym. assumption. + - eapply cconv_trans. all: eauto. + - eapply param_scoping in H, H0. cbn in *. + apply cconv_irr. all: rewrite csc_param_ctx. all: assumption. +Qed. diff --git a/theories/param/Param.v b/theories/param/Param.v new file mode 100644 index 0000000..4c2176c --- /dev/null +++ b/theories/param/Param.v @@ -0,0 +1,9 @@ + +From GhostTT.param Require Export Term. +From GhostTT.param Require Export Scope. +From GhostTT.param Require Export Renaming. +From GhostTT.param Require Export Substitution. +From GhostTT.param Require Export Conversion. +From GhostTT.param Require Export Cast. +From GhostTT.param Require Export Typing. + diff --git a/theories/param/Renaming.v b/theories/param/Renaming.v new file mode 100644 index 0000000..81ae506 --- /dev/null +++ b/theories/param/Renaming.v @@ -0,0 +1,463 @@ +From Coq Require Import Utf8 List Bool Lia. +From Equations Require Import Equations. +From GhostTT.autosubst Require Import CCAST GAST core unscoped. +From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl + CScoping Scoping CTyping TermMode Typing BasicMetaTheory CCMetaTheory Erasure + Revival. +From Coq Require Import Setoid Morphisms Relation_Definitions. +From GhostTT.param Require Export Term. + +Import ListNotations. +Import CombineNotations. + +Set Default Goal Selector "!". +Set Equations Transparent. + +(** Parametricity commutes with renaming + + For this we define pren ρ which is basically the following function: + pren ρ (2 * n) = 2 * (ρ n) + pren ρ (2 * n + 1) = 2 * (ρ n) + 1 + +**) + +Definition pren (ρ : nat → nat) := + λ n, PeanoNat.Nat.b2n (Nat.odd n) + 2 * ρ (Nat.div2 n). + +Lemma pren_even : + ∀ ρ n, + pren ρ (n * 2) = (ρ n) * 2. +Proof. + intros ρ n. + unfold pren. + replace (n * 2) with (2 * n) by lia. + rewrite PeanoNat.Nat.div2_double. + rewrite PeanoNat.Nat.odd_mul. cbn. lia. +Qed. + +Lemma pren_odd : + ∀ ρ n, + pren ρ (S (n * 2)) = S ((ρ n) * 2). +Proof. + intros ρ n. + unfold pren. + replace (n * 2) with (2 * n) by lia. + rewrite PeanoNat.Nat.div2_succ_double. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_mul. cbn. lia. +Qed. + +Lemma div2_SS : + ∀ n, Nat.div2 (S (S n)) = S (Nat.div2 n). +Proof. + intro n. + destruct (PeanoNat.Nat.Even_Odd_dec n) as [h | h]. + - rewrite <- PeanoNat.Nat.Odd_div2. + 2:{ apply PeanoNat.Nat.Odd_succ. assumption. } + rewrite <- PeanoNat.Nat.Even_div2. 2: assumption. + reflexivity. + - rewrite <- PeanoNat.Nat.Even_div2. + 2:{ apply PeanoNat.Nat.Even_succ. assumption. } + rewrite <- PeanoNat.Nat.Odd_div2. 2: assumption. + reflexivity. +Qed. + +Lemma pren_SS : + ∀ ρ n, pren ρ (S (S n)) = pren (S >> ρ) n. +Proof. + intros ρ n. + unfold pren. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_succ. + rewrite div2_SS. reflexivity. +Qed. + +Lemma pren_comp_S : + ∀ ρ n, pren (ρ >> S) n = S (S (pren ρ n)). +Proof. + intros ρ n. + unfold pren. ssimpl. lia. +Qed. + +Lemma pren_id : + ∀ n, pren id n = n. +Proof. + intros n. + unfold pren. + rewrite PeanoNat.Nat.div2_div. + symmetry. etransitivity. 1: eapply PeanoNat.Nat.div_mod_eq with (y := 2). + rewrite <- PeanoNat.Nat.bit0_mod. + rewrite PeanoNat.Nat.bit0_odd. + unfold id. unfold Datatypes.id. lia. +Qed. + +Transparent epm_lift rpm_lift. + +Lemma pren_epm_lift : + ∀ ρ t, + pren ρ ⋅ epm_lift t = epm_lift (ρ ⋅ t). +Proof. + intros ρ t. + unfold epm_lift. ssimpl. + eapply extRen_cterm. intro x. + unfold vreg, pren. ssimpl. + replace (x * 2) with (2 * x) by lia. + rewrite PeanoNat.Nat.div2_succ_double. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_mul. cbn. lia. +Qed. + +Lemma pren_rpm_lift : + ∀ ρ t, + pren ρ ⋅ rpm_lift t = rpm_lift (ρ ⋅ t). +Proof. + intros ρ t. + apply pren_epm_lift. +Qed. + +Opaque epm_lift rpm_lift. + +Ltac cEl_ren := + change (cEl (?ρ ⋅ ?t)) with (ρ ⋅ (cEl t)). + +Lemma param_ren : + ∀ Γ Δ ρ t, + rscoping Γ ρ Δ → + rscoping_comp Γ ρ Δ → + ⟦ Γ | ρ ⋅ t ⟧p = (pren ρ) ⋅ ⟦ Δ | t ⟧p. +Proof. + intros Γ Δ ρ t hρ hcρ. + induction t in Γ, Δ, ρ, hρ, hcρ |- *. + - cbn. + destruct (nth_error Δ n) eqn:e. + + eapply hρ in e as e'. rewrite e'. + destruct_if e1. + * unfold vreg, pren. ssimpl. f_equal. + replace (n * 2) with (2 * n) by lia. + rewrite PeanoNat.Nat.div2_succ_double. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_mul. cbn. lia. + * unfold pren, vpar. ssimpl. f_equal. + replace (n * 2) with (2 * n) by lia. + rewrite PeanoNat.Nat.div2_double. + rewrite PeanoNat.Nat.odd_mul. cbn. lia. + + eapply hcρ in e as e'. rewrite e'. reflexivity. + - cbn. destruct_ifs. all: reflexivity. + - cbn. unfold pmPi, pmPiP, pmPiNP, pPi. + erewrite IHt1. 2,3: eassumption. + erewrite IHt2. + 2:{ eapply rscoping_upren. eassumption. } + 2:{ eapply rscoping_comp_upren. eassumption. } + erewrite ?erase_ren, ?revive_ren. + 2-5: eauto using rscoping_upren, rscoping_comp_upren. + rewrite ?pren_epm_lift, ?pren_rpm_lift. + destruct m, m0. all: cbn in *. (* all: try reflexivity. *) + + f_equal. all: f_equal. + 1:{ rewrite pren_epm_lift. cbn. reflexivity. } + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. all: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ rewrite pren_epm_lift. cbn. reflexivity. } + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. all: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ + rewrite pren_epm_lift. cbn. f_equal. + unfold close. ssimpl. reflexivity. + } + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. all: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ + rewrite pren_epm_lift. cbn. f_equal. + unfold close. ssimpl. reflexivity. + } + 1:{ ssimpl. reflexivity. } + f_equal. + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ rewrite pren_epm_lift. cbn. reflexivity. } + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. all: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ rewrite pren_epm_lift. cbn. reflexivity. } + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. all: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ + rewrite pren_epm_lift. cbn. f_equal. + unfold close. ssimpl. reflexivity. + } + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. all: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ + rewrite pren_epm_lift. cbn. f_equal. + unfold close. ssimpl. reflexivity. + } + 1:{ ssimpl. reflexivity. } + f_equal. + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ rewrite pren_epm_lift. cbn. reflexivity. } + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. all: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ rewrite pren_epm_lift. cbn. reflexivity. } + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. all: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ + rewrite pren_epm_lift. cbn. f_equal. + unfold close. ssimpl. reflexivity. + } + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. all: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ + rewrite pren_epm_lift. cbn. f_equal. + unfold close. ssimpl. reflexivity. + } + 1:{ ssimpl. reflexivity. } + f_equal. + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ rewrite pren_epm_lift. cbn. reflexivity. } + 1:{ ssimpl. reflexivity. } + ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + f_equal. all: f_equal. + 1:{ rewrite pren_epm_lift. reflexivity. } + 1:{ ssimpl. reflexivity. } + ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. reflexivity. + + f_equal. all: f_equal. + 1:{ rewrite pren_epm_lift. reflexivity. } + 1:{ ssimpl. reflexivity. } + ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. reflexivity. + + f_equal. unfold close. ssimpl. + eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + - cbn. + erewrite IHt1. 2,3: eassumption. + erewrite IHt2. + 2:{ eapply rscoping_upren. eassumption. } + 2:{ eapply rscoping_comp_upren. eassumption. } + unfold plam. + erewrite erase_ren. 2,3: eassumption. + destruct_ifs. all: mode_eqs. + + cbn. unfold close. ssimpl. f_equal. + eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. + + cbn. rewrite pren_epm_lift. ssimpl. f_equal. f_equal. + eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. reflexivity. + + cbn. rewrite pren_epm_lift. ssimpl. f_equal. f_equal. + eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. reflexivity. + - cbn. + erewrite md_ren. 2,3: eassumption. + erewrite IHt1. 2,3: eassumption. + erewrite IHt2. 2,3: eassumption. + erewrite erase_ren. 2,3: eassumption. + erewrite revive_ren. 2,3: eassumption. + rewrite <- pren_epm_lift. rewrite <- pren_rpm_lift. + destruct_ifs. all: reflexivity. + - cbn. + erewrite md_ren. 2,3: eassumption. + destruct_ifs. all: eauto. + - cbn. + erewrite md_ren. 2,3: eassumption. + destruct_ifs. all: eauto. + - cbn. + erewrite md_ren. 2,3: eassumption. + destruct_ifs. 1: reflexivity. + cbn. erewrite IHt3. 2,3: eassumption. + erewrite IHt1. 2,3: eassumption. + erewrite ?erase_ren, ?revive_ren. 2,3: eassumption. + rewrite !pren_rpm_lift. reflexivity. + - cbn. + erewrite md_ren. 2,3: eassumption. + destruct_ifs. 2: reflexivity. + cbn. erewrite IHt2. 2,3: eassumption. + erewrite IHt1. 2,3: eassumption. + erewrite ?erase_ren, ?revive_ren. 2,3: eassumption. + rewrite !pren_rpm_lift. reflexivity. + - cbn. eauto. + - cbn. eauto. + - cbn. + erewrite ?erase_ren, ?revive_ren. 2-7: eassumption. + rewrite ?pren_epm_lift. reflexivity. + - cbn. + erewrite ?erase_ren, ?revive_ren. 2-5: eassumption. + rewrite ?pren_epm_lift. reflexivity. + - cbn. + erewrite md_ren. 2,3: eassumption. + erewrite ?erase_ren, ?revive_ren. 2-11: eassumption. + destruct (md _ _). + + eauto. + + unfold pcastTG. cbn. + erewrite IHt1. 2,3: eassumption. + erewrite IHt3. 2,3: eassumption. + erewrite IHt4. 2,3: eassumption. + erewrite IHt5. 2,3: eassumption. + erewrite IHt6. 2,3: eassumption. + rewrite ?pren_epm_lift, ?pren_rpm_lift. + f_equal. all: f_equal. all: f_equal. + 2:{ rewrite <- pren_epm_lift. ssimpl. reflexivity. } + 3:{ ssimpl. reflexivity. } + all: f_equal. + 2:{ ssimpl. reflexivity. } + all: f_equal. + 1:{ ssimpl. reflexivity. } + 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + 1:{ cEl_ren. rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + 3:{ ssimpl. reflexivity. } + all: f_equal. + 3:{ ssimpl. reflexivity. } + 3:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + all: f_equal. + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + all: f_equal. + 1:{ ssimpl. reflexivity. } + 2:{ rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. f_equal. + ssimpl. reflexivity. + + unfold pcastTG. cbn. + erewrite IHt1. 2,3: eassumption. + erewrite IHt3. 2,3: eassumption. + erewrite IHt4. 2,3: eassumption. + erewrite IHt5. 2,3: eassumption. + erewrite IHt6. 2,3: eassumption. + rewrite ?pren_epm_lift, ?pren_rpm_lift. + f_equal. all: f_equal. all: f_equal. + 2:{ rewrite <- pren_epm_lift. ssimpl. reflexivity. } + 3:{ ssimpl. reflexivity. } + all: f_equal. + 2:{ ssimpl. reflexivity. } + all: f_equal. + 1:{ ssimpl. reflexivity. } + 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + 1:{ cEl_ren. rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + 3:{ ssimpl. reflexivity. } + all: f_equal. + 3:{ ssimpl. reflexivity. } + 3:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + all: f_equal. + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + all: f_equal. + 1:{ ssimpl. reflexivity. } + 2:{ rewrite <- pren_epm_lift. ssimpl. reflexivity. } + f_equal. f_equal. + ssimpl. reflexivity. + + unfold pcastP. cbn. + erewrite IHt1. 2,3: eassumption. + erewrite IHt3. 2,3: eassumption. + erewrite IHt4. 2,3: eassumption. + erewrite IHt5. 2,3: eassumption. + erewrite IHt6. 2,3: eassumption. + rewrite ?pren_epm_lift, ?pren_rpm_lift. + f_equal. all: f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 3:{ ssimpl. reflexivity. } + all: f_equal. + 1:{ ssimpl. reflexivity. } + 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + all: f_equal. + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + 3:{ ssimpl. reflexivity. } + all: f_equal. + 3:{ ssimpl. reflexivity. } + 3:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + all: f_equal. + 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } + 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } + all: f_equal. + 1:{ ssimpl. reflexivity. } + f_equal. + ssimpl. reflexivity. + - cbn. reflexivity. + - cbn. reflexivity. + - cbn. reflexivity. + - cbn. + erewrite IHt1, IHt2, IHt3, IHt4. 2-9: eassumption. + erewrite ?erase_ren, ?revive_ren. 2-13: eassumption. + rewrite <- !pren_epm_lift. + change (epm_lift ⟦ ?G | ?t ⟧v) with (⟦ G | t⟧pv). + destruct m. 4: reflexivity. + + unfold pmifK. unfold perif. cbn. f_equal. f_equal. + all: f_equal. 1,4: f_equal. + 1,2: f_equal. 1-4: f_equal. 1,2,5-7,10: f_equal. + 2,3,5,6: f_equal. 2,4: f_equal. + all: ssimpl. all: reflexivity. + + cbn. unfold pmif, plam. cbn. f_equal. f_equal. f_equal. + ssimpl. reflexivity. + + cbn. unfold pmif, plam. cbn. f_equal. f_equal. f_equal. + ssimpl. reflexivity. + - cbn. reflexivity. + - cbn. reflexivity. + - cbn. f_equal. eauto. + - cbn. + erewrite IHt1, IHt2, IHt3, IHt4. 2-9: eassumption. + erewrite ?erase_ren, ?revive_ren. 2-13: eassumption. + rewrite <- !pren_epm_lift. + change (epm_lift ⟦ ?G | ?t ⟧v) with (⟦ G | t⟧pv). + destruct m. all: reflexivity. + - cbn. erewrite IHt1, IHt2. 2-5: eassumption. + erewrite ?erase_ren, ?revive_ren. 2-5: eassumption. + rewrite <- !pren_epm_lift. + reflexivity. + - cbn. erewrite IHt. 2-3: eassumption. + reflexivity. + - cbn. erewrite IHt1, IHt2, IHt3. 2-7: eassumption. + reflexivity. + - cbn. + erewrite IHt1, IHt2, IHt3, IHt4, IHt5, IHt6. 2-13: eassumption. + erewrite ?erase_ren, ?revive_ren. 2-17: eassumption. + rewrite <- !pren_epm_lift. + change (epm_lift ⟦ ?G | ?t ⟧v) with (⟦ G | t⟧pv). + destruct m. all: reflexivity. + - cbn. reflexivity. + - cbn. destruct_ifs. all: mode_eqs. + + cbn. f_equal. all: eauto. + + cbn. f_equal. 1: f_equal. all: eauto. + erewrite erase_ren. 2,3: eassumption. + rewrite pren_epm_lift. reflexivity. + + cbn. f_equal. 1: f_equal. all: eauto. + erewrite erase_ren. 2,3: eassumption. + rewrite pren_epm_lift. reflexivity. +Qed. + diff --git a/theories/param/Scope.v b/theories/param/Scope.v new file mode 100644 index 0000000..f071edb --- /dev/null +++ b/theories/param/Scope.v @@ -0,0 +1,424 @@ +From Coq Require Import Utf8 List Bool Lia. +From Equations Require Import Equations. +From GhostTT.autosubst Require Import CCAST GAST core unscoped. +From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl + CScoping Scoping CTyping TermMode Typing BasicMetaTheory CCMetaTheory Erasure + Revival. +From Coq Require Import Setoid Morphisms Relation_Definitions. +From GhostTT.param Require Export Term. + +Import ListNotations. +Import CombineNotations. + +Set Default Goal Selector "!". +Set Equations Transparent. + +Equations param_ctx (Γ : context) : ccontext := { + ⟦ [] ⟧p := [] ; + ⟦ Γ ,, (mx, A) ⟧p := + if isProp mx then None :: Some (cProp, ⟦ sc Γ | A ⟧p) :: ⟦ Γ ⟧p + else if isKind mx then + Some (cType, capp (S ⋅ ⟦ sc Γ | A ⟧p) (cvar 0)) :: + Some (cType, ⟦ sc Γ | A ⟧pτ) :: ⟦ Γ ⟧p + else + Some (cProp, capp (S ⋅ ⟦ sc Γ | A ⟧p) (cvar 0)) :: + Some (cType, ⟦ sc Γ | A ⟧pτ) :: ⟦ Γ ⟧p +} +where "⟦ G '⟧p'" := (param_ctx G). + +Fixpoint param_sc (Γ : scope) : cscope := + match Γ with + | [] => [] + | m :: Γ => + if isProp m then None :: Some cProp :: param_sc Γ + else if isKind m then Some cType :: Some cType :: param_sc Γ + else Some cProp :: Some cType :: param_sc Γ + end. + +(** Scope lookups **) + +Lemma nth_error_param_vreg : + ∀ Γ x, + nth_error (param_sc Γ) (vreg x) = + option_map (λ m, if isProp m then Some cProp else Some cType) (nth_error Γ x). +Proof. + intros Γ x. + induction Γ as [| m Γ ih] in x |- *. + - destruct x. all: reflexivity. + - destruct x. + + cbn. destruct_ifs. all: reflexivity. + + unfold vreg. simpl "*". remember (S (x * 2)) as y eqn:e. + cbn. subst. destruct_ifs. all: eapply ih. +Qed. + +Lemma nth_error_param_vpar : + ∀ Γ x, + nth_error (param_sc Γ) (vpar x) = + option_map (λ m, + if isProp m then None + else if isKind m then Some cType + else Some cProp + ) (nth_error Γ x). +Proof. + intros Γ x. + induction Γ as [| m Γ ih] in x |- *. + - destruct x. all: reflexivity. + - destruct x. + + cbn. destruct_ifs. all: reflexivity. + + cbn. destruct_ifs. all: eapply ih. +Qed. + +(** ⟦ Γ ⟧v is a sub-context of ⟦ Γ ⟧p **) + +Lemma scoping_rev_sub_param : + ∀ Γ, + crscoping (param_sc Γ) vreg (revive_sc Γ). +Proof. + intros Γ. intros x m e. + unfold revive_sc in e. rewrite nth_error_map in e. + rewrite nth_error_param_vreg. + destruct (nth_error Γ x) as [mx|] eqn:ex. 2: discriminate. + cbn in e. cbn. + destruct_ifs. 1: discriminate. + assumption. +Qed. + +Hint Resolve scoping_rev_sub_param : cc_scope. + +Lemma typing_rev_sub_param : + ∀ Γ, + crtyping ⟦ Γ ⟧p vreg ⟦ Γ ⟧v. +Proof. + intros Γ. intros x m A e. + assert (h : nth_error ⟦ Γ ⟧p (vreg x) = Some (Some (m, vreg ⋅ A))). + { induction Γ as [| [my B] Γ ih] in x, m, A, e |- *. + 1:{ destruct x. all: discriminate. } + destruct x. + - cbn in e. + destruct (isProp my) eqn:ey. 1: discriminate. + noconf e. cbn. rewrite ey. + destruct_if e1. all: reflexivity. + - cbn in e. + unfold vreg. simpl "*". remember (S (x * 2)) as z eqn:ez. + cbn. subst. + destruct_if ey. + + eapply ih. assumption. + + destruct_if e1. + * eapply ih. assumption. + * eapply ih. assumption. + } + eexists. split. + - eassumption. + - ssimpl. unfold vreg. cbn. ssimpl. eapply extRen_cterm. + intro. ssimpl. lia. +Qed. + +(** ⟦ Γ ⟧ε is a sub-context of ⟦ Γ ⟧p **) + +Lemma scoping_er_sub_param : + ∀ Γ, + crscoping (param_sc Γ) vreg (erase_sc Γ). +Proof. + intros Γ. + pose proof (scoping_sub_rev Γ) as lem. + eapply crscoping_comp in lem. 2: eapply scoping_rev_sub_param. + eapply crscoping_morphism. all: eauto. + intros x. reflexivity. +Qed. + +Hint Resolve scoping_er_sub_param : cc_scope. + +Lemma typing_er_sub_param : + ∀ Γ, + crtyping ⟦ Γ ⟧p vreg ⟦ Γ ⟧ε. +Proof. + intros Γ. + pose proof (typing_sub_rev Γ) as lem. + eapply crtyping_comp in lem. 2: eapply typing_rev_sub_param. + eapply crtyping_morphism. all: eauto. + intros x. reflexivity. +Qed. + +(** Parametricity preserves scoping **) + +Lemma scoping_epm_lift : + ∀ Γ Γ' t, + ccscoping (erase_sc Γ) t cType → + Γ' = param_sc Γ → + ccscoping Γ' (epm_lift t) cType. +Proof. + intros Γ Γ' t h ->. + eapply cscoping_ren. + - eapply scoping_er_sub_param. + - assumption. +Qed. + +(* Hint Resolve scoping_epm_lift | 1000 : cc_scope. *) + +Lemma pscoping_erase_term : + ∀ Γ Γ' t, + Γ' = param_sc Γ → + ccscoping Γ' ⟦ Γ | t ⟧pε cType. +Proof. + intros Γ Γ' t ->. + eapply scoping_epm_lift. + - eapply erase_scoping. + - reflexivity. +Qed. + +Hint Resolve pscoping_erase_term : cc_scope. + +Lemma pscoping_erase_type : + ∀ Γ Γ' t, + Γ' = param_sc Γ → + ccscoping Γ' ⟦ Γ | t ⟧pτ cType. +Proof. + intros Γ Γ' t ->. + eapply scoping_epm_lift. + - constructor. eapply erase_scoping. + - reflexivity. +Qed. + +Hint Resolve pscoping_erase_type : cc_scope. + +Lemma pscoping_erase_err : + ∀ Γ Γ' t, + Γ' = param_sc Γ → + ccscoping Γ' ⟦ Γ | t ⟧p∅ cType. +Proof. + intros Γ Γ' t ->. + eapply scoping_epm_lift. + - constructor. eapply erase_scoping. + - reflexivity. +Qed. + +Hint Resolve pscoping_erase_err : cc_scope. + +Lemma scoping_rpm_lift : + ∀ Γ Γ' t, + ccscoping (revive_sc Γ) t cType → + Γ' = param_sc Γ → + ccscoping Γ' (rpm_lift t) cType. +Proof. + intros Γ Γ' t h ->. + eapply cscoping_ren. + - eapply scoping_rev_sub_param. + - assumption. +Qed. + +(* Hint Resolve scoping_rpm_lift | 1000 : cc_scope. *) + +Lemma pscoping_revive : + ∀ Γ Γ' t, + Γ' = param_sc Γ → + ccscoping Γ' ⟦ Γ | t ⟧pv cType. +Proof. + intros Γ Γ' t ->. + eapply scoping_rpm_lift. + - eapply revive_scoping. + - reflexivity. +Qed. + +Hint Resolve pscoping_revive : cc_scope. + +Lemma erase_scoping_eq : + ∀ Γ Γ' t, + Γ' = erase_sc Γ → + ccscoping Γ' ⟦ Γ | t ⟧ε cType. +Proof. + intros Γ ? t ->. + eapply erase_scoping. +Qed. + +Lemma revive_scoping_eq : + ∀ Γ Γ' t, + Γ' = revive_sc Γ → + ccscoping Γ' ⟦ Γ | t ⟧v cType. +Proof. + intros Γ ? t ->. + eapply revive_scoping. +Qed. + +Hint Resolve erase_scoping_eq : cc_scope. +Hint Resolve revive_scoping_eq : cc_scope. +Hint Resolve revive_scoping : cc_scope. +Hint Resolve crscoping_plus : cc_scope. + +Lemma pPi_scoping : + ∀ Γ mx A B C, + ccscoping Γ A cType → + ccscoping Γ B cType → + ccscoping (Some mx :: Some cType :: Γ) C cType → + ccscoping Γ (pPi mx A B C) cType. +Proof. + intros Γ mx A B C hA hB hC. + unshelve eauto with cc_scope shelvedb ; shelve_unifiable. + constructor. reflexivity. +Qed. + +Hint Resolve pPi_scoping : cc_scope. + +(* So that they're not unfolded too eagerly *) +Opaque epm_lift rpm_lift. + +Lemma param_scoping : + ∀ Γ t m, + scoping Γ t m → + ccscoping (param_sc Γ) ⟦ Γ | t ⟧p (if isKind m then cType else cProp). +Proof. + intros Γ t m h. + induction h. + all: try solve [ cbn ; eauto with cc_scope ]. + all: try solve [ cbn ; destruct_ifs ; eauto with cc_scope ]. + - cbn. rewrite H. destruct_if e. + + mode_eqs. cbn. constructor. + rewrite nth_error_param_vreg. rewrite H. reflexivity. + + constructor. rewrite nth_error_param_vpar. rewrite H. + cbn. rewrite e. destruct_ifs. all: reflexivity. + - cbn. + destruct m, mx. all: cbn in *. + all: try solve [ typeclasses eauto 50 with cc_scope ]. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + 1:{ + eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + } + eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + reflexivity. + * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + reflexivity. + * eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + reflexivity. + * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + reflexivity. + * eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + * eapply scoping_epm_lift. 2: reflexivity. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + reflexivity. + * eapply crscoping_shift. eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + - cbn in *. destruct_ifs. all: mode_eqs. all: cbn in *. + all: try solve [ typeclasses eauto 50 with cc_scope ]. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + - cbn in *. + erewrite scoping_md. 2: eassumption. + cbn. assumption. + - cbn in *. + erewrite scoping_md. 2: eassumption. + destruct_ifs. all: mode_eqs. all: try intuition discriminate. + 1:{ destruct m. all: intuition discriminate. } + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + reflexivity. + - cbn in *. + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: reflexivity. + - cbn in *. + erewrite scoping_md. 2: eassumption. + destruct m. 1: contradiction. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + all: repeat try eapply crscoping_shift. + all: eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + all: repeat try eapply crscoping_shift. + all: eauto with cc_scope. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + all: try reflexivity. + all: repeat try eapply crscoping_shift. + all: eauto with cc_scope. + - cbn in *. + destruct m. + + cbn in *. escope. all: reflexivity. + + cbn in *. escope. all: reflexivity. + + cbn in *. escope. all: reflexivity. + + cbn in *. escope. + - cbn in *. + destruct m. + + contradiction. + + cbn in *. escope. all: reflexivity. + + cbn in *. escope. all: reflexivity. + + cbn in *. escope. all: reflexivity. + - cbn in *. + destruct m. + + contradiction. + + cbn in *. escope. all: reflexivity. + + cbn in *. escope. all: reflexivity. + + cbn in *. escope. all: reflexivity. + - cbn in *. + destruct_ifs. all: mode_eqs. all: try discriminate. + all: try solve [ typeclasses eauto 50 with cc_scope ]. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + reflexivity. + + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. + reflexivity. +Qed. + diff --git a/theories/param/Substitution.v b/theories/param/Substitution.v new file mode 100644 index 0000000..8707285 --- /dev/null +++ b/theories/param/Substitution.v @@ -0,0 +1,561 @@ +From Coq Require Import Utf8 List Bool Lia. +From Equations Require Import Equations. +From GhostTT.autosubst Require Import CCAST GAST core unscoped. +From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl + CScoping Scoping CTyping TermMode Typing BasicMetaTheory CCMetaTheory Erasure + Revival. +From Coq Require Import Setoid Morphisms Relation_Definitions. +From GhostTT Require Export Term. +From GhostTT Require Import Scope Renaming. + +Import ListNotations. +Import CombineNotations. + +Set Default Goal Selector "!". +Set Equations Transparent. + +Opaque epm_lift rpm_lift. + +(** Parametricity commutes with substitution + + As for revival we need to craft a new substitution that gets the scopes as + input in order to determine the mode of the various variables. + +**) + +Definition psubst Δ Γ σ n := + let p := Nat.div2 n in + match nth_error Δ p with + | Some m => + if relm m then ( + if Nat.odd n then ⟦ Γ | σ p ⟧pε + else ⟦ Γ | σ p ⟧p + ) + else if isGhost m then ( + if Nat.odd n then ⟦ Γ | σ p ⟧pv + else ⟦ Γ | σ p ⟧p + ) + else ( + if Nat.odd n then ⟦ Γ | σ p ⟧p + else cDummy + ) + | None => cDummy + end. + +Lemma div2_vreg : + ∀ n, Nat.div2 (vreg n) = n. +Proof. + intros n. + unfold vreg. replace (n * 2) with (2 * n) by lia. + apply PeanoNat.Nat.div2_succ_double. +Qed. + +Lemma div2_vpar : + ∀ n, Nat.div2 (vpar n) = n. +Proof. + intros n. + unfold vpar. replace (n * 2) with (2 * n) by lia. + apply PeanoNat.Nat.div2_double. +Qed. + +Lemma odd_vreg : + ∀ n, Nat.odd (vreg n) = true. +Proof. + intros n. + unfold vreg. replace (n * 2) with (2 * n) by lia. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_mul. reflexivity. +Qed. + +Lemma odd_vpar : + ∀ n, Nat.odd (vpar n) = false. +Proof. + intros n. + unfold vpar. replace (n * 2) with (2 * n) by lia. + rewrite PeanoNat.Nat.odd_mul. reflexivity. +Qed. + +Lemma psubst_SS : + ∀ m Δ Γ σ n, + psubst (m :: Δ) (m :: Γ) (up_term σ) (S (S n)) = + plus 2 ⋅ psubst Δ Γ σ n. +Proof. + intros m Δ Γ σ n. + unfold psubst. rewrite div2_SS. cbn. + destruct nth_error eqn:e. 2: reflexivity. + rewrite PeanoNat.Nat.odd_succ. + rewrite PeanoNat.Nat.even_succ. + destruct_ifs. all: mode_eqs. + - ssimpl. erewrite erase_ren. + 2: eapply rscoping_S. + 2: eapply rscoping_comp_S. + ssimpl. rewrite <- pren_epm_lift. + ssimpl. eapply extRen_cterm. + intro x. unfold shift. change (pren S) with (pren (id >> S)). + rewrite pren_comp_S. ssimpl. rewrite pren_id. reflexivity. + - ssimpl. erewrite param_ren. + 2: eapply rscoping_S. + 2: eapply rscoping_comp_S. + ssimpl. eapply extRen_cterm. + intro x. unfold shift. change (pren S) with (pren (id >> S)). + rewrite pren_comp_S. ssimpl. rewrite pren_id. reflexivity. + - ssimpl. erewrite revive_ren. + 2: eapply rscoping_S. + 2: eapply rscoping_comp_S. + ssimpl. rewrite <- pren_rpm_lift. + eapply extRen_cterm. + intro x. unfold shift. change (pren S) with (pren (id >> S)). + rewrite pren_comp_S. ssimpl. rewrite pren_id. reflexivity. + - ssimpl. erewrite param_ren. + 2: eapply rscoping_S. + 2: eapply rscoping_comp_S. + ssimpl. eapply extRen_cterm. + intro x. unfold shift. change (pren S) with (pren (id >> S)). + rewrite pren_comp_S. ssimpl. rewrite pren_id. reflexivity. + - ssimpl. erewrite param_ren. + 2: eapply rscoping_S. + 2: eapply rscoping_comp_S. + ssimpl. eapply extRen_cterm. + intro x. unfold shift. change (pren S) with (pren (id >> S)). + rewrite pren_comp_S. ssimpl. rewrite pren_id. reflexivity. + - reflexivity. +Qed. + +Transparent epm_lift rpm_lift. + +Lemma psubst_epm_lift : + ∀ Γ Δ σ t, + ccscoping (erase_sc Δ) t cType → + (epm_lift t) <[ psubst Δ Γ σ ] = epm_lift (t <[ σ >> erase_term Γ ]). +Proof. + intros Γ Δ σ t ht. + unfold epm_lift. ssimpl. + eapply ext_cterm_scoped. 1: eassumption. + intros x hx. + ssimpl. unfold psubst. rewrite div2_vreg. + unfold inscope in hx. unfold erase_sc in hx. + rewrite nth_error_map in hx. + destruct (nth_error Δ x) eqn:e. 2: discriminate. + cbn in hx. destruct (relm m) eqn:e1. 2: discriminate. + rewrite odd_vreg. reflexivity. +Qed. + +Lemma psubst_rpm_lift : + ∀ Γ Δ σ t, + ccscoping (revive_sc Δ) t cType → + (rpm_lift t) <[ psubst Δ Γ σ ] = rpm_lift (t <[ rev_subst Δ Γ σ ]). +Proof. + intros Γ Δ σ t ht. + unfold rpm_lift. ssimpl. + eapply ext_cterm_scoped. 1: eassumption. + intros x hx. + ssimpl. unfold psubst. rewrite div2_vreg. + unfold rev_subst. unfold ghv. + unfold inscope in hx. unfold revive_sc in hx. + rewrite nth_error_map in hx. + destruct (nth_error Δ x) eqn:e. 2: discriminate. + cbn in hx. destruct (isProp m) eqn:e1. 1: discriminate. + rewrite odd_vreg. + destruct_ifs. all: mode_eqs. all: try discriminate. + all: try reflexivity. + destruct m. all: discriminate. +Qed. + +Opaque epm_lift rpm_lift. + +Lemma param_subst : + ∀ Γ Δ σ t, + sscoping Γ σ Δ → + sscoping_comp Γ σ Δ → + ⟦ Γ | t <[ σ ] ⟧p = ⟦ Δ | t ⟧p <[ psubst Δ Γ σ ]. +Proof. + intros Γ Δ σ t hσ hcσ. + induction t in Γ, Δ, σ, hσ, hcσ |- *. + - cbn. destruct (nth_error Δ n) eqn:e. + + destruct_if e1. + * mode_eqs. cbn. unfold psubst. rewrite div2_vreg. + rewrite e. cbn. rewrite odd_vreg. reflexivity. + * cbn. unfold psubst. rewrite div2_vpar. rewrite e. + rewrite odd_vpar. + destruct_ifs. all: try reflexivity. + destruct m. all: discriminate. + + eapply hcσ in e as e'. destruct e' as [k [e1 e2]]. + rewrite e1. cbn. rewrite e2. reflexivity. + - cbn. destruct_ifs. all: reflexivity. + - cbn. + unfold pmPi, pmPiP, pmPiNP, pPi. + erewrite IHt1. 2,3: eassumption. + erewrite IHt2. + 2:{ eapply sscoping_shift. eassumption. } + 2:{ eapply sscoping_comp_shift. assumption. } + erewrite !erase_subst. + 2-5: eauto using sscoping_shift, sscoping_comp_shift with cc_scope. + change (cEl (?t <[ ?σ ])) with ((cEl t) <[ σ ]). + erewrite <- psubst_epm_lift. 2:{ econstructor. eapply erase_scoping. } + destruct m, m0. all: cbn in *. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. all: cbn. all: f_equal. + 2:{ ssimpl. reflexivity. } + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + } + f_equal. all: f_equal. all: f_equal. + all: eapply ext_cterm. all: ssimpl. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. all: cbn. all: f_equal. + 2:{ ssimpl. reflexivity. } + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + } + f_equal. all: f_equal. all: f_equal. + all: eapply ext_cterm. all: ssimpl. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. all: cbn. all: f_equal. + 2:{ ssimpl. reflexivity. } + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + } + unfold cty_lift. f_equal. all: f_equal. + all: unfold close. all: ssimpl. + all: eapply ext_cterm. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + all: ssimpl. all: reflexivity. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. + rewrite rinstInst'_cterm. reflexivity. + } + cbn. unfold cty_lift. f_equal. f_equal. all: f_equal. + all: unfold close. all: ssimpl. + all: eapply ext_cterm. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + all: ssimpl. all: reflexivity. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + } + cbn. f_equal. f_equal. all: f_equal. all: f_equal. + all: ssimpl. + all: eapply ext_cterm. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + } + cbn. f_equal. f_equal. all: f_equal. all: f_equal. + all: ssimpl. + all: eapply ext_cterm. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + } + cbn. unfold cty_lift. f_equal. f_equal. all: f_equal. all: unfold close. + all: ssimpl. + all: eapply ext_cterm. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + all: ssimpl. all: reflexivity. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. + erewrite rinstInst'_cterm. reflexivity. + } + cbn. unfold cty_lift. f_equal. f_equal. all: f_equal. all: unfold close. + all: ssimpl. + all: eapply ext_cterm. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + all: ssimpl. all: reflexivity. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + } + cbn. f_equal. f_equal. all: f_equal. all: f_equal. + all: ssimpl. + all: eapply ext_cterm. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + } + cbn. f_equal. f_equal. all: f_equal. all: f_equal. + all: ssimpl. + all: eapply ext_cterm. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + } + cbn. f_equal. f_equal. all: f_equal. all: f_equal. + all: ssimpl. + all: eapply ext_cterm. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + all: ssimpl. + * rewrite rinstInst'_cterm. reflexivity. + * rewrite rinstInst'_cterm. reflexivity. + + f_equal. all: f_equal. + 2:{ ssimpl. reflexivity. } + 1: rewrite psubst_epm_lift. + 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } + all: f_equal. + 2:{ + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. + erewrite rinstInst'_cterm. reflexivity. + } + cbn. unfold cty_lift. f_equal. f_equal. all: f_equal. all: unfold close. + all: ssimpl. + all: eapply ext_cterm. all: intros []. + all: cbn. 1,3: reflexivity. + all: ssimpl. + all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. + all: ssimpl. all: reflexivity. + + f_equal. all: f_equal. 1: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + + f_equal. all: f_equal. 1: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + + f_equal. all: f_equal. 1: f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. reflexivity. + + f_equal. unfold close. + ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. + rewrite rinstInst'_cterm. reflexivity. + - cbn. + erewrite IHt1. 2,3: eassumption. + erewrite IHt2. + 2:{ eapply sscoping_shift. eassumption. } + 2:{ eapply sscoping_comp_shift. assumption. } + erewrite erase_subst. 2,3: eassumption. + change (cEl (?t <[ ?σ ])) with ((cEl t) <[ σ ]). + erewrite <- psubst_epm_lift. 2:{ econstructor. eapply erase_scoping. } + destruct_ifs. all: mode_eqs. + + cbn. f_equal. unfold close. ssimpl. + eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + ssimpl. rewrite psubst_SS. ssimpl. + erewrite rinstInst'_cterm. reflexivity. + + cbn. f_equal. unfold plam. f_equal. f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. + rewrite psubst_SS. ssimpl. reflexivity. + + cbn. unfold plam. f_equal. f_equal. + * ssimpl. reflexivity. + * ssimpl. eapply ext_cterm. intros [| []]. all: cbn. + --- destruct_ifs. all: mode_eqs. all: try discriminate. + all: try reflexivity. + destruct m. all: discriminate. + --- destruct_ifs. all: mode_eqs. all: try discriminate. + all: try reflexivity. + destruct m. all: discriminate. + --- rewrite psubst_SS. ssimpl. reflexivity. + - cbn. + erewrite md_subst. 2,3: eassumption. + erewrite IHt1. 2,3: eassumption. + erewrite IHt2. 2,3: eassumption. + erewrite erase_subst. 2,3: eassumption. + erewrite revive_subst. 2,3: eassumption. + erewrite <- psubst_rpm_lift. 2: eapply revive_scoping. + erewrite <- psubst_epm_lift. 2: eapply erase_scoping. + destruct_ifs. all: reflexivity. + - cbn. + erewrite md_subst. 2,3: eassumption. + erewrite IHt. 2,3: eassumption. + destruct_ifs. all: reflexivity. + - cbn. + erewrite md_subst. 2,3: eassumption. + erewrite IHt. 2,3: eassumption. + destruct_ifs. all: reflexivity. + - cbn. + erewrite md_subst. 2,3: eassumption. + erewrite IHt1. 2,3: eassumption. + erewrite IHt3. 2,3: eassumption. + erewrite revive_subst. 2,3: eassumption. + erewrite <- psubst_rpm_lift. 2: eapply revive_scoping. + destruct_ifs. all: reflexivity. + - cbn. + erewrite md_subst. 2,3: eassumption. + erewrite IHt1. 2,3: eassumption. + erewrite IHt2. 2,3: eassumption. + erewrite revive_subst. 2,3: eassumption. + erewrite <- psubst_rpm_lift. 2: eapply revive_scoping. + destruct_ifs. all: reflexivity. + - cbn. eauto. + - cbn. eauto. + - cbn. + erewrite !revive_subst. 2-5: eassumption. + erewrite !erase_subst. 2,3: eassumption. + erewrite <- !psubst_rpm_lift. 2,3: eapply revive_scoping. + change (cEl (?t <[ ?σ ])) with ((cEl t) <[ σ ]). + erewrite <- psubst_epm_lift. 2:{ econstructor. eapply erase_scoping. } + reflexivity. + - cbn. + erewrite erase_subst. 2,3: eassumption. + erewrite revive_subst. 2,3: eassumption. + erewrite <- psubst_rpm_lift. 2: eapply revive_scoping. + change (cEl (?t <[ ?σ ])) with ((cEl t) <[ σ ]). + erewrite <- psubst_epm_lift. 2:{ econstructor. eapply erase_scoping. } + reflexivity. + - cbn. + erewrite md_subst. 2,3: eassumption. + erewrite IHt1. 2,3: eassumption. + erewrite IHt3. 2,3: eassumption. + erewrite IHt4. 2,3: eassumption. + erewrite IHt5. 2,3: eassumption. + erewrite IHt6. 2,3: eassumption. + erewrite !erase_subst. 2-5: eassumption. + erewrite !revive_subst. 2-7: eassumption. + erewrite <- !psubst_rpm_lift. 2-4: eapply revive_scoping. + change (cEl (?t <[ ?σ ])) with ((cEl t) <[ σ ]). + erewrite <- !psubst_epm_lift. + 2: eapply erase_scoping. + 2:{ econstructor. eapply erase_scoping. } + destruct md eqn:e. + + reflexivity. + + unfold pcastTG. cbn. ssimpl. reflexivity. + + unfold pcastTG. cbn. ssimpl. reflexivity. + + unfold pcastP. cbn. ssimpl. reflexivity. + - cbn. reflexivity. + - cbn. reflexivity. + - cbn. reflexivity. + - cbn. + erewrite IHt1, IHt2, IHt3, IHt4. 2-9: eassumption. + erewrite !erase_subst. 2-9: eassumption. + erewrite !revive_subst. 2-5: eassumption. + erewrite <- !psubst_rpm_lift. 2-3: eapply revive_scoping. + erewrite <- !psubst_epm_lift. 2-5: eapply erase_scoping. + destruct m. + + unfold pmifK. unfold perif. cbn. f_equal. f_equal. all: f_equal. + 1,4: f_equal. 1,2: f_equal. 1-4: f_equal. 1,2,5-7,10: f_equal. + 2,3,5,6: f_equal. 2,4: f_equal. + all: ssimpl. all: reflexivity. + + unfold pmif, plam. cbn. f_equal. f_equal. f_equal. + ssimpl. reflexivity. + + unfold pmif, plam. cbn. f_equal. f_equal. f_equal. + ssimpl. reflexivity. + + reflexivity. + - cbn. reflexivity. + - cbn. reflexivity. + - cbn. f_equal. eauto. + - cbn. + erewrite IHt1, IHt2, IHt3, IHt4. 2-9: eassumption. + erewrite !erase_subst. 2-9: eassumption. + erewrite !revive_subst. 2-5: eassumption. + erewrite <- !psubst_rpm_lift. 2-3: eapply revive_scoping. + erewrite <- !psubst_epm_lift. 2-5: eapply erase_scoping. + destruct m. all: reflexivity. + - cbn. erewrite IHt1, IHt2. 2-5: eassumption. + erewrite !erase_subst, !revive_subst. 2-5: eassumption. + erewrite <- !psubst_epm_lift, <- !psubst_rpm_lift. + 2: eapply revive_scoping. + 2: eapply erase_scoping. + reflexivity. + - cbn. erewrite IHt. 2-3: eassumption. + reflexivity. + - cbn. erewrite IHt1, IHt2, IHt3. 2-7: eassumption. + reflexivity. + - cbn. + erewrite IHt1, IHt2, IHt3, IHt4, IHt5, IHt6. 2-13: eassumption. + erewrite !erase_subst. 2-11: eassumption. + erewrite !revive_subst. 2-7: eassumption. + erewrite <- !psubst_rpm_lift. 2-4: eapply revive_scoping. + erewrite <- !psubst_epm_lift. 2-6: eapply erase_scoping. + destruct m. all: reflexivity. + - cbn. reflexivity. + - cbn. + erewrite IHt1. 2,3: eassumption. + erewrite IHt2. 2,3: eassumption. + erewrite erase_subst. 2,3: eassumption. + destruct_ifs. all: try reflexivity. + + cbn. f_equal. f_equal. + rewrite psubst_epm_lift. 2: eauto with cc_scope. + reflexivity. + + cbn. f_equal. f_equal. + rewrite psubst_epm_lift. 2: eauto with cc_scope. + reflexivity. +Qed. diff --git a/theories/param/Term.v b/theories/param/Term.v new file mode 100644 index 0000000..dd8b6dd --- /dev/null +++ b/theories/param/Term.v @@ -0,0 +1,322 @@ +(*** Parametricity ***) + +From Coq Require Import Utf8 List Bool Lia. +From Equations Require Import Equations. +From GhostTT.autosubst Require Import CCAST GAST unscoped. +From GhostTT Require Import BasicAST SubstNotations ContextDecl + CTyping TermMode Typing CCMetaTheory Erasure Revival. + +Import ListNotations. + +Set Default Goal Selector "!". +Set Equations Transparent. + +(** Variables and parametricity + + x : A in the context is translated to x : A, xP : AP when A is not a Prop. + When x : P : Prop then x is translated to only one variable. To keep the + context regular we will still make use of our flexible contexts. + Variables are then either odd and regular or even and correspond to + parametricity. + +**) + +Definition vreg x := S (x * 2). +Definition vpar x := x * 2. + +(** Lifting erasure and revival + + Because erasure and revival produce terms in ⟦ Γ ⟧ε and ⟦ Γ ⟧v respectively + when we expect ⟦ Γ ⟧p, we need to do some lifting. Thankfully this lifting + can be done simply by using vreg as a renaming. + We define handy notations to make it all work. + +**) + +Definition epm_lift (t : cterm) := + vreg ⋅ t. + +Definition rpm_lift (t : cterm) := + vreg ⋅ t. + +Notation "⟦ G | u '⟧pε'":= + (epm_lift ⟦ G | u ⟧ε) (at level 9, G, u at next level). + +Notation "⟦ G | u '⟧pτ'":= + (epm_lift ⟦ G | u ⟧τ) (at level 9, G, u at next level). + +Notation "⟦ G | u '⟧p∅'":= + (epm_lift ⟦ G | u ⟧∅) (at level 9, G, u at next level). + +Notation "⟦ G | u '⟧pv'":= + (rpm_lift ⟦ G | u ⟧v) (at level 9, G, u at next level). + +(** Parametricity translation + + We start by defining useful shorthands. + +**) + +Definition pKind i := + clam cType (cty i) (cEl (cvar 0) ⇒[ cType ] cSort cType i). + +Definition pType i := + clam cType (cty i) (cEl (cvar 0) ⇒[ cType ] cSort cProp 0). + +Definition pProp := + clam cType cunit (cSort cProp 0). + +(* ∀ (x : A) (x@mp : B x). C *) +Definition pPi mp A B C := + cPi cType A (cPi mp (capp (S ⋅ B) (cvar 0)) C). + +Definition plam mp A B t := + clam cType A (clam mp (capp (S ⋅ B) (cvar 0)) t). + +Definition pcastTG Ae AP uv vv vP eP PP te tP := + sq_elim + eP + (clam cProp (squash (teq Ae uv vv)) (S ⋅ (capp (capp (capp PP vv) vP) te))) + (clam cType (teq Ae uv vv) ( + capp + (tJ + (cvar 0) + (S ⋅ (clam cType Ae ( + clam cType (teq (S ⋅ Ae) (S ⋅ uv) (cvar 0)) ( + cPi cProp (capp (plus 2 ⋅ AP) (cvar 1)) ( + capp (capp (capp (plus 3 ⋅ PP) (cvar 2)) (cvar 0)) (plus 3 ⋅ te) + ) + ) + ))) + (S ⋅ (clam cProp (capp AP uv) (S ⋅ tP)))) + (S ⋅ vP) + )). + +Definition pcastP Ae AP uv vv vP eP PP tP := + sq_elim + eP + (clam cProp (squash (teq Ae uv vv)) (S ⋅ (capp (capp PP vv) vP))) + (clam cType (teq Ae uv vv) ( + capp + (tJ + (cvar 0) + (S ⋅ (clam cType Ae ( + clam cType (teq (S ⋅ Ae) (S ⋅ uv) (cvar 0)) ( + cPi cProp (capp (plus 2 ⋅ AP) (cvar 1)) ( + capp (capp (plus 3 ⋅ PP) (cvar 2)) (cvar 0) + ) + ) + ))) + (S ⋅ (clam cProp (capp AP uv) (S ⋅ tP)))) + (S ⋅ vP) + )). + +Reserved Notation "⟦ G | u '⟧p'" (at level 9, G, u at next level). + +(** Translation of Pi types, to factorise a bit **) + +(* Prop case *) +Definition pmPiP mx Ae Ap Bp := + if isProp mx then cPi cProp Ap (close Bp) + else if isKind mx then pPi cType Ae Ap Bp + else pPi cProp Ae Ap Bp. + +(* Non-prop case *) +Definition pmPiNP mx m Te Ae Ap Bp := + let cm := if isKind mx then cType else cProp in + clam cType Te ( + if isProp mx then cPi cProp (S ⋅ Ap) (capp ((up_ren S) ⋅ (close Bp)) (cvar 1)) + else ( + pPi cm (S ⋅ Ae) (S ⋅ Ap) (capp ((up_ren (up_ren S)) ⋅ Bp) ( + if isGhost mx && relm m then cvar 2 + else capp (cvar 2) (cvar 1) + )) + ) + ). + +(* General case *) +Definition pmPi mx m Te Ae Ap Bp := + if isProp m then pmPiP mx Ae Ap Bp + else pmPiNP mx m Te Ae Ap Bp. + +(* Parametricity for if *) + +Definition perif be Pe te fe := + eif cType be + (clam cType ebool (cEl (capp (S ⋅ Pe) (cvar 0)))) + te fe (cErr (capp Pe bool_err)). + +Definition pmif bP Pe PP te tP fe fP := + pif bP + (plam cProp ebool pbool ( + capp + (capp (capp (S ⋅ S ⋅ PP) (cvar 1)) (cvar 0)) + (S ⋅ (perif (cvar 0) (S ⋅ Pe) (S ⋅ te) (S ⋅ fe))) + )) + tP fP. + +(* Proves cbot from pbool bool_err *) +Definition pbool_bool_err_inv h := + pif h (clam cType ebool ( + clam cProp (capp pbool (cvar 0)) ( + eif cType (cvar 1) (clam cType ebool (cSort cProp 0)) ctop ctop cbot + ) + )) cstar cstar. + +Definition pmifK be bP Pe PP te tP fe fP := + capp ( + eif cType be + (clam cType ebool ( + cPi cProp (capp pbool (cvar 0)) ( + capps (S ⋅ S ⋅ PP) [ + cvar 1 ; + cvar 0 ; + perif (cvar 1) (S ⋅ S ⋅ Pe) (S ⋅ S ⋅ te) (S ⋅ S ⋅ fe) + ] + ) + )) + (clam cProp (capp pbool etrue) (S ⋅ tP)) + (clam cProp (capp pbool efalse) (S ⋅ fP)) + (clam cProp (capp pbool bool_err) ( + cbot_elim cType ( + capps (S ⋅ PP) [ + bool_err ; + cvar 0 ; + perif bool_err (S ⋅ Pe) (S ⋅ te) (S ⋅ fe) + ] + ) (pbool_bool_err_inv (cvar 0)) + )) + ) bP. + +Equations param_term (Γ : scope) (t : term) : cterm := { + ⟦ Γ | var x ⟧p := + match nth_error Γ x with + | Some m => cvar (if isProp m then vreg x else vpar x) + | None => cDummy + end ; + ⟦ Γ | Sort m i ⟧p := + if isKind m then pKind i + else if isProp m then pProp + else pType i ; + ⟦ Γ | Pi i j m mx A B ⟧p := + let Te := ⟦ Γ | Pi i j m mx A B ⟧pτ in + let Ae := ⟦ Γ | A ⟧pτ in + let Ap := ⟦ Γ | A ⟧p in + let Bp := ⟦ mx :: Γ | B ⟧p in + pmPi mx m Te Ae Ap Bp ; + ⟦ Γ | lam mx A t ⟧p := + if isProp mx then clam cProp ⟦ Γ | A ⟧p (close ⟦ mx :: Γ | t ⟧p) + else ( + let cm := if isKind mx then cType else cProp in + plam cm ⟦ Γ | A ⟧pτ ⟦ Γ | A ⟧p ⟦ mx :: Γ | t ⟧p + ) ; + ⟦ Γ | app u v ⟧p := + if relm (md Γ v) then capp (capp ⟦ Γ | u ⟧p ⟦ Γ | v ⟧pε) ⟦ Γ | v ⟧p + else if isGhost (md Γ v) then capp (capp ⟦ Γ | u ⟧p ⟦ Γ | v ⟧pv) ⟦ Γ | v ⟧p + else capp ⟦ Γ | u ⟧p ⟦ Γ | v ⟧p + ; + ⟦ Γ | Erased A ⟧p := + if isKind (md Γ A) then ⟦ Γ | A ⟧p else cDummy ; + ⟦ Γ | hide t ⟧p := + if isType (md Γ t) then ⟦ Γ | t ⟧p else cDummy ; + ⟦ Γ | reveal t P p ⟧p := + if relm (md Γ p) then cDummy + else capp (capp ⟦ Γ | p ⟧p ⟦ Γ | t ⟧pv) ⟦ Γ | t ⟧p ; + ⟦ Γ | Reveal t p ⟧p := + if isKind (md Γ p) then capp (capp ⟦ Γ | p ⟧p ⟦ Γ | t ⟧pv) ⟦ Γ | t ⟧p + else cDummy ; + ⟦ Γ | toRev t p u ⟧p := ⟦ Γ | u ⟧p ; + ⟦ Γ | fromRev t p u ⟧p := ⟦ Γ | u ⟧p ; + ⟦ Γ | gheq A u v ⟧p := squash (teq ⟦ Γ | A ⟧pτ ⟦ Γ | u ⟧pv ⟦ Γ | v ⟧pv) ; + ⟦ Γ | ghrefl A u ⟧p := sq (trefl ⟦ Γ | A ⟧pτ ⟦ Γ | u ⟧pv) ; + ⟦ Γ | ghcast A u v e P t ⟧p := + let eP := ⟦ Γ | e ⟧p in + let PP := ⟦ Γ | P ⟧p in + let uv := ⟦ Γ | u ⟧pv in + let vv := ⟦ Γ | v ⟧pv in + let vP := ⟦ Γ | v ⟧p in + let Ae := ⟦ Γ | A ⟧pτ in + let AP := ⟦ Γ | A ⟧p in + let te := ⟦ Γ | t ⟧pε in + let tv := ⟦ Γ | t ⟧pv in + let tP := ⟦ Γ | t ⟧p in + match md Γ t with + | mKind => cDummy + | mType => pcastTG Ae AP uv vv vP eP PP te tP + | mGhost => pcastTG Ae AP uv vv vP eP PP tv tP + | mProp => pcastP Ae AP uv vv vP eP PP tP + end ; + ⟦ Γ | tbool ⟧p := pbool ; + ⟦ Γ | ttrue ⟧p := ptrue ; + ⟦ Γ | tfalse ⟧p := pfalse ; + ⟦ Γ | tif m b P t f ⟧p := + let be := ⟦ Γ | b ⟧pε in + let bP := ⟦ Γ | b ⟧p in + let Pe := ⟦ Γ | P ⟧pε in + let PP := ⟦ Γ | P ⟧p in + let te := ⟦ Γ | t ⟧pε in + let tv := ⟦ Γ | t ⟧pv in + let tP := ⟦ Γ | t ⟧p in + let fe := ⟦ Γ | f ⟧pε in + let fv := ⟦ Γ | f ⟧pv in + let fP := ⟦ Γ | f ⟧p in + match m with + | mKind => pmifK be bP Pe PP te tP fe fP + | mType => pmif bP Pe PP te tP fe fP + | mGhost => pmif bP Pe PP tv tP fv fP + | mProp => pif bP PP tP fP + end ; + ⟦ Γ | tnat ⟧p := pnat ; + ⟦ Γ | tzero ⟧p := pzero ; + ⟦ Γ | tsucc n ⟧p := psucc ⟦ Γ | n ⟧p ; + ⟦ Γ | tnat_elim m n P z s ⟧p := + let ne := ⟦ Γ | n ⟧pε in + let nP := ⟦ Γ | n ⟧p in + let Pe := ⟦ Γ | P ⟧pε in + let PP := ⟦ Γ | P ⟧p in + let ze := ⟦ Γ | z ⟧pε in + let zv := ⟦ Γ | z ⟧pv in + let zP := ⟦ Γ | z ⟧p in + let se := ⟦ Γ | s ⟧pε in + let sv := ⟦ Γ | s ⟧pv in + let sP := ⟦ Γ | s ⟧p in + match m with + | mKind => cDummy + | mType => pnat_elim ne nP Pe PP ze zP se sP + | mGhost => pnat_elim ne nP Pe PP zv zP sv sP + | mProp => pnat_elimP ne nP Pe PP zP sP + end ; + ⟦ Γ | tvec A n ⟧p := pvec ⟦ Γ | A ⟧pε ⟦ Γ | A ⟧p ⟦ Γ | n ⟧pv ⟦ Γ | n ⟧p ; + ⟦ Γ | tvnil A ⟧p := pvnil ⟦ Γ | A ⟧p ; + ⟦ Γ | tvcons a n v ⟧p := pvcons ⟦ Γ | a ⟧p ⟦ Γ | n ⟧p ⟦ Γ | v ⟧p ; + ⟦ Γ | tvec_elim m A n v P z s ⟧p := + let Ae := ⟦ Γ | A ⟧pε in + let AP := ⟦ Γ | A ⟧p in + let nv := ⟦ Γ | n ⟧pv in + let nP := ⟦ Γ | n ⟧p in + let ve := ⟦ Γ | v ⟧pε in + let vP := ⟦ Γ | v ⟧p in + let Pe := ⟦ Γ | P ⟧pε in + let PP := ⟦ Γ | P ⟧p in + let ze := ⟦ Γ | z ⟧pε in + let zv := ⟦ Γ | z ⟧pv in + let zP := ⟦ Γ | z ⟧p in + let se := ⟦ Γ | s ⟧pε in + let sv := ⟦ Γ | s ⟧pv in + let sP := ⟦ Γ | s ⟧p in + match m with + | mKind => cDummy + | mType => pvec_elim Ae AP nv nP ve vP Pe PP ze zP se sP + | mGhost => pvec_elimG Ae AP nv nP ve vP Pe PP zv zP sv sP + | mProp => pvec_elimP Ae AP nv nP ve vP Pe PP zP sP + end ; + ⟦ Γ | bot ⟧p := cbot ; + ⟦ Γ | bot_elim m A p ⟧p := + if isProp m then cbot_elim cProp ⟦ Γ | A ⟧p ⟦ Γ | p ⟧p + else if isKind m then cbot_elim cType (capp ⟦ Γ | A ⟧p ⟦ Γ | A ⟧p∅) ⟦ Γ | p ⟧p + else cbot_elim cProp (capp ⟦ Γ | A ⟧p ⟦ Γ | A ⟧p∅) ⟦ Γ | p ⟧p +} +where "⟦ G | u '⟧p'" := (param_term G u). + +Reserved Notation "⟦ G '⟧p'" (at level 9, G at next level). + diff --git a/theories/Param.v b/theories/param/Typing.v similarity index 70% rename from theories/Param.v rename to theories/param/Typing.v index 75c9fcf..03dc944 100644 --- a/theories/Param.v +++ b/theories/param/Typing.v @@ -1,5 +1,3 @@ -(*** Parametricity ***) - From Coq Require Import Utf8 List Bool Lia. From Equations Require Import Equations. From GhostTT.autosubst Require Import CCAST GAST core unscoped. @@ -7,6 +5,8 @@ From GhostTT Require Import Util BasicAST CastRemoval SubstNotations ContextDecl CScoping Scoping CTyping TermMode Typing BasicMetaTheory CCMetaTheory Erasure Revival. From Coq Require Import Setoid Morphisms Relation_Definitions. +From GhostTT.param Require Export Term Scope. +From GhostTT.param Require Import Renaming Substitution Conversion Cast. Import ListNotations. Import CombineNotations. @@ -14,2571 +14,7 @@ Import CombineNotations. Set Default Goal Selector "!". Set Equations Transparent. -(** Variables and parametricity - - x : A in the context is translated to x : A, xP : AP when A is not a Prop. - When x : P : Prop then x is translated to only one variable. To keep the - context regular we will still make use of our flexible contexts. - Variables are then either odd and regular or even and correspond to - parametricity. - -**) - -Definition vreg x := S (x * 2). -Definition vpar x := x * 2. - -(** Lifting erasure and revival - - Because erasure and revival produce terms in ⟦ Γ ⟧ε and ⟦ Γ ⟧v respectively - when we expect ⟦ Γ ⟧p, we need to do some lifting. Thankfully this lifting - can be done simply by using vreg as a renaming. - We define handy notations to make it all work. - -**) - -Definition epm_lift (t : cterm) := - vreg ⋅ t. - -Definition rpm_lift (t : cterm) := - vreg ⋅ t. - -Notation "⟦ G | u '⟧pε'":= - (epm_lift ⟦ G | u ⟧ε) (at level 9, G, u at next level). - -Notation "⟦ G | u '⟧pτ'":= - (epm_lift ⟦ G | u ⟧τ) (at level 9, G, u at next level). - -Notation "⟦ G | u '⟧p∅'":= - (epm_lift ⟦ G | u ⟧∅) (at level 9, G, u at next level). - -Notation "⟦ G | u '⟧pv'":= - (rpm_lift ⟦ G | u ⟧v) (at level 9, G, u at next level). - -(** Parametricity translation - - We start by defining useful shorthands. - -**) - -Definition pKind i := - clam cType (cty i) (cEl (cvar 0) ⇒[ cType ] cSort cType i). - -Definition pType i := - clam cType (cty i) (cEl (cvar 0) ⇒[ cType ] cSort cProp 0). - -Definition pProp := - clam cType cunit (cSort cProp 0). - -(* ∀ (x : A) (x@mp : B x). C *) -Definition pPi mp A B C := - cPi cType A (cPi mp (capp (S ⋅ B) (cvar 0)) C). - -Definition plam mp A B t := - clam cType A (clam mp (capp (S ⋅ B) (cvar 0)) t). - -Definition pcastTG Ae AP uv vv vP eP PP te tP := - sq_elim - eP - (clam cProp (squash (teq Ae uv vv)) (S ⋅ (capp (capp (capp PP vv) vP) te))) - (clam cType (teq Ae uv vv) ( - capp - (tJ - (cvar 0) - (S ⋅ (clam cType Ae ( - clam cType (teq (S ⋅ Ae) (S ⋅ uv) (cvar 0)) ( - cPi cProp (capp (plus 2 ⋅ AP) (cvar 1)) ( - capp (capp (capp (plus 3 ⋅ PP) (cvar 2)) (cvar 0)) (plus 3 ⋅ te) - ) - ) - ))) - (S ⋅ (clam cProp (capp AP uv) (S ⋅ tP)))) - (S ⋅ vP) - )). - -Definition pcastP Ae AP uv vv vP eP PP tP := - sq_elim - eP - (clam cProp (squash (teq Ae uv vv)) (S ⋅ (capp (capp PP vv) vP))) - (clam cType (teq Ae uv vv) ( - capp - (tJ - (cvar 0) - (S ⋅ (clam cType Ae ( - clam cType (teq (S ⋅ Ae) (S ⋅ uv) (cvar 0)) ( - cPi cProp (capp (plus 2 ⋅ AP) (cvar 1)) ( - capp (capp (plus 3 ⋅ PP) (cvar 2)) (cvar 0) - ) - ) - ))) - (S ⋅ (clam cProp (capp AP uv) (S ⋅ tP)))) - (S ⋅ vP) - )). - -Reserved Notation "⟦ G | u '⟧p'" (at level 9, G, u at next level). - -(** Translation of Pi types, to factorise a bit **) - -(* Prop case *) -Definition pmPiP mx Ae Ap Bp := - if isProp mx then cPi cProp Ap (close Bp) - else if isKind mx then pPi cType Ae Ap Bp - else pPi cProp Ae Ap Bp. - -(* Non-prop case *) -Definition pmPiNP mx m Te Ae Ap Bp := - let cm := if isKind mx then cType else cProp in - clam cType Te ( - if isProp mx then cPi cProp (S ⋅ Ap) (capp ((up_ren S) ⋅ (close Bp)) (cvar 1)) - else ( - pPi cm (S ⋅ Ae) (S ⋅ Ap) (capp ((up_ren (up_ren S)) ⋅ Bp) ( - if isGhost mx && relm m then cvar 2 - else capp (cvar 2) (cvar 1) - )) - ) - ). - -(* General case *) -Definition pmPi mx m Te Ae Ap Bp := - if isProp m then pmPiP mx Ae Ap Bp - else pmPiNP mx m Te Ae Ap Bp. - -(* Parametricity for if *) - -Definition perif be Pe te fe := - eif cType be - (clam cType ebool (cEl (capp (S ⋅ Pe) (cvar 0)))) - te fe (cErr (capp Pe bool_err)). - -Definition pmif bP Pe PP te tP fe fP := - pif bP - (plam cProp ebool pbool ( - capp - (capp (capp (S ⋅ S ⋅ PP) (cvar 1)) (cvar 0)) - (S ⋅ (perif (cvar 0) (S ⋅ Pe) (S ⋅ te) (S ⋅ fe))) - )) - tP fP. - -(* Proves cbot from pbool bool_err *) -Definition pbool_bool_err_inv h := - pif h (clam cType ebool ( - clam cProp (capp pbool (cvar 0)) ( - eif cType (cvar 1) (clam cType ebool (cSort cProp 0)) ctop ctop cbot - ) - )) cstar cstar. - -Definition pmifK be bP Pe PP te tP fe fP := - capp ( - eif cType be - (clam cType ebool ( - cPi cProp (capp pbool (cvar 0)) ( - capps (S ⋅ S ⋅ PP) [ - cvar 1 ; - cvar 0 ; - perif (cvar 1) (S ⋅ S ⋅ Pe) (S ⋅ S ⋅ te) (S ⋅ S ⋅ fe) - ] - ) - )) - (clam cProp (capp pbool etrue) (S ⋅ tP)) - (clam cProp (capp pbool efalse) (S ⋅ fP)) - (clam cProp (capp pbool bool_err) ( - cbot_elim cType ( - capps (S ⋅ PP) [ - bool_err ; - cvar 0 ; - perif bool_err (S ⋅ Pe) (S ⋅ te) (S ⋅ fe) - ] - ) (pbool_bool_err_inv (cvar 0)) - )) - ) bP. - -Equations param_term (Γ : scope) (t : term) : cterm := { - ⟦ Γ | var x ⟧p := - match nth_error Γ x with - | Some m => cvar (if isProp m then vreg x else vpar x) - | None => cDummy - end ; - ⟦ Γ | Sort m i ⟧p := - if isKind m then pKind i - else if isProp m then pProp - else pType i ; - ⟦ Γ | Pi i j m mx A B ⟧p := - let Te := ⟦ Γ | Pi i j m mx A B ⟧pτ in - let Ae := ⟦ Γ | A ⟧pτ in - let Ap := ⟦ Γ | A ⟧p in - let Bp := ⟦ mx :: Γ | B ⟧p in - pmPi mx m Te Ae Ap Bp ; - ⟦ Γ | lam mx A t ⟧p := - if isProp mx then clam cProp ⟦ Γ | A ⟧p (close ⟦ mx :: Γ | t ⟧p) - else ( - let cm := if isKind mx then cType else cProp in - plam cm ⟦ Γ | A ⟧pτ ⟦ Γ | A ⟧p ⟦ mx :: Γ | t ⟧p - ) ; - ⟦ Γ | app u v ⟧p := - if relm (md Γ v) then capp (capp ⟦ Γ | u ⟧p ⟦ Γ | v ⟧pε) ⟦ Γ | v ⟧p - else if isGhost (md Γ v) then capp (capp ⟦ Γ | u ⟧p ⟦ Γ | v ⟧pv) ⟦ Γ | v ⟧p - else capp ⟦ Γ | u ⟧p ⟦ Γ | v ⟧p - ; - ⟦ Γ | Erased A ⟧p := - if isKind (md Γ A) then ⟦ Γ | A ⟧p else cDummy ; - ⟦ Γ | hide t ⟧p := - if isType (md Γ t) then ⟦ Γ | t ⟧p else cDummy ; - ⟦ Γ | reveal t P p ⟧p := - if relm (md Γ p) then cDummy - else capp (capp ⟦ Γ | p ⟧p ⟦ Γ | t ⟧pv) ⟦ Γ | t ⟧p ; - ⟦ Γ | Reveal t p ⟧p := - if isKind (md Γ p) then capp (capp ⟦ Γ | p ⟧p ⟦ Γ | t ⟧pv) ⟦ Γ | t ⟧p - else cDummy ; - ⟦ Γ | toRev t p u ⟧p := ⟦ Γ | u ⟧p ; - ⟦ Γ | fromRev t p u ⟧p := ⟦ Γ | u ⟧p ; - ⟦ Γ | gheq A u v ⟧p := squash (teq ⟦ Γ | A ⟧pτ ⟦ Γ | u ⟧pv ⟦ Γ | v ⟧pv) ; - ⟦ Γ | ghrefl A u ⟧p := sq (trefl ⟦ Γ | A ⟧pτ ⟦ Γ | u ⟧pv) ; - ⟦ Γ | ghcast A u v e P t ⟧p := - let eP := ⟦ Γ | e ⟧p in - let PP := ⟦ Γ | P ⟧p in - let uv := ⟦ Γ | u ⟧pv in - let vv := ⟦ Γ | v ⟧pv in - let vP := ⟦ Γ | v ⟧p in - let Ae := ⟦ Γ | A ⟧pτ in - let AP := ⟦ Γ | A ⟧p in - let te := ⟦ Γ | t ⟧pε in - let tv := ⟦ Γ | t ⟧pv in - let tP := ⟦ Γ | t ⟧p in - match md Γ t with - | mKind => cDummy - | mType => pcastTG Ae AP uv vv vP eP PP te tP - | mGhost => pcastTG Ae AP uv vv vP eP PP tv tP - | mProp => pcastP Ae AP uv vv vP eP PP tP - end ; - ⟦ Γ | tbool ⟧p := pbool ; - ⟦ Γ | ttrue ⟧p := ptrue ; - ⟦ Γ | tfalse ⟧p := pfalse ; - ⟦ Γ | tif m b P t f ⟧p := - let be := ⟦ Γ | b ⟧pε in - let bP := ⟦ Γ | b ⟧p in - let Pe := ⟦ Γ | P ⟧pε in - let PP := ⟦ Γ | P ⟧p in - let te := ⟦ Γ | t ⟧pε in - let tv := ⟦ Γ | t ⟧pv in - let tP := ⟦ Γ | t ⟧p in - let fe := ⟦ Γ | f ⟧pε in - let fv := ⟦ Γ | f ⟧pv in - let fP := ⟦ Γ | f ⟧p in - match m with - | mKind => pmifK be bP Pe PP te tP fe fP - | mType => pmif bP Pe PP te tP fe fP - | mGhost => pmif bP Pe PP tv tP fv fP - | mProp => pif bP PP tP fP - end ; - ⟦ Γ | tnat ⟧p := pnat ; - ⟦ Γ | tzero ⟧p := pzero ; - ⟦ Γ | tsucc n ⟧p := psucc ⟦ Γ | n ⟧p ; - ⟦ Γ | tnat_elim m n P z s ⟧p := - let ne := ⟦ Γ | n ⟧pε in - let nP := ⟦ Γ | n ⟧p in - let Pe := ⟦ Γ | P ⟧pε in - let PP := ⟦ Γ | P ⟧p in - let ze := ⟦ Γ | z ⟧pε in - let zv := ⟦ Γ | z ⟧pv in - let zP := ⟦ Γ | z ⟧p in - let se := ⟦ Γ | s ⟧pε in - let sv := ⟦ Γ | s ⟧pv in - let sP := ⟦ Γ | s ⟧p in - match m with - | mKind => cDummy - | mType => pnat_elim ne nP Pe PP ze zP se sP - | mGhost => pnat_elim ne nP Pe PP zv zP sv sP - | mProp => pnat_elimP ne nP Pe PP zP sP - end ; - ⟦ Γ | tvec A n ⟧p := pvec ⟦ Γ | A ⟧pε ⟦ Γ | A ⟧p ⟦ Γ | n ⟧pv ⟦ Γ | n ⟧p ; - ⟦ Γ | tvnil A ⟧p := pvnil ⟦ Γ | A ⟧p ; - ⟦ Γ | tvcons a n v ⟧p := pvcons ⟦ Γ | a ⟧p ⟦ Γ | n ⟧p ⟦ Γ | v ⟧p ; - ⟦ Γ | tvec_elim m A n v P z s ⟧p := - let Ae := ⟦ Γ | A ⟧pε in - let AP := ⟦ Γ | A ⟧p in - let nv := ⟦ Γ | n ⟧pv in - let nP := ⟦ Γ | n ⟧p in - let ve := ⟦ Γ | v ⟧pε in - let vP := ⟦ Γ | v ⟧p in - let Pe := ⟦ Γ | P ⟧pε in - let PP := ⟦ Γ | P ⟧p in - let ze := ⟦ Γ | z ⟧pε in - let zv := ⟦ Γ | z ⟧pv in - let zP := ⟦ Γ | z ⟧p in - let se := ⟦ Γ | s ⟧pε in - let sv := ⟦ Γ | s ⟧pv in - let sP := ⟦ Γ | s ⟧p in - match m with - | mKind => cDummy - | mType => pvec_elim Ae AP nv nP ve vP Pe PP ze zP se sP - | mGhost => pvec_elimG Ae AP nv nP ve vP Pe PP zv zP sv sP - | mProp => pvec_elimP Ae AP nv nP ve vP Pe PP zP sP - end ; - ⟦ Γ | bot ⟧p := cbot ; - ⟦ Γ | bot_elim m A p ⟧p := - if isProp m then cbot_elim cProp ⟦ Γ | A ⟧p ⟦ Γ | p ⟧p - else if isKind m then cbot_elim cType (capp ⟦ Γ | A ⟧p ⟦ Γ | A ⟧p∅) ⟦ Γ | p ⟧p - else cbot_elim cProp (capp ⟦ Γ | A ⟧p ⟦ Γ | A ⟧p∅) ⟦ Γ | p ⟧p -} -where "⟦ G | u '⟧p'" := (param_term G u). - -Reserved Notation "⟦ G '⟧p'" (at level 9, G at next level). - -Equations param_ctx (Γ : context) : ccontext := { - ⟦ [] ⟧p := [] ; - ⟦ Γ ,, (mx, A) ⟧p := - if isProp mx then None :: Some (cProp, ⟦ sc Γ | A ⟧p) :: ⟦ Γ ⟧p - else if isKind mx then - Some (cType, capp (S ⋅ ⟦ sc Γ | A ⟧p) (cvar 0)) :: - Some (cType, ⟦ sc Γ | A ⟧pτ) :: ⟦ Γ ⟧p - else - Some (cProp, capp (S ⋅ ⟦ sc Γ | A ⟧p) (cvar 0)) :: - Some (cType, ⟦ sc Γ | A ⟧pτ) :: ⟦ Γ ⟧p -} -where "⟦ G '⟧p'" := (param_ctx G). - -Fixpoint param_sc (Γ : scope) : cscope := - match Γ with - | [] => [] - | m :: Γ => - if isProp m then None :: Some cProp :: param_sc Γ - else if isKind m then Some cType :: Some cType :: param_sc Γ - else Some cProp :: Some cType :: param_sc Γ - end. - -(** Scope lookups **) - -Lemma nth_error_param_vreg : - ∀ Γ x, - nth_error (param_sc Γ) (vreg x) = - option_map (λ m, if isProp m then Some cProp else Some cType) (nth_error Γ x). -Proof. - intros Γ x. - induction Γ as [| m Γ ih] in x |- *. - - destruct x. all: reflexivity. - - destruct x. - + cbn. destruct_ifs. all: reflexivity. - + unfold vreg. simpl "*". remember (S (x * 2)) as y eqn:e. - cbn. subst. destruct_ifs. all: eapply ih. -Qed. - -Lemma nth_error_param_vpar : - ∀ Γ x, - nth_error (param_sc Γ) (vpar x) = - option_map (λ m, - if isProp m then None - else if isKind m then Some cType - else Some cProp - ) (nth_error Γ x). -Proof. - intros Γ x. - induction Γ as [| m Γ ih] in x |- *. - - destruct x. all: reflexivity. - - destruct x. - + cbn. destruct_ifs. all: reflexivity. - + cbn. destruct_ifs. all: eapply ih. -Qed. - -(** ⟦ Γ ⟧v is a sub-context of ⟦ Γ ⟧p **) - -Lemma scoping_rev_sub_param : - ∀ Γ, - crscoping (param_sc Γ) vreg (revive_sc Γ). -Proof. - intros Γ. intros x m e. - unfold revive_sc in e. rewrite nth_error_map in e. - rewrite nth_error_param_vreg. - destruct (nth_error Γ x) as [mx|] eqn:ex. 2: discriminate. - cbn in e. cbn. - destruct_ifs. 1: discriminate. - assumption. -Qed. - -Hint Resolve scoping_rev_sub_param : cc_scope. - -Lemma typing_rev_sub_param : - ∀ Γ, - crtyping ⟦ Γ ⟧p vreg ⟦ Γ ⟧v. -Proof. - intros Γ. intros x m A e. - assert (h : nth_error ⟦ Γ ⟧p (vreg x) = Some (Some (m, vreg ⋅ A))). - { induction Γ as [| [my B] Γ ih] in x, m, A, e |- *. - 1:{ destruct x. all: discriminate. } - destruct x. - - cbn in e. - destruct (isProp my) eqn:ey. 1: discriminate. - noconf e. cbn. rewrite ey. - destruct_if e1. all: reflexivity. - - cbn in e. - unfold vreg. simpl "*". remember (S (x * 2)) as z eqn:ez. - cbn. subst. - destruct_if ey. - + eapply ih. assumption. - + destruct_if e1. - * eapply ih. assumption. - * eapply ih. assumption. - } - eexists. split. - - eassumption. - - ssimpl. unfold vreg. cbn. ssimpl. eapply extRen_cterm. - intro. ssimpl. lia. -Qed. - -(** ⟦ Γ ⟧ε is a sub-context of ⟦ Γ ⟧p **) - -Lemma scoping_er_sub_param : - ∀ Γ, - crscoping (param_sc Γ) vreg (erase_sc Γ). -Proof. - intros Γ. - pose proof (scoping_sub_rev Γ) as lem. - eapply crscoping_comp in lem. 2: eapply scoping_rev_sub_param. - eapply crscoping_morphism. all: eauto. - intros x. reflexivity. -Qed. - -Hint Resolve scoping_er_sub_param : cc_scope. - -Lemma typing_er_sub_param : - ∀ Γ, - crtyping ⟦ Γ ⟧p vreg ⟦ Γ ⟧ε. -Proof. - intros Γ. - pose proof (typing_sub_rev Γ) as lem. - eapply crtyping_comp in lem. 2: eapply typing_rev_sub_param. - eapply crtyping_morphism. all: eauto. - intros x. reflexivity. -Qed. - -(** Parametricity preserves scoping **) - -Lemma scoping_epm_lift : - ∀ Γ Γ' t, - ccscoping (erase_sc Γ) t cType → - Γ' = param_sc Γ → - ccscoping Γ' (epm_lift t) cType. -Proof. - intros Γ Γ' t h ->. - eapply cscoping_ren. - - eapply scoping_er_sub_param. - - assumption. -Qed. - -(* Hint Resolve scoping_epm_lift | 1000 : cc_scope. *) - -Lemma pscoping_erase_term : - ∀ Γ Γ' t, - Γ' = param_sc Γ → - ccscoping Γ' ⟦ Γ | t ⟧pε cType. -Proof. - intros Γ Γ' t ->. - eapply scoping_epm_lift. - - eapply erase_scoping. - - reflexivity. -Qed. - -Hint Resolve pscoping_erase_term : cc_scope. - -Lemma pscoping_erase_type : - ∀ Γ Γ' t, - Γ' = param_sc Γ → - ccscoping Γ' ⟦ Γ | t ⟧pτ cType. -Proof. - intros Γ Γ' t ->. - eapply scoping_epm_lift. - - constructor. eapply erase_scoping. - - reflexivity. -Qed. - -Hint Resolve pscoping_erase_type : cc_scope. - -Lemma pscoping_erase_err : - ∀ Γ Γ' t, - Γ' = param_sc Γ → - ccscoping Γ' ⟦ Γ | t ⟧p∅ cType. -Proof. - intros Γ Γ' t ->. - eapply scoping_epm_lift. - - constructor. eapply erase_scoping. - - reflexivity. -Qed. - -Hint Resolve pscoping_erase_err : cc_scope. - -Lemma scoping_rpm_lift : - ∀ Γ Γ' t, - ccscoping (revive_sc Γ) t cType → - Γ' = param_sc Γ → - ccscoping Γ' (rpm_lift t) cType. -Proof. - intros Γ Γ' t h ->. - eapply cscoping_ren. - - eapply scoping_rev_sub_param. - - assumption. -Qed. - -(* Hint Resolve scoping_rpm_lift | 1000 : cc_scope. *) - -Lemma pscoping_revive : - ∀ Γ Γ' t, - Γ' = param_sc Γ → - ccscoping Γ' ⟦ Γ | t ⟧pv cType. -Proof. - intros Γ Γ' t ->. - eapply scoping_rpm_lift. - - eapply revive_scoping. - - reflexivity. -Qed. - -Hint Resolve pscoping_revive : cc_scope. - -Lemma erase_scoping_eq : - ∀ Γ Γ' t, - Γ' = erase_sc Γ → - ccscoping Γ' ⟦ Γ | t ⟧ε cType. -Proof. - intros Γ ? t ->. - eapply erase_scoping. -Qed. - -Lemma revive_scoping_eq : - ∀ Γ Γ' t, - Γ' = revive_sc Γ → - ccscoping Γ' ⟦ Γ | t ⟧v cType. -Proof. - intros Γ ? t ->. - eapply revive_scoping. -Qed. - -Hint Resolve erase_scoping_eq : cc_scope. -Hint Resolve revive_scoping_eq : cc_scope. -Hint Resolve revive_scoping : cc_scope. -Hint Resolve crscoping_plus : cc_scope. - -Lemma pPi_scoping : - ∀ Γ mx A B C, - ccscoping Γ A cType → - ccscoping Γ B cType → - ccscoping (Some mx :: Some cType :: Γ) C cType → - ccscoping Γ (pPi mx A B C) cType. -Proof. - intros Γ mx A B C hA hB hC. - unshelve eauto with cc_scope shelvedb ; shelve_unifiable. - constructor. reflexivity. -Qed. - -Hint Resolve pPi_scoping : cc_scope. - -(* So that they're not unfolded too eagerly *) -Opaque epm_lift rpm_lift. - -Lemma param_scoping : - ∀ Γ t m, - scoping Γ t m → - ccscoping (param_sc Γ) ⟦ Γ | t ⟧p (if isKind m then cType else cProp). -Proof. - intros Γ t m h. - induction h. - all: try solve [ cbn ; eauto with cc_scope ]. - all: try solve [ cbn ; destruct_ifs ; eauto with cc_scope ]. - - cbn. rewrite H. destruct_if e. - + mode_eqs. cbn. constructor. - rewrite nth_error_param_vreg. rewrite H. reflexivity. - + constructor. rewrite nth_error_param_vpar. rewrite H. - cbn. rewrite e. destruct_ifs. all: reflexivity. - - cbn. - destruct m, mx. all: cbn in *. - all: try solve [ typeclasses eauto 50 with cc_scope ]. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - 1:{ - eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - } - eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - reflexivity. - * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - reflexivity. - * eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - reflexivity. - * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - reflexivity. - * eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - * eapply crscoping_shift. eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - * eapply scoping_epm_lift. 2: reflexivity. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - reflexivity. - * eapply crscoping_shift. eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - - cbn in *. destruct_ifs. all: mode_eqs. all: cbn in *. - all: try solve [ typeclasses eauto 50 with cc_scope ]. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - - cbn in *. - erewrite scoping_md. 2: eassumption. - cbn. assumption. - - cbn in *. - erewrite scoping_md. 2: eassumption. - destruct_ifs. all: mode_eqs. all: try intuition discriminate. - 1:{ destruct m. all: intuition discriminate. } - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - reflexivity. - - cbn in *. - unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: reflexivity. - - cbn in *. - erewrite scoping_md. 2: eassumption. - destruct m. 1: contradiction. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - all: repeat try eapply crscoping_shift. - all: eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - all: repeat try eapply crscoping_shift. - all: eauto with cc_scope. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - all: try reflexivity. - all: repeat try eapply crscoping_shift. - all: eauto with cc_scope. - - cbn in *. - destruct m. - + cbn in *. escope. all: reflexivity. - + cbn in *. escope. all: reflexivity. - + cbn in *. escope. all: reflexivity. - + cbn in *. escope. - - cbn in *. - destruct m. - + contradiction. - + cbn in *. escope. all: reflexivity. - + cbn in *. escope. all: reflexivity. - + cbn in *. escope. all: reflexivity. - - cbn in *. - destruct m. - + contradiction. - + cbn in *. escope. all: reflexivity. - + cbn in *. escope. all: reflexivity. - + cbn in *. escope. all: reflexivity. - - cbn in *. - destruct_ifs. all: mode_eqs. all: try discriminate. - all: try solve [ typeclasses eauto 50 with cc_scope ]. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - reflexivity. - + unshelve typeclasses eauto 50 with cc_scope shelvedb ; shelve_unifiable. - reflexivity. -Qed. - -(** Parametricity commutes with renaming - - For this we define pren ρ which is basically the following function: - pren ρ (2 * n) = 2 * (ρ n) - pren ρ (2 * n + 1) = 2 * (ρ n) + 1 - -**) - -Definition pren (ρ : nat → nat) := - λ n, PeanoNat.Nat.b2n (Nat.odd n) + 2 * ρ (Nat.div2 n). - -Lemma pren_even : - ∀ ρ n, - pren ρ (n * 2) = (ρ n) * 2. -Proof. - intros ρ n. - unfold pren. - replace (n * 2) with (2 * n) by lia. - rewrite PeanoNat.Nat.div2_double. - rewrite PeanoNat.Nat.odd_mul. cbn. lia. -Qed. - -Lemma pren_odd : - ∀ ρ n, - pren ρ (S (n * 2)) = S ((ρ n) * 2). -Proof. - intros ρ n. - unfold pren. - replace (n * 2) with (2 * n) by lia. - rewrite PeanoNat.Nat.div2_succ_double. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_mul. cbn. lia. -Qed. - -Lemma div2_SS : - ∀ n, Nat.div2 (S (S n)) = S (Nat.div2 n). -Proof. - intro n. - destruct (PeanoNat.Nat.Even_Odd_dec n) as [h | h]. - - rewrite <- PeanoNat.Nat.Odd_div2. - 2:{ apply PeanoNat.Nat.Odd_succ. assumption. } - rewrite <- PeanoNat.Nat.Even_div2. 2: assumption. - reflexivity. - - rewrite <- PeanoNat.Nat.Even_div2. - 2:{ apply PeanoNat.Nat.Even_succ. assumption. } - rewrite <- PeanoNat.Nat.Odd_div2. 2: assumption. - reflexivity. -Qed. - -Lemma pren_SS : - ∀ ρ n, pren ρ (S (S n)) = pren (S >> ρ) n. -Proof. - intros ρ n. - unfold pren. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_succ. - rewrite div2_SS. reflexivity. -Qed. - -Lemma pren_comp_S : - ∀ ρ n, pren (ρ >> S) n = S (S (pren ρ n)). -Proof. - intros ρ n. - unfold pren. ssimpl. lia. -Qed. - -Lemma pren_id : - ∀ n, pren id n = n. -Proof. - intros n. - unfold pren. - rewrite PeanoNat.Nat.div2_div. - symmetry. etransitivity. 1: eapply PeanoNat.Nat.div_mod_eq with (y := 2). - rewrite <- PeanoNat.Nat.bit0_mod. - rewrite PeanoNat.Nat.bit0_odd. - unfold id. unfold Datatypes.id. lia. -Qed. - -Transparent epm_lift rpm_lift. - -Lemma pren_epm_lift : - ∀ ρ t, - pren ρ ⋅ epm_lift t = epm_lift (ρ ⋅ t). -Proof. - intros ρ t. - unfold epm_lift. ssimpl. - eapply extRen_cterm. intro x. - unfold vreg, pren. ssimpl. - replace (x * 2) with (2 * x) by lia. - rewrite PeanoNat.Nat.div2_succ_double. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_mul. cbn. lia. -Qed. - -Lemma pren_rpm_lift : - ∀ ρ t, - pren ρ ⋅ rpm_lift t = rpm_lift (ρ ⋅ t). -Proof. - intros ρ t. - apply pren_epm_lift. -Qed. - -Opaque epm_lift rpm_lift. - -Ltac cEl_ren := - change (cEl (?ρ ⋅ ?t)) with (ρ ⋅ (cEl t)). - -Lemma param_ren : - ∀ Γ Δ ρ t, - rscoping Γ ρ Δ → - rscoping_comp Γ ρ Δ → - ⟦ Γ | ρ ⋅ t ⟧p = (pren ρ) ⋅ ⟦ Δ | t ⟧p. -Proof. - intros Γ Δ ρ t hρ hcρ. - induction t in Γ, Δ, ρ, hρ, hcρ |- *. - - cbn. - destruct (nth_error Δ n) eqn:e. - + eapply hρ in e as e'. rewrite e'. - destruct_if e1. - * unfold vreg, pren. ssimpl. f_equal. - replace (n * 2) with (2 * n) by lia. - rewrite PeanoNat.Nat.div2_succ_double. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_mul. cbn. lia. - * unfold pren, vpar. ssimpl. f_equal. - replace (n * 2) with (2 * n) by lia. - rewrite PeanoNat.Nat.div2_double. - rewrite PeanoNat.Nat.odd_mul. cbn. lia. - + eapply hcρ in e as e'. rewrite e'. reflexivity. - - cbn. destruct_ifs. all: reflexivity. - - cbn. unfold pmPi, pmPiP, pmPiNP, pPi. - erewrite IHt1. 2,3: eassumption. - erewrite IHt2. - 2:{ eapply rscoping_upren. eassumption. } - 2:{ eapply rscoping_comp_upren. eassumption. } - erewrite ?erase_ren, ?revive_ren. - 2-5: eauto using rscoping_upren, rscoping_comp_upren. - rewrite ?pren_epm_lift, ?pren_rpm_lift. - destruct m, m0. all: cbn in *. (* all: try reflexivity. *) - + f_equal. all: f_equal. - 1:{ rewrite pren_epm_lift. cbn. reflexivity. } - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. all: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ rewrite pren_epm_lift. cbn. reflexivity. } - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. all: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ - rewrite pren_epm_lift. cbn. f_equal. - unfold close. ssimpl. reflexivity. - } - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. all: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ - rewrite pren_epm_lift. cbn. f_equal. - unfold close. ssimpl. reflexivity. - } - 1:{ ssimpl. reflexivity. } - f_equal. - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ rewrite pren_epm_lift. cbn. reflexivity. } - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. all: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ rewrite pren_epm_lift. cbn. reflexivity. } - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. all: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ - rewrite pren_epm_lift. cbn. f_equal. - unfold close. ssimpl. reflexivity. - } - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. all: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ - rewrite pren_epm_lift. cbn. f_equal. - unfold close. ssimpl. reflexivity. - } - 1:{ ssimpl. reflexivity. } - f_equal. - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ rewrite pren_epm_lift. cbn. reflexivity. } - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. all: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ rewrite pren_epm_lift. cbn. reflexivity. } - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. all: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ - rewrite pren_epm_lift. cbn. f_equal. - unfold close. ssimpl. reflexivity. - } - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. all: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ - rewrite pren_epm_lift. cbn. f_equal. - unfold close. ssimpl. reflexivity. - } - 1:{ ssimpl. reflexivity. } - f_equal. - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ rewrite pren_epm_lift. cbn. reflexivity. } - 1:{ ssimpl. reflexivity. } - ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + f_equal. all: f_equal. - 1:{ rewrite pren_epm_lift. reflexivity. } - 1:{ ssimpl. reflexivity. } - ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. reflexivity. - + f_equal. all: f_equal. - 1:{ rewrite pren_epm_lift. reflexivity. } - 1:{ ssimpl. reflexivity. } - ssimpl. eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. reflexivity. - + f_equal. unfold close. ssimpl. - eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - - cbn. - erewrite IHt1. 2,3: eassumption. - erewrite IHt2. - 2:{ eapply rscoping_upren. eassumption. } - 2:{ eapply rscoping_comp_upren. eassumption. } - unfold plam. - erewrite erase_ren. 2,3: eassumption. - destruct_ifs. all: mode_eqs. - + cbn. unfold close. ssimpl. f_equal. - eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. cbn. reflexivity. - + cbn. rewrite pren_epm_lift. ssimpl. f_equal. f_equal. - eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. reflexivity. - + cbn. rewrite pren_epm_lift. ssimpl. f_equal. f_equal. - eapply extRen_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite pren_SS. ssimpl. rewrite pren_comp_S. reflexivity. - - cbn. - erewrite md_ren. 2,3: eassumption. - erewrite IHt1. 2,3: eassumption. - erewrite IHt2. 2,3: eassumption. - erewrite erase_ren. 2,3: eassumption. - erewrite revive_ren. 2,3: eassumption. - rewrite <- pren_epm_lift. rewrite <- pren_rpm_lift. - destruct_ifs. all: reflexivity. - - cbn. - erewrite md_ren. 2,3: eassumption. - destruct_ifs. all: eauto. - - cbn. - erewrite md_ren. 2,3: eassumption. - destruct_ifs. all: eauto. - - cbn. - erewrite md_ren. 2,3: eassumption. - destruct_ifs. 1: reflexivity. - cbn. erewrite IHt3. 2,3: eassumption. - erewrite IHt1. 2,3: eassumption. - erewrite ?erase_ren, ?revive_ren. 2,3: eassumption. - rewrite !pren_rpm_lift. reflexivity. - - cbn. - erewrite md_ren. 2,3: eassumption. - destruct_ifs. 2: reflexivity. - cbn. erewrite IHt2. 2,3: eassumption. - erewrite IHt1. 2,3: eassumption. - erewrite ?erase_ren, ?revive_ren. 2,3: eassumption. - rewrite !pren_rpm_lift. reflexivity. - - cbn. eauto. - - cbn. eauto. - - cbn. - erewrite ?erase_ren, ?revive_ren. 2-7: eassumption. - rewrite ?pren_epm_lift. reflexivity. - - cbn. - erewrite ?erase_ren, ?revive_ren. 2-5: eassumption. - rewrite ?pren_epm_lift. reflexivity. - - cbn. - erewrite md_ren. 2,3: eassumption. - erewrite ?erase_ren, ?revive_ren. 2-11: eassumption. - destruct (md _ _). - + eauto. - + unfold pcastTG. cbn. - erewrite IHt1. 2,3: eassumption. - erewrite IHt3. 2,3: eassumption. - erewrite IHt4. 2,3: eassumption. - erewrite IHt5. 2,3: eassumption. - erewrite IHt6. 2,3: eassumption. - rewrite ?pren_epm_lift, ?pren_rpm_lift. - f_equal. all: f_equal. all: f_equal. - 2:{ rewrite <- pren_epm_lift. ssimpl. reflexivity. } - 3:{ ssimpl. reflexivity. } - all: f_equal. - 2:{ ssimpl. reflexivity. } - all: f_equal. - 1:{ ssimpl. reflexivity. } - 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - 1:{ cEl_ren. rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - 3:{ ssimpl. reflexivity. } - all: f_equal. - 3:{ ssimpl. reflexivity. } - 3:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - all: f_equal. - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - all: f_equal. - 1:{ ssimpl. reflexivity. } - 2:{ rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. f_equal. - ssimpl. reflexivity. - + unfold pcastTG. cbn. - erewrite IHt1. 2,3: eassumption. - erewrite IHt3. 2,3: eassumption. - erewrite IHt4. 2,3: eassumption. - erewrite IHt5. 2,3: eassumption. - erewrite IHt6. 2,3: eassumption. - rewrite ?pren_epm_lift, ?pren_rpm_lift. - f_equal. all: f_equal. all: f_equal. - 2:{ rewrite <- pren_epm_lift. ssimpl. reflexivity. } - 3:{ ssimpl. reflexivity. } - all: f_equal. - 2:{ ssimpl. reflexivity. } - all: f_equal. - 1:{ ssimpl. reflexivity. } - 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - 1:{ cEl_ren. rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - 3:{ ssimpl. reflexivity. } - all: f_equal. - 3:{ ssimpl. reflexivity. } - 3:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - all: f_equal. - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - all: f_equal. - 1:{ ssimpl. reflexivity. } - 2:{ rewrite <- pren_epm_lift. ssimpl. reflexivity. } - f_equal. f_equal. - ssimpl. reflexivity. - + unfold pcastP. cbn. - erewrite IHt1. 2,3: eassumption. - erewrite IHt3. 2,3: eassumption. - erewrite IHt4. 2,3: eassumption. - erewrite IHt5. 2,3: eassumption. - erewrite IHt6. 2,3: eassumption. - rewrite ?pren_epm_lift, ?pren_rpm_lift. - f_equal. all: f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 3:{ ssimpl. reflexivity. } - all: f_equal. - 1:{ ssimpl. reflexivity. } - 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - all: f_equal. - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - 3:{ ssimpl. reflexivity. } - all: f_equal. - 3:{ ssimpl. reflexivity. } - 3:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - all: f_equal. - 1:{ cEl_ren. rewrite <- pren_epm_lift. ssimpl. reflexivity. } - 1:{ rewrite <- pren_rpm_lift. ssimpl. reflexivity. } - all: f_equal. - 1:{ ssimpl. reflexivity. } - f_equal. - ssimpl. reflexivity. - - cbn. reflexivity. - - cbn. reflexivity. - - cbn. reflexivity. - - cbn. - erewrite IHt1, IHt2, IHt3, IHt4. 2-9: eassumption. - erewrite ?erase_ren, ?revive_ren. 2-13: eassumption. - rewrite <- !pren_epm_lift. - change (epm_lift ⟦ ?G | ?t ⟧v) with (⟦ G | t⟧pv). - destruct m. 4: reflexivity. - + unfold pmifK. unfold perif. cbn. f_equal. f_equal. - all: f_equal. 1,4: f_equal. - 1,2: f_equal. 1-4: f_equal. 1,2,5-7,10: f_equal. - 2,3,5,6: f_equal. 2,4: f_equal. - all: ssimpl. all: reflexivity. - + cbn. unfold pmif, plam. cbn. f_equal. f_equal. f_equal. - ssimpl. reflexivity. - + cbn. unfold pmif, plam. cbn. f_equal. f_equal. f_equal. - ssimpl. reflexivity. - - cbn. reflexivity. - - cbn. reflexivity. - - cbn. f_equal. eauto. - - cbn. - erewrite IHt1, IHt2, IHt3, IHt4. 2-9: eassumption. - erewrite ?erase_ren, ?revive_ren. 2-13: eassumption. - rewrite <- !pren_epm_lift. - change (epm_lift ⟦ ?G | ?t ⟧v) with (⟦ G | t⟧pv). - destruct m. all: reflexivity. - - cbn. erewrite IHt1, IHt2. 2-5: eassumption. - erewrite ?erase_ren, ?revive_ren. 2-5: eassumption. - rewrite <- !pren_epm_lift. - reflexivity. - - cbn. erewrite IHt. 2-3: eassumption. - reflexivity. - - cbn. erewrite IHt1, IHt2, IHt3. 2-7: eassumption. - reflexivity. - - cbn. - erewrite IHt1, IHt2, IHt3, IHt4, IHt5, IHt6. 2-13: eassumption. - erewrite ?erase_ren, ?revive_ren. 2-17: eassumption. - rewrite <- !pren_epm_lift. - change (epm_lift ⟦ ?G | ?t ⟧v) with (⟦ G | t⟧pv). - destruct m. all: reflexivity. - - cbn. reflexivity. - - cbn. destruct_ifs. all: mode_eqs. - + cbn. f_equal. all: eauto. - + cbn. f_equal. 1: f_equal. all: eauto. - erewrite erase_ren. 2,3: eassumption. - rewrite pren_epm_lift. reflexivity. - + cbn. f_equal. 1: f_equal. all: eauto. - erewrite erase_ren. 2,3: eassumption. - rewrite pren_epm_lift. reflexivity. -Qed. - -(** Parametricity commutes with substitution - - As for revival we need to craft a new substitution that gets the scopes as - input in order to determine the mode of the various variables. - -**) - -Definition psubst Δ Γ σ n := - let p := Nat.div2 n in - match nth_error Δ p with - | Some m => - if relm m then ( - if Nat.odd n then ⟦ Γ | σ p ⟧pε - else ⟦ Γ | σ p ⟧p - ) - else if isGhost m then ( - if Nat.odd n then ⟦ Γ | σ p ⟧pv - else ⟦ Γ | σ p ⟧p - ) - else ( - if Nat.odd n then ⟦ Γ | σ p ⟧p - else cDummy - ) - | None => cDummy - end. - -Lemma div2_vreg : - ∀ n, Nat.div2 (vreg n) = n. -Proof. - intros n. - unfold vreg. replace (n * 2) with (2 * n) by lia. - apply PeanoNat.Nat.div2_succ_double. -Qed. - -Lemma div2_vpar : - ∀ n, Nat.div2 (vpar n) = n. -Proof. - intros n. - unfold vpar. replace (n * 2) with (2 * n) by lia. - apply PeanoNat.Nat.div2_double. -Qed. - -Lemma odd_vreg : - ∀ n, Nat.odd (vreg n) = true. -Proof. - intros n. - unfold vreg. replace (n * 2) with (2 * n) by lia. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_mul. reflexivity. -Qed. - -Lemma odd_vpar : - ∀ n, Nat.odd (vpar n) = false. -Proof. - intros n. - unfold vpar. replace (n * 2) with (2 * n) by lia. - rewrite PeanoNat.Nat.odd_mul. reflexivity. -Qed. - -Lemma psubst_SS : - ∀ m Δ Γ σ n, - psubst (m :: Δ) (m :: Γ) (up_term σ) (S (S n)) = - plus 2 ⋅ psubst Δ Γ σ n. -Proof. - intros m Δ Γ σ n. - unfold psubst. rewrite div2_SS. cbn. - destruct nth_error eqn:e. 2: reflexivity. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_succ. - destruct_ifs. all: mode_eqs. - - ssimpl. erewrite erase_ren. - 2: eapply rscoping_S. - 2: eapply rscoping_comp_S. - ssimpl. rewrite <- pren_epm_lift. - ssimpl. eapply extRen_cterm. - intro x. unfold shift. change (pren S) with (pren (id >> S)). - rewrite pren_comp_S. ssimpl. rewrite pren_id. reflexivity. - - ssimpl. erewrite param_ren. - 2: eapply rscoping_S. - 2: eapply rscoping_comp_S. - ssimpl. eapply extRen_cterm. - intro x. unfold shift. change (pren S) with (pren (id >> S)). - rewrite pren_comp_S. ssimpl. rewrite pren_id. reflexivity. - - ssimpl. erewrite revive_ren. - 2: eapply rscoping_S. - 2: eapply rscoping_comp_S. - ssimpl. rewrite <- pren_rpm_lift. - eapply extRen_cterm. - intro x. unfold shift. change (pren S) with (pren (id >> S)). - rewrite pren_comp_S. ssimpl. rewrite pren_id. reflexivity. - - ssimpl. erewrite param_ren. - 2: eapply rscoping_S. - 2: eapply rscoping_comp_S. - ssimpl. eapply extRen_cterm. - intro x. unfold shift. change (pren S) with (pren (id >> S)). - rewrite pren_comp_S. ssimpl. rewrite pren_id. reflexivity. - - ssimpl. erewrite param_ren. - 2: eapply rscoping_S. - 2: eapply rscoping_comp_S. - ssimpl. eapply extRen_cterm. - intro x. unfold shift. change (pren S) with (pren (id >> S)). - rewrite pren_comp_S. ssimpl. rewrite pren_id. reflexivity. - - reflexivity. -Qed. - -Transparent epm_lift rpm_lift. - -Lemma psubst_epm_lift : - ∀ Γ Δ σ t, - ccscoping (erase_sc Δ) t cType → - (epm_lift t) <[ psubst Δ Γ σ ] = epm_lift (t <[ σ >> erase_term Γ ]). -Proof. - intros Γ Δ σ t ht. - unfold epm_lift. ssimpl. - eapply ext_cterm_scoped. 1: eassumption. - intros x hx. - ssimpl. unfold psubst. rewrite div2_vreg. - unfold inscope in hx. unfold erase_sc in hx. - rewrite nth_error_map in hx. - destruct (nth_error Δ x) eqn:e. 2: discriminate. - cbn in hx. destruct (relm m) eqn:e1. 2: discriminate. - rewrite odd_vreg. reflexivity. -Qed. - -Lemma psubst_rpm_lift : - ∀ Γ Δ σ t, - ccscoping (revive_sc Δ) t cType → - (rpm_lift t) <[ psubst Δ Γ σ ] = rpm_lift (t <[ rev_subst Δ Γ σ ]). -Proof. - intros Γ Δ σ t ht. - unfold rpm_lift. ssimpl. - eapply ext_cterm_scoped. 1: eassumption. - intros x hx. - ssimpl. unfold psubst. rewrite div2_vreg. - unfold rev_subst. unfold ghv. - unfold inscope in hx. unfold revive_sc in hx. - rewrite nth_error_map in hx. - destruct (nth_error Δ x) eqn:e. 2: discriminate. - cbn in hx. destruct (isProp m) eqn:e1. 1: discriminate. - rewrite odd_vreg. - destruct_ifs. all: mode_eqs. all: try discriminate. - all: try reflexivity. - destruct m. all: discriminate. -Qed. - -Opaque epm_lift rpm_lift. - -Lemma param_subst : - ∀ Γ Δ σ t, - sscoping Γ σ Δ → - sscoping_comp Γ σ Δ → - ⟦ Γ | t <[ σ ] ⟧p = ⟦ Δ | t ⟧p <[ psubst Δ Γ σ ]. -Proof. - intros Γ Δ σ t hσ hcσ. - induction t in Γ, Δ, σ, hσ, hcσ |- *. - - cbn. destruct (nth_error Δ n) eqn:e. - + destruct_if e1. - * mode_eqs. cbn. unfold psubst. rewrite div2_vreg. - rewrite e. cbn. rewrite odd_vreg. reflexivity. - * cbn. unfold psubst. rewrite div2_vpar. rewrite e. - rewrite odd_vpar. - destruct_ifs. all: try reflexivity. - destruct m. all: discriminate. - + eapply hcσ in e as e'. destruct e' as [k [e1 e2]]. - rewrite e1. cbn. rewrite e2. reflexivity. - - cbn. destruct_ifs. all: reflexivity. - - cbn. - unfold pmPi, pmPiP, pmPiNP, pPi. - erewrite IHt1. 2,3: eassumption. - erewrite IHt2. - 2:{ eapply sscoping_shift. eassumption. } - 2:{ eapply sscoping_comp_shift. assumption. } - erewrite !erase_subst. - 2-5: eauto using sscoping_shift, sscoping_comp_shift with cc_scope. - change (cEl (?t <[ ?σ ])) with ((cEl t) <[ σ ]). - erewrite <- psubst_epm_lift. 2:{ econstructor. eapply erase_scoping. } - destruct m, m0. all: cbn in *. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. all: cbn. all: f_equal. - 2:{ ssimpl. reflexivity. } - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - } - f_equal. all: f_equal. all: f_equal. - all: eapply ext_cterm. all: ssimpl. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. all: cbn. all: f_equal. - 2:{ ssimpl. reflexivity. } - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - } - f_equal. all: f_equal. all: f_equal. - all: eapply ext_cterm. all: ssimpl. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. all: cbn. all: f_equal. - 2:{ ssimpl. reflexivity. } - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - } - unfold cty_lift. f_equal. all: f_equal. - all: unfold close. all: ssimpl. - all: eapply ext_cterm. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - all: ssimpl. all: reflexivity. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. - rewrite rinstInst'_cterm. reflexivity. - } - cbn. unfold cty_lift. f_equal. f_equal. all: f_equal. - all: unfold close. all: ssimpl. - all: eapply ext_cterm. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - all: ssimpl. all: reflexivity. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - } - cbn. f_equal. f_equal. all: f_equal. all: f_equal. - all: ssimpl. - all: eapply ext_cterm. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - } - cbn. f_equal. f_equal. all: f_equal. all: f_equal. - all: ssimpl. - all: eapply ext_cterm. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - } - cbn. unfold cty_lift. f_equal. f_equal. all: f_equal. all: unfold close. - all: ssimpl. - all: eapply ext_cterm. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - all: ssimpl. all: reflexivity. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. - erewrite rinstInst'_cterm. reflexivity. - } - cbn. unfold cty_lift. f_equal. f_equal. all: f_equal. all: unfold close. - all: ssimpl. - all: eapply ext_cterm. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - all: ssimpl. all: reflexivity. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - } - cbn. f_equal. f_equal. all: f_equal. all: f_equal. - all: ssimpl. - all: eapply ext_cterm. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - } - cbn. f_equal. f_equal. all: f_equal. all: f_equal. - all: ssimpl. - all: eapply ext_cterm. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - } - cbn. f_equal. f_equal. all: f_equal. all: f_equal. - all: ssimpl. - all: eapply ext_cterm. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - all: ssimpl. - * rewrite rinstInst'_cterm. reflexivity. - * rewrite rinstInst'_cterm. reflexivity. - + f_equal. all: f_equal. - 2:{ ssimpl. reflexivity. } - 1: rewrite psubst_epm_lift. - 2:{ unshelve typeclasses eauto with cc_scope shelvedb. all: reflexivity. } - all: f_equal. - 2:{ - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. - erewrite rinstInst'_cterm. reflexivity. - } - cbn. unfold cty_lift. f_equal. f_equal. all: f_equal. all: unfold close. - all: ssimpl. - all: eapply ext_cterm. all: intros []. - all: cbn. 1,3: reflexivity. - all: ssimpl. - all: erewrite erase_ren ; eauto using rscoping_S, rscoping_comp_S. - all: ssimpl. all: reflexivity. - + f_equal. all: f_equal. 1: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - + f_equal. all: f_equal. 1: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - + f_equal. all: f_equal. 1: f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. reflexivity. - + f_equal. unfold close. - ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. - rewrite rinstInst'_cterm. reflexivity. - - cbn. - erewrite IHt1. 2,3: eassumption. - erewrite IHt2. - 2:{ eapply sscoping_shift. eassumption. } - 2:{ eapply sscoping_comp_shift. assumption. } - erewrite erase_subst. 2,3: eassumption. - change (cEl (?t <[ ?σ ])) with ((cEl t) <[ σ ]). - erewrite <- psubst_epm_lift. 2:{ econstructor. eapply erase_scoping. } - destruct_ifs. all: mode_eqs. - + cbn. f_equal. unfold close. ssimpl. - eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - ssimpl. rewrite psubst_SS. ssimpl. - erewrite rinstInst'_cterm. reflexivity. - + cbn. f_equal. unfold plam. f_equal. f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply ext_cterm. intros [| []]. all: cbn. 1,2: reflexivity. - rewrite psubst_SS. ssimpl. reflexivity. - + cbn. unfold plam. f_equal. f_equal. - * ssimpl. reflexivity. - * ssimpl. eapply ext_cterm. intros [| []]. all: cbn. - --- destruct_ifs. all: mode_eqs. all: try discriminate. - all: try reflexivity. - destruct m. all: discriminate. - --- destruct_ifs. all: mode_eqs. all: try discriminate. - all: try reflexivity. - destruct m. all: discriminate. - --- rewrite psubst_SS. ssimpl. reflexivity. - - cbn. - erewrite md_subst. 2,3: eassumption. - erewrite IHt1. 2,3: eassumption. - erewrite IHt2. 2,3: eassumption. - erewrite erase_subst. 2,3: eassumption. - erewrite revive_subst. 2,3: eassumption. - erewrite <- psubst_rpm_lift. 2: eapply revive_scoping. - erewrite <- psubst_epm_lift. 2: eapply erase_scoping. - destruct_ifs. all: reflexivity. - - cbn. - erewrite md_subst. 2,3: eassumption. - erewrite IHt. 2,3: eassumption. - destruct_ifs. all: reflexivity. - - cbn. - erewrite md_subst. 2,3: eassumption. - erewrite IHt. 2,3: eassumption. - destruct_ifs. all: reflexivity. - - cbn. - erewrite md_subst. 2,3: eassumption. - erewrite IHt1. 2,3: eassumption. - erewrite IHt3. 2,3: eassumption. - erewrite revive_subst. 2,3: eassumption. - erewrite <- psubst_rpm_lift. 2: eapply revive_scoping. - destruct_ifs. all: reflexivity. - - cbn. - erewrite md_subst. 2,3: eassumption. - erewrite IHt1. 2,3: eassumption. - erewrite IHt2. 2,3: eassumption. - erewrite revive_subst. 2,3: eassumption. - erewrite <- psubst_rpm_lift. 2: eapply revive_scoping. - destruct_ifs. all: reflexivity. - - cbn. eauto. - - cbn. eauto. - - cbn. - erewrite !revive_subst. 2-5: eassumption. - erewrite !erase_subst. 2,3: eassumption. - erewrite <- !psubst_rpm_lift. 2,3: eapply revive_scoping. - change (cEl (?t <[ ?σ ])) with ((cEl t) <[ σ ]). - erewrite <- psubst_epm_lift. 2:{ econstructor. eapply erase_scoping. } - reflexivity. - - cbn. - erewrite erase_subst. 2,3: eassumption. - erewrite revive_subst. 2,3: eassumption. - erewrite <- psubst_rpm_lift. 2: eapply revive_scoping. - change (cEl (?t <[ ?σ ])) with ((cEl t) <[ σ ]). - erewrite <- psubst_epm_lift. 2:{ econstructor. eapply erase_scoping. } - reflexivity. - - cbn. - erewrite md_subst. 2,3: eassumption. - erewrite IHt1. 2,3: eassumption. - erewrite IHt3. 2,3: eassumption. - erewrite IHt4. 2,3: eassumption. - erewrite IHt5. 2,3: eassumption. - erewrite IHt6. 2,3: eassumption. - erewrite !erase_subst. 2-5: eassumption. - erewrite !revive_subst. 2-7: eassumption. - erewrite <- !psubst_rpm_lift. 2-4: eapply revive_scoping. - change (cEl (?t <[ ?σ ])) with ((cEl t) <[ σ ]). - erewrite <- !psubst_epm_lift. - 2: eapply erase_scoping. - 2:{ econstructor. eapply erase_scoping. } - destruct md eqn:e. - + reflexivity. - + unfold pcastTG. cbn. ssimpl. reflexivity. - + unfold pcastTG. cbn. ssimpl. reflexivity. - + unfold pcastP. cbn. ssimpl. reflexivity. - - cbn. reflexivity. - - cbn. reflexivity. - - cbn. reflexivity. - - cbn. - erewrite IHt1, IHt2, IHt3, IHt4. 2-9: eassumption. - erewrite !erase_subst. 2-9: eassumption. - erewrite !revive_subst. 2-5: eassumption. - erewrite <- !psubst_rpm_lift. 2-3: eapply revive_scoping. - erewrite <- !psubst_epm_lift. 2-5: eapply erase_scoping. - destruct m. - + unfold pmifK. unfold perif. cbn. f_equal. f_equal. all: f_equal. - 1,4: f_equal. 1,2: f_equal. 1-4: f_equal. 1,2,5-7,10: f_equal. - 2,3,5,6: f_equal. 2,4: f_equal. - all: ssimpl. all: reflexivity. - + unfold pmif, plam. cbn. f_equal. f_equal. f_equal. - ssimpl. reflexivity. - + unfold pmif, plam. cbn. f_equal. f_equal. f_equal. - ssimpl. reflexivity. - + reflexivity. - - cbn. reflexivity. - - cbn. reflexivity. - - cbn. f_equal. eauto. - - cbn. - erewrite IHt1, IHt2, IHt3, IHt4. 2-9: eassumption. - erewrite !erase_subst. 2-9: eassumption. - erewrite !revive_subst. 2-5: eassumption. - erewrite <- !psubst_rpm_lift. 2-3: eapply revive_scoping. - erewrite <- !psubst_epm_lift. 2-5: eapply erase_scoping. - destruct m. all: reflexivity. - - cbn. erewrite IHt1, IHt2. 2-5: eassumption. - erewrite !erase_subst, !revive_subst. 2-5: eassumption. - erewrite <- !psubst_epm_lift, <- !psubst_rpm_lift. - 2: eapply revive_scoping. - 2: eapply erase_scoping. - reflexivity. - - cbn. erewrite IHt. 2-3: eassumption. - reflexivity. - - cbn. erewrite IHt1, IHt2, IHt3. 2-7: eassumption. - reflexivity. - - cbn. - erewrite IHt1, IHt2, IHt3, IHt4, IHt5, IHt6. 2-13: eassumption. - erewrite !erase_subst. 2-11: eassumption. - erewrite !revive_subst. 2-7: eassumption. - erewrite <- !psubst_rpm_lift. 2-4: eapply revive_scoping. - erewrite <- !psubst_epm_lift. 2-6: eapply erase_scoping. - destruct m. all: reflexivity. - - cbn. reflexivity. - - cbn. - erewrite IHt1. 2,3: eassumption. - erewrite IHt2. 2,3: eassumption. - erewrite erase_subst. 2,3: eassumption. - destruct_ifs. all: try reflexivity. - + cbn. f_equal. f_equal. - rewrite psubst_epm_lift. 2: eauto with cc_scope. - reflexivity. - + cbn. f_equal. f_equal. - rewrite psubst_epm_lift. 2: eauto with cc_scope. - reflexivity. -Qed. - -(** Parametricity preserves conversion **) - -Lemma vreg_vpar_dec : - ∀ n, { n = vpar (Nat.div2 n) } + { n = vreg (Nat.div2 n) }. -Proof. - intros n. - destruct (PeanoNat.Nat.Even_Odd_dec n). - - left. unfold vpar. - etransitivity. - + apply PeanoNat.Nat.Even_double. assumption. - + unfold Nat.double. lia. - - right. unfold vreg. - etransitivity. - + apply PeanoNat.Nat.Odd_double. assumption. - + unfold Nat.double. lia. -Qed. - -Lemma ccong_pPi : - ∀ Γ mx A B C A' B' C', - Γ ⊢ᶜ A ≡ A' → - Γ ⊢ᶜ B ≡ B' → - Some (mx, capp (S ⋅ B) (cvar 0)) :: Some (cType, A) :: Γ ⊢ᶜ C ≡ C' → - Γ ⊢ᶜ pPi mx A B C ≡ pPi mx A' B' C'. -Proof. - intros Γ mx A B C A' B' C' hA hB hC. - unfold pPi. - econv. - eapply cconv_ren. 2: eassumption. - apply crtyping_S. -Qed. - -Hint Resolve ccong_pPi : cc_conv. - -Lemma ccong_plam : - ∀ Γ mp A B t A' B' t', - Γ ⊢ᶜ A ≡ A' → - Γ ⊢ᶜ B ≡ B' → - Some (mp, capp (S ⋅ B) (cvar 0)) :: Some (cType, A) :: Γ ⊢ᶜ t ≡ t' → - Γ ⊢ᶜ plam mp A B t ≡ plam mp A' B' t'. -Proof. - intros Γ mp A B t A' B' t' hA hB ht. - econv. - eapply cconv_ren. 2: eassumption. - apply crtyping_S. -Qed. - -Hint Resolve ccong_plam : cc_conv. - -Transparent epm_lift rpm_lift. - -Lemma ccong_epm_lift : - ∀ Γ Γ' t t', - ⟦ Γ ⟧ε ⊢ᶜ t ≡ t' → - Γ' = ⟦ Γ ⟧p → - Γ' ⊢ᶜ epm_lift t ≡ epm_lift t'. -Proof. - intros Γ Γ' t t' ht ->. - unfold epm_lift. eapply cconv_ren. - - apply typing_er_sub_param. - - assumption. -Qed. - -Lemma ccong_rpm_lift : - ∀ Γ Γ' t t', - ⟦ Γ ⟧v ⊢ᶜ t ≡ t' → - Γ' = ⟦ Γ ⟧p → - Γ' ⊢ᶜ rpm_lift t ≡ rpm_lift t'. -Proof. - intros Γ Γ' t t' ht ->. - unfold rpm_lift. eapply cconv_ren. - - apply typing_rev_sub_param. - - assumption. -Qed. - -(* Need to have them opaque so that they can't be confused. *) -Hint Opaque epm_lift rpm_lift : cc_scope cc_conv cc_type. -Hint Resolve ccong_epm_lift ccong_rpm_lift : cc_conv. - Opaque epm_lift rpm_lift. - -Hint Resolve cconv_ren cconv_subst : cc_conv. - -Lemma erase_conv_eq : - ∀ Γ Γ' Γ'' u v, - Γ ⊢ u ≡ v → - Γ' = ⟦ Γ ⟧ε → - Γ'' = sc Γ → - Γ' ⊢ᶜ ⟦ Γ'' | u ⟧ε ≡ ⟦ Γ'' | v ⟧ε. -Proof. - intros Γ Γ' Γ'' u v h -> ->. - apply erase_conv. assumption. -Qed. - -Hint Resolve erase_conv_eq : cc_conv. - -Lemma revive_conv_eq : - ∀ Γ Γ' Γ'' u v, - Γ ⊢ u ≡ v → - Γ' = ⟦ Γ ⟧v → - Γ'' = sc Γ → - Γ' ⊢ᶜ ⟦ Γ'' | u ⟧v ≡ ⟦ Γ'' | v ⟧v. -Proof. - intros Γ Γ' Γ'' u v h -> ->. - apply revive_conv. assumption. -Qed. - -Hint Resolve revive_conv_eq : cc_conv. - -Lemma crtyping_shift_eq : - ∀ Γ Δ Ξ mx A ρ, - crtyping Γ ρ Δ → - Ξ = Some (mx, ρ ⋅ A) :: Γ → - crtyping Ξ (up_ren ρ) (Some (mx, A) :: Δ). -Proof. - intros Γ Δ Ξ mx A ρ hρ ->. - apply crtyping_shift. assumption. -Qed. - -Lemma csc_param_ctx : - ∀ Γ, - csc ⟦ Γ ⟧p = param_sc (sc Γ). -Proof. - intros Γ. - induction Γ as [| [m A] Γ ih]. - - cbn. reflexivity. - - cbn. destruct_ifs. all: cbn. - + f_equal. f_equal. apply ih. - + f_equal. f_equal. apply ih. - + f_equal. f_equal. apply ih. -Qed. - -Lemma ccong_perif : - ∀ Γ be be' Pe Pe' te te' fe fe', - Γ ⊢ᶜ be ≡ be' → - Γ ⊢ᶜ Pe ≡ Pe' → - Γ ⊢ᶜ te ≡ te' → - Γ ⊢ᶜ fe ≡ fe' → - Γ ⊢ᶜ perif be Pe te fe ≡ perif be' Pe' te' fe'. -Proof. - intros Γ be be' Pe Pe' te te' fe fe' hb hP ht hf. - unfold perif. econv. apply crtyping_S. -Qed. - -Hint Opaque perif : cc_conv. -Hint Resolve ccong_perif : cc_conv. - -Lemma ccong_pmif : - ∀ Γ bP Pe PP te tP fe fP bP' Pe' PP' te' tP' fe' fP', - Γ ⊢ᶜ bP ≡ bP' → - Γ ⊢ᶜ Pe ≡ Pe' → - Γ ⊢ᶜ PP ≡ PP' → - Γ ⊢ᶜ te ≡ te' → - Γ ⊢ᶜ tP ≡ tP' → - Γ ⊢ᶜ fe ≡ fe' → - Γ ⊢ᶜ fP ≡ fP' → - Γ ⊢ᶜ pmif bP Pe PP te tP fe fP ≡ pmif bP' Pe' PP' te' tP' fe' fP'. -Proof. - intros Γ bP Pe PP te tP fe fP bP' Pe' PP' te' tP' fe' fP'. - intros hbP hPe hPP hte htP hfe hfP. - unfold pmif. econv. all: apply crtyping_S. -Qed. - -Hint Opaque pmif : cc_conv. -Hint Resolve ccong_pmif : cc_conv. - -Lemma ccong_pmifK : - ∀ Γ be bP Pe PP te tP fe fP be' bP' Pe' PP' te' tP' fe' fP', - Γ ⊢ᶜ be ≡ be' → - Γ ⊢ᶜ bP ≡ bP' → - Γ ⊢ᶜ Pe ≡ Pe' → - Γ ⊢ᶜ PP ≡ PP' → - Γ ⊢ᶜ te ≡ te' → - Γ ⊢ᶜ tP ≡ tP' → - Γ ⊢ᶜ fe ≡ fe' → - Γ ⊢ᶜ fP ≡ fP' → - Γ ⊢ᶜ pmifK be bP Pe PP te tP fe fP ≡ pmifK be' bP' Pe' PP' te' tP' fe' fP'. -Proof. - intros Γ be bP Pe PP te tP fe fP be' bP' Pe' PP' te' tP' fe' fP'. - intros hbe hbP hPe hPP hte htP hfe hfP. - unfold pmifK. econv. all: apply crtyping_S. -Qed. - -Hint Opaque pmifK : cc_conv. -Hint Resolve ccong_pmifK : cc_conv. - -Lemma pren_S : - ∀ n, pren S n = S (S n). -Proof. - intro n. - change (pren S n) with (pren (id >> S) n). - rewrite pren_comp_S. rewrite pren_id. reflexivity. -Qed. - -Lemma pren_S_pw : - pointwise_relation _ eq (pren S) (S >> S). -Proof. - intro n. apply pren_S. -Qed. - -Lemma param_conv : - ∀ Γ u v, - Γ ⊢ u ≡ v → - ⟦ Γ ⟧p ⊢ᶜ ⟦ sc Γ | u ⟧p ≡ ⟦ sc Γ | v ⟧p. -Proof. - intros Γ u v h. - induction h. - (* all: try solve [ cbn ; destruct_ifs ; eauto with cc_conv ]. *) - - cbn. - erewrite scoping_md. 2: eassumption. - destruct_ifs. all: mode_eqs. all: try discriminate. - + eapply cconv_trans. - 1:{ constructor. 2: apply cconv_refl. constructor. } - cbn. - eapply cconv_trans. 1: constructor. - ssimpl. apply ccmeta_refl. - erewrite param_subst. - 2:{ eapply sscoping_one. eassumption. } - 2: eapply sscoping_comp_one. - ssimpl. eapply ext_cterm_scoped. - 1:{ eapply param_scoping. eassumption. } - (* The following we do basically four times, but we don't know how - to factorise. - *) - intros [| []] hx. all: cbn. 1,2: reflexivity. - unfold inscope in hx. cbn in hx. - unfold psubst. rewrite div2_SS. cbn. - destruct (vreg_vpar_dec n) as [en | en]. - * rewrite en in hx. rewrite nth_error_param_vpar in hx. - destruct nth_error as [mx|] eqn:e1. 2: discriminate. - cbn in hx. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_succ. - destruct PeanoNat.Nat.odd eqn:eodd. - 1:{ rewrite en in eodd. rewrite odd_vpar in eodd. discriminate. } - destruct (isProp mx) eqn:e2. 1: discriminate. - destruct (isKind mx) eqn:e3. all: mode_eqs. - -- cbn. f_equal. assumption. - -- destruct mx. all: try discriminate. - ++ cbn. f_equal. assumption. - ++ cbn. f_equal. assumption. - * set (p := Nat.div2 n) in *. - rewrite en in hx. rewrite nth_error_param_vreg in hx. - destruct nth_error as [mx|] eqn:e1. 2: discriminate. - cbn in hx. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_succ. - destruct PeanoNat.Nat.odd eqn:eodd. - 2:{ rewrite en in eodd. rewrite odd_vreg in eodd. discriminate. } - destruct (isProp mx) eqn:e2. - -- mode_eqs. cbn. f_equal. assumption. - -- unfold relv, ghv. rewrite e1. - destruct_ifs. - ++ rewrite en. reflexivity. - ++ rewrite en. reflexivity. - ++ destruct mx. all: discriminate. - + destruct (isType mx) eqn:e2. 2:{ destruct mx. all: discriminate. } - mode_eqs. - eapply cconv_trans. - 1:{ constructor. 2: apply cconv_refl. constructor. } - cbn. - eapply cconv_trans. 1: constructor. - ssimpl. apply ccmeta_refl. - erewrite param_subst. - 2:{ eapply sscoping_one. eassumption. } - 2: eapply sscoping_comp_one. - ssimpl. eapply ext_cterm_scoped. - 1:{ eapply param_scoping. eassumption. } - (* Basically same as above, is there a nice lemma to state? *) - intros [| []] hx. all: cbn. 1,2: reflexivity. - unfold inscope in hx. cbn in hx. - unfold psubst. rewrite div2_SS. cbn. - destruct (vreg_vpar_dec n) as [en | en]. - * set (p := Nat.div2 n) in *. - rewrite en in hx. rewrite nth_error_param_vpar in hx. - destruct nth_error as [mx|] eqn:emx. 2: discriminate. - cbn in hx. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_succ. - destruct PeanoNat.Nat.odd eqn:eodd. - 1:{ rewrite en in eodd. rewrite odd_vpar in eodd. discriminate. } - destruct (isProp mx) eqn:e2. 1: discriminate. - destruct (isKind mx) eqn:e3. all: mode_eqs. - -- cbn. f_equal. assumption. - -- destruct mx. all: try discriminate. - ++ cbn. f_equal. assumption. - ++ cbn. f_equal. assumption. - * set (p := Nat.div2 n) in *. - rewrite en in hx. rewrite nth_error_param_vreg in hx. - destruct nth_error as [mx|] eqn:emx. 2: discriminate. - cbn in hx. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_succ. - destruct PeanoNat.Nat.odd eqn:eodd. - 2:{ rewrite en in eodd. rewrite odd_vreg in eodd. discriminate. } - destruct (isProp mx) eqn:e2. - -- mode_eqs. cbn. f_equal. assumption. - -- unfold relv, ghv. rewrite emx. - destruct_ifs. - ++ rewrite en. reflexivity. - ++ rewrite en. reflexivity. - ++ destruct mx. all: discriminate. - + eapply cconv_trans. - 1:{ constructor. 2: apply cconv_refl. constructor. } - cbn. - eapply cconv_trans. 1: constructor. - ssimpl. apply ccmeta_refl. - erewrite param_subst. - 2:{ eapply sscoping_one. eassumption. } - 2: eapply sscoping_comp_one. - ssimpl. eapply ext_cterm_scoped. - 1:{ eapply param_scoping. eassumption. } - (* Basically same as above, is there a nice lemma to state? *) - intros [| []] hx. all: cbn. 1,2: reflexivity. - unfold inscope in hx. cbn in hx. - unfold psubst. rewrite div2_SS. cbn. - destruct (vreg_vpar_dec n) as [en | en]. - * set (p := Nat.div2 n) in *. - rewrite en in hx. rewrite nth_error_param_vpar in hx. - destruct nth_error as [mx|] eqn:emx. 2: discriminate. - cbn in hx. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_succ. - destruct PeanoNat.Nat.odd eqn:eodd. - 1:{ rewrite en in eodd. rewrite odd_vpar in eodd. discriminate. } - destruct (isProp mx) eqn:e3. 1: discriminate. - destruct (isKind mx) eqn:e4. all: mode_eqs. - -- cbn. f_equal. assumption. - -- destruct mx. all: try discriminate. - ++ cbn. f_equal. assumption. - ++ cbn. f_equal. assumption. - * set (p := Nat.div2 n) in *. - rewrite en in hx. rewrite nth_error_param_vreg in hx. - destruct nth_error as [mx|] eqn:emx. 2: discriminate. - cbn in hx. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_succ. - destruct PeanoNat.Nat.odd eqn:eodd. - 2:{ rewrite en in eodd. rewrite odd_vreg in eodd. discriminate. } - destruct (isProp mx) eqn:e3. - -- mode_eqs. cbn. f_equal. assumption. - -- unfold relv, ghv. rewrite emx. - destruct_ifs. - ++ rewrite en. reflexivity. - ++ rewrite en. reflexivity. - ++ destruct mx. all: discriminate. - + eapply cconv_trans. 1: constructor. - unfold close. ssimpl. apply ccmeta_refl. - erewrite param_subst. - 2:{ eapply sscoping_one. eassumption. } - 2: eapply sscoping_comp_one. - ssimpl. eapply ext_cterm_scoped. - 1:{ eapply param_scoping. eassumption. } - (* Basically same as above, is there a nice lemma to state? *) - intros [| []] hx. all: cbn. 1,2: reflexivity. - unfold inscope in hx. cbn in hx. - unfold psubst. rewrite div2_SS. cbn. - destruct (vreg_vpar_dec n) as [en | en]. - * set (p := Nat.div2 n) in *. - rewrite en in hx. rewrite nth_error_param_vpar in hx. - destruct nth_error as [mx|] eqn:emx. 2: discriminate. - cbn in hx. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_succ. - destruct PeanoNat.Nat.odd eqn:eodd. - 1:{ rewrite en in eodd. rewrite odd_vpar in eodd. discriminate. } - destruct (isProp mx) eqn:e3. 1: discriminate. - destruct (isKind mx) eqn:e4. all: mode_eqs. - -- cbn. f_equal. assumption. - -- destruct mx. all: try discriminate. - ++ cbn. f_equal. assumption. - ++ cbn. f_equal. assumption. - * set (p := Nat.div2 n) in *. - rewrite en in hx. rewrite nth_error_param_vreg in hx. - destruct nth_error as [mx|] eqn:emx. 2: discriminate. - cbn in hx. - rewrite PeanoNat.Nat.odd_succ. - rewrite PeanoNat.Nat.even_succ. - destruct PeanoNat.Nat.odd eqn:eodd. - 2:{ rewrite en in eodd. rewrite odd_vreg in eodd. discriminate. } - destruct (isProp mx) eqn:e3. - -- mode_eqs. cbn. f_equal. assumption. - -- unfold relv, ghv. rewrite emx. - destruct_ifs. - ++ rewrite en. reflexivity. - ++ rewrite en. reflexivity. - ++ destruct mx. all: discriminate. - + destruct mx. all: discriminate. - - cbn. - erewrite scoping_md. 2: eassumption. - erewrite scoping_md. 2: eassumption. - destruct_if e. - 1:{ - destruct H2 as [emp | [ emp | ]]. 3: contradiction. - all: subst. all: discriminate. - } - cbn. apply cconv_refl. - - cbn. destruct m. - + unfold pmifK. change (epm_lift etrue) with etrue. - eapply cconv_trans. - 1:{ - econstructor. 2: econv. - constructor. - } - eapply cconv_trans. 1: constructor. - ssimpl. econv. - + unfold pmif. constructor. - + unfold pmif. constructor. - + constructor. - - cbn. destruct m. - + unfold pmifK. change (epm_lift etrue) with etrue. - eapply cconv_trans. - 1:{ - econstructor. 2: econv. - constructor. - } - eapply cconv_trans. 1: constructor. - ssimpl. econv. - + unfold pmif. constructor. - + unfold pmif. constructor. - + constructor. - - cbn. eapply param_scoping in H0, H1, H2. - rewrite <- csc_param_ctx in H0, H1, H2. - destruct m. - + contradiction. - + cbn in *. - eapply cconv_irr. - * escope. all: eauto using csc_param_ctx. - eapply scoping_epm_lift. 1: escope. apply csc_param_ctx. - * assumption. - + cbn in *. - eapply cconv_irr. - * escope. all: eauto using csc_param_ctx. - eapply scoping_epm_lift. 1: escope. apply csc_param_ctx. - * assumption. - + cbn in *. - eapply cconv_irr. - * escope. 2: eauto using csc_param_ctx. - eapply scoping_epm_lift. 1: escope. apply csc_param_ctx. - * assumption. - - cbn. - remd. - eapply param_scoping in H0, H1, H2, H3. - rewrite <- csc_param_ctx in H0, H1, H2, H3. - destruct m. - + contradiction. - + cbn in *. - eapply cconv_irr. - * escope. all: eauto using csc_param_ctx. - eapply scoping_epm_lift. 1: escope. 1: reflexivity. - apply csc_param_ctx. - * escope. all: eauto using csc_param_ctx. - eapply scoping_epm_lift. 1: escope. - all: try reflexivity. - apply csc_param_ctx. - + cbn in *. - eapply cconv_irr. - * escope. all: eauto using csc_param_ctx. - eapply scoping_epm_lift. 1: escope. 1: reflexivity. - apply csc_param_ctx. - * { - escope. all: eauto using csc_param_ctx. - eapply scoping_rpm_lift. 2: apply csc_param_ctx. - econstructor. - 1,2: eapply scoping_to_rev. - all: escope. all: reflexivity. - } - + cbn in *. - eapply cconv_irr. - * escope. 2: eauto using csc_param_ctx. - eapply scoping_epm_lift. 1: escope. 1: reflexivity. - apply csc_param_ctx. - * escope. all: eauto using csc_param_ctx. - - cbn. eapply param_scoping in H0, H1, H2, H3, H4. - rewrite <- csc_param_ctx in H0, H1, H2, H3, H4. - destruct m. - + contradiction. - + cbn in *. eapply cconv_irr. - * { - escope. all: eauto using csc_param_ctx. - - change (rpm_lift ?t) with t. escope. - - change (epm_lift ?t) with (vreg ⋅ t). cbn. - change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - escope. apply csc_param_ctx. - } - * auto. - + cbn in *. eapply cconv_irr. - * { - escope. all: eauto using csc_param_ctx. - - change (rpm_lift ?t) with t. escope. - - change (epm_lift ?t) with (vreg ⋅ t). cbn. - change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - escope. apply csc_param_ctx. - } - * auto. - + cbn in *. eapply cconv_irr. - * { - escope. all: eauto using csc_param_ctx. - - change (rpm_lift ?t) with t. escope. - - change (epm_lift ?t) with (vreg ⋅ t). cbn. - change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - escope. apply csc_param_ctx. - } - * auto. - - simpl. remd. - eapply param_scoping in H0, H1, H2, H3, H4, H5, H6, H7. - rewrite <- csc_param_ctx in H0, H1, H2, H3, H4, H5, H6, H7. - destruct m. - + contradiction. - + cbn in *. eapply cconv_irr. - * { - escope. all: eauto using csc_param_ctx. - change (epm_lift ?t) with (vreg ⋅ t). cbn. - change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - escope. all: apply csc_param_ctx. - } - * change (epm_lift ?t) with (vreg ⋅ t). cbn. - change (rpm_lift ?t) with (vreg ⋅ t). cbn. - erewrite !erase_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. - erewrite !param_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. - change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - change (vreg ⋅ ⟦ ?G | ?t ⟧v) with (⟦ G | t ⟧pv). - ssimpl. rewrite pren_S_pw. ssimpl. - rewrite <- !rinstInst'_cterm. - change (S >> vreg) with (vreg >> S >> S). - rewrite <- !funcomp_assoc. - change (S >> vreg) with (vreg >> S >> S). - rewrite !funcomp_assoc. - rewrite <- !renRen_cterm. - change (ren_cterm vreg ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - escope. all: cbn. all: eauto using csc_param_ctx. - + cbn in *. eapply cconv_irr. - * { - escope. all: eauto using csc_param_ctx. - change (epm_lift ?t) with (vreg ⋅ t). cbn. - change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - escope. all: apply csc_param_ctx. - } - * change (epm_lift ?t) with (vreg ⋅ t). cbn. - change (rpm_lift ?t) with (vreg ⋅ t). cbn. - erewrite !erase_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. - erewrite !param_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. - change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - change (vreg ⋅ ⟦ ?G | ?t ⟧v) with (⟦ G | t ⟧pv). - ssimpl. rewrite pren_S_pw. ssimpl. - rewrite <- !rinstInst'_cterm. - change (S >> vreg) with (vreg >> S >> S). - rewrite <- !funcomp_assoc. - change (S >> vreg) with (vreg >> S >> S). - rewrite !funcomp_assoc. - rewrite <- !renRen_cterm. - change (ren_cterm vreg ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - change (ren_cterm vreg ⟦ ?G | ?t ⟧v) with (⟦ G | t ⟧pv). - escope. all: cbn. all: eauto using csc_param_ctx. - + cbn in *. eapply cconv_irr. - * { - escope. all: eauto using csc_param_ctx. - change (epm_lift ?t) with (vreg ⋅ t). cbn. - change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - escope. all: apply csc_param_ctx. - } - * change (epm_lift ?t) with (vreg ⋅ t). cbn. - change (rpm_lift ?t) with (vreg ⋅ t). cbn. - erewrite !erase_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. - erewrite !param_ren. 2-7: eauto using rscoping_S, rscoping_comp_S. - change (vreg ⋅ ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - change (vreg ⋅ ⟦ ?G | ?t ⟧v) with (⟦ G | t ⟧pv). - ssimpl. rewrite pren_S_pw. ssimpl. - rewrite <- !rinstInst'_cterm. - change (S >> vreg) with (vreg >> S >> S). - rewrite <- !funcomp_assoc. - change (S >> vreg) with (vreg >> S >> S). - rewrite !funcomp_assoc. - rewrite <- !renRen_cterm. - change (ren_cterm vreg ⟦ ?G | ?t ⟧ε) with (⟦ G | t ⟧pε). - change (ren_cterm vreg ⟦ ?G | ?t ⟧v) with (⟦ G | t ⟧pv). - escope. all: cbn. all: eauto using csc_param_ctx. - - cbn. apply cconv_refl. - - cbn. - destruct m, mx. all: simpl. - + econv. all: try reflexivity. - all: eauto using crtyping_S. - cbn. eapply crtyping_shift_eq. - * apply crtyping_shift. apply crtyping_S. - * cbn. f_equal. ssimpl. reflexivity. - + econv. all: try reflexivity. - all: eauto using crtyping_S. - cbn. eapply crtyping_shift_eq. - * apply crtyping_shift. apply crtyping_S. - * cbn. f_equal. ssimpl. reflexivity. - + econv. all: try reflexivity. - all: eauto using crtyping_S, cstyping_one_none. - cbn. eapply crtyping_shift_eq. - * apply crtyping_shift. apply crtyping_S. - * cbn. f_equal. ssimpl. reflexivity. - + econv. all: try reflexivity. - all: eauto using crtyping_S, cstyping_one_none. - * cbn. apply crtyping_shift. apply crtyping_S. - * eapply cstyping_one_none. - + econv. all: try reflexivity. - all: eauto using crtyping_S. - cbn. eapply crtyping_shift_eq. - * apply crtyping_shift. apply crtyping_S. - * cbn. f_equal. ssimpl. reflexivity. - + econv. all: try reflexivity. - all: eauto using crtyping_S. - cbn. eapply crtyping_shift_eq. - * apply crtyping_shift. apply crtyping_S. - * cbn. f_equal. ssimpl. reflexivity. - + econv. all: try reflexivity. - all: eauto using crtyping_S, cstyping_one_none. - cbn. eapply crtyping_shift_eq. - * apply crtyping_shift. apply crtyping_S. - * cbn. f_equal. ssimpl. reflexivity. - + econv. all: try reflexivity. - all: eauto using crtyping_S, cstyping_one_none. - * apply crtyping_shift. apply crtyping_S. - * apply cstyping_one_none. - + econv. all: try reflexivity. - all: eauto using crtyping_S. - cbn. eapply crtyping_shift_eq. - * apply crtyping_shift. apply crtyping_S. - * cbn. f_equal. ssimpl. reflexivity. - + econv. all: try reflexivity. - all: eauto using crtyping_S. - cbn. eapply crtyping_shift_eq. - * apply crtyping_shift. apply crtyping_S. - * cbn. f_equal. ssimpl. reflexivity. - + econv. all: try reflexivity. - all: eauto using crtyping_S, cstyping_one_none. - * eapply cstyping_nones. - * cbn. eapply crtyping_shift_eq. - -- apply crtyping_shift. apply crtyping_S. - -- cbn. f_equal. ssimpl. reflexivity. - + econv. all: try reflexivity. - all: eauto using crtyping_S, cstyping_one_none. - * apply crtyping_shift. apply crtyping_S. - * apply cstyping_one_none. - + econv. all: try reflexivity. - eauto using crtyping_S. - + econv. all: try reflexivity. - eauto using crtyping_S. - + econv. all: try reflexivity. - eauto using crtyping_S. - + econv. apply cstyping_one_none. - - cbn in *. destruct_ifs. - + econv. apply cstyping_one_none. - + econv. all: try reflexivity. - apply crtyping_S. - + econv. all: try reflexivity. - apply crtyping_S. - - cbn. - eapply conv_md in h2 as e2. rewrite <- e2. - destruct_ifs. - + econv. all: reflexivity. - + econv. all: reflexivity. - + econv. - - cbn. - eapply conv_md in h as e. rewrite <- e. - destruct_ifs. all: econv. - - cbn. - eapply conv_md in h as e. rewrite <- e. - destruct_ifs. all: econv. - - cbn. - eapply conv_md in h3 as e3. rewrite <- e3. - destruct_ifs. all: econv. all: reflexivity. - - cbn. - eapply conv_md in h2 as e2. rewrite <- e2. - destruct_ifs. all: econv. all: reflexivity. - - cbn. - econv. all: reflexivity. - - cbn. destruct m. - + econv. all: reflexivity. - + econv. all: reflexivity. - + econv. all: reflexivity. - + econv. - - cbn. econv. - - cbn. destruct m. - all: econv. all: reflexivity. - - cbn. econv. all: reflexivity. - - cbn. econv. - - cbn. econv. - - cbn. destruct m. - + econv. - + econv. all: reflexivity. - + econv. all: reflexivity. - + econv. all: reflexivity. - - cbn. - destruct_ifs. all: econv. all: reflexivity. - - econv. - - apply cconv_sym. assumption. - - eapply cconv_trans. all: eauto. - - eapply param_scoping in H, H0. cbn in *. - apply cconv_irr. all: rewrite csc_param_ctx. all: assumption. -Qed. - -(** Parametricity ignores casts **) - -Lemma ccong_pmPi : - ∀ Γ mx m Te Ae Ap Bp Te' Ae' Ap' Bp', - Γ ⊢ᶜ Te ≡ Te' → - Γ ⊢ᶜ Ae ≡ Ae' → - Γ ⊢ᶜ Ap ≡ Ap' → - let Γ' := - if isProp mx then - None :: Some (cProp, Ap) :: Γ - else if isKind mx then - Some (cType, capp (S ⋅ Ap) (cvar 0)) :: Some (cType, Ae) :: Γ - else - Some (cProp, capp (S ⋅ Ap) (cvar 0)) :: Some (cType, Ae) :: Γ - in - Γ' ⊢ᶜ Bp ≡ Bp' → - Γ ⊢ᶜ pmPi mx m Te Ae Ap Bp ≡ pmPi mx m Te' Ae' Ap' Bp'. -Proof. - cbn zeta. - intros Γ mx m Te Ae Ap Bp Te' Ae' Ap' Bp' hTe hAe hAp hBp. - unfold pmPi. destruct (isProp m) eqn:ep. - - unfold pmPiP. destruct_ifs. all: econv. - all: try apply crtyping_S. - apply cstyping_one_none. - - unfold pmPiNP. econv. - destruct (isProp mx) eqn:epx. all: econv. - all: try apply crtyping_S. - + eapply crtyping_shift. apply crtyping_S. - + eapply cstyping_one_none. - + destruct (isKind mx) eqn:ekx. - * { - eapply crtyping_shift_eq. - - eapply crtyping_shift. apply crtyping_S. - - f_equal. f_equal. f_equal. cbn. ssimpl. reflexivity. - } - * { - eapply crtyping_shift_eq. - - eapply crtyping_shift. apply crtyping_S. - - f_equal. f_equal. f_equal. cbn. ssimpl. reflexivity. - } -Qed. - -Hint Opaque pmPi : cc_conv. -Hint Resolve ccong_pmPi : cc_conv. - -Lemma meta_ctx_conv_conv : - ∀ Γ Δ u v, - Γ ⊢ᶜ u ≡ v → - Γ = Δ → - Δ ⊢ᶜ u ≡ v. -Proof. - intros Γ ? u v h ->. - assumption. -Qed. - -Lemma meta_ctx_param_conv : - ∀ sΓ Γp sΔ Δp u v, - Δp ⊢ᶜ ⟦ sΔ | u ⟧p ≡ ⟦ sΔ | v ⟧p → - Γp = Δp → - sΓ = sΔ → - Γp ⊢ᶜ ⟦ sΓ | u ⟧p ≡ ⟦ sΓ | v ⟧p. -Proof. - intros sΓ Γp sΔ Δp u v h -> ->. - assumption. -Qed. - -Hint Opaque plam : cc_conv. - -Lemma meta_ccscoping_conv : - ∀ Γ t m m', - ccscoping Γ t m → - m = m' → - ccscoping Γ t m'. -Proof. - intros Γ t m m' h ->. - assumption. -Qed. - -Lemma param_castrm : - ∀ Γ t m, - cscoping Γ t m → - ⟦ Γ ⟧p ⊢ᶜ ⟦ sc Γ | (ε|t|) ⟧p ≡ ⟦ sc Γ | t ⟧p. -Proof. - intros Δ t m h. - remember (sc Δ) as Γ eqn:e in *. - induction h in Δ, e |- *. - all: try solve [ econv ]. - all: try solve [ - cbn ; rewrite <- ?md_castrm ; - rewrite ?erase_castrm, ?revive_castrm ; - destruct_ifs ; econv - ]. - - cbn. rewrite !erase_castrm. econv. - subst. - eapply meta_ctx_conv_conv. - + eapply IHh2 with (Δ := Δ ,, (mx, ε|A|)). - reflexivity. - + cbn. rewrite !erase_castrm. reflexivity. - - cbn. rewrite !erase_castrm. - destruct_ifs. - + econstructor. 1: eauto. - eapply cconv_close. - eapply meta_ctx_conv_conv. - * eapply IHh2 with (Δ := Δ ,, (mx, ε|A|)). - subst. reflexivity. - * cbn. subst. rewrite !erase_castrm. rewrite e0. reflexivity. - + econv. - eapply meta_ctx_conv_conv. - * eapply IHh2 with (Δ := Δ ,, (mx, ε|A|)). - subst. reflexivity. - * cbn. subst. rewrite !erase_castrm. rewrite e0,e1. reflexivity. - + econv. - eapply meta_ctx_conv_conv. - * eapply IHh2 with (Δ := Δ ,, (mx, ε|A|)). - subst. reflexivity. - * cbn. subst. rewrite !erase_castrm. rewrite e0,e1. reflexivity. - - eapply cconv_trans. 1:{ eapply IHh6. eassumption. } - destruct (isKind m) eqn:ek. 1:{ mode_eqs. contradiction. } - subst. - apply cconv_irr. - + eapply param_scoping in h6. rewrite ek in h6. - rewrite <- csc_param_ctx in h6. assumption. - + rewrite csc_param_ctx. eapply meta_ccscoping_conv. - * eapply param_scoping. constructor. all: eassumption. - * rewrite ek. reflexivity. - - cbn. rewrite !erase_castrm, !revive_castrm. - destruct m. - + eapply ccong_pmifK. all: eauto. all: econv. - + eapply ccong_pmif. all: eauto. all: econv. - + eapply ccong_pmif. all: eauto. all: econv. - + econv. - - cbn. rewrite !erase_castrm, !revive_castrm. - destruct m. - + contradiction. - + econv. - + econv. - + econv. - - cbn. rewrite !erase_castrm, !revive_castrm. - destruct m. - + contradiction. - + econv. - + econv. - + econv. -Qed. - -Lemma param_castrm_conv : - ∀ Γ u v mu mv, - cscoping Γ u mu → - cscoping Γ v mv → - Γ ⊢ u ε≡ v → - ⟦ Γ ⟧p ⊢ᶜ ⟦ sc Γ | u ⟧p ≡ ⟦ sc Γ | v ⟧p. -Proof. - intros Γ u v mu mv hu hv h. - eapply param_conv in h. - eapply cconv_trans. - - apply cconv_sym. eapply param_castrm. eassumption. - - eapply cconv_trans. 1: eassumption. - eapply param_castrm. eassumption. -Qed. - (** Parametricity preserves typing **) Definition ptype Γ t A := From 12b8b027014a9ccf80f8eb82b74dd0329fca13bc Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Thu, 27 Jun 2024 14:27:26 +0200 Subject: [PATCH 12/15] ignore vim files --- .gitignore | 1 + 1 file changed, 1 insertion(+) diff --git a/.gitignore b/.gitignore index 2834be7..11041fa 100644 --- a/.gitignore +++ b/.gitignore @@ -1,3 +1,4 @@ +*.*~ .*.aux .*.d *.a From 792aacc79e7a59f8e76837188504778d7c3eae1d Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Fri, 28 Jun 2024 15:24:01 +0200 Subject: [PATCH 13/15] tri notations MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit modifié : _CoqProject modifié : theories/reduction/Notations.v modifié : theories/reduction/onestep/Properties.v modifié : theories/reduction/onestep/Reduction.v modifié : theories/reduction/onestep/SubjectReduction.v modifié : theories/reduction/wrapping/Core.v nouveau fichier : theories/reduction/wrapping/Notations.v modifié : theories/reduction/wrapping/Properties.v modifié : _CoqProject modifié : theories/reduction/Notations.v modifié : theories/reduction/onestep/Properties.v modifié : theories/reduction/onestep/Reduction.v modifié : theories/reduction/onestep/SubjectReduction.v modifié : theories/reduction/wrapping/Core.v nouveau fichier : theories/reduction/wrapping/Notations.v modifié : theories/reduction/wrapping/Properties.v modifié : theories/reduction/Notations.v modifié : _CoqProject modifié : theories/reduction/Notations.v modifié : theories/reduction/onestep/Properties.v modifié : theories/reduction/onestep/Reduction.v modifié : theories/reduction/onestep/SubjectReduction.v modifié : theories/reduction/wrapping/Core.v nouveau fichier : theories/reduction/wrapping/Notations.v modifié : theories/reduction/wrapping/Properties.v --- _CoqProject | 1 + theories/reduction/Notations.v | 64 +++++++++++++++---- theories/reduction/onestep/Properties.v | 2 +- theories/reduction/onestep/Reduction.v | 7 +- theories/reduction/onestep/SubjectReduction.v | 3 +- theories/reduction/wrapping/Core.v | 38 +++-------- theories/reduction/wrapping/Notations.v | 19 ++++++ theories/reduction/wrapping/Properties.v | 13 ++-- 8 files changed, 90 insertions(+), 57 deletions(-) create mode 100644 theories/reduction/wrapping/Notations.v diff --git a/_CoqProject b/_CoqProject index f9da9d3..c1088b4 100644 --- a/_CoqProject +++ b/_CoqProject @@ -37,6 +37,7 @@ theories/reduction/Model.v theories/Admissible.v theories/reduction/wrapping/Core.v +theories/reduction/wrapping/Notations.v theories/reduction/wrapping/Properties.v theories/reduction/onestep/Reduction.v diff --git a/theories/reduction/Notations.v b/theories/reduction/Notations.v index 1797470..e8384e4 100644 --- a/theories/reduction/Notations.v +++ b/theories/reduction/Notations.v @@ -1,30 +1,45 @@ +(* All global notations defined in the reduction folder *) +(* Some notations used in this folder come from other parts of the repository + and aren't listed here (for example from GhostTT.SubstNotations or Coq.Utf8) *) +(* List of special caracters used : + · ℙ 𝔾 𝕋 𝕂 ⊥ ⊤ ⋆ ⊢ ⊨ ∷ ⇶ ε □ ¹ ² ↣ ⇜ *) + From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. -From GhostTT Require Import BasicAST ContextDecl Scoping TermNotations. +From GhostTT Require Import BasicAST ContextDecl Scoping Typing. +From GhostTT Require Export SubstNotations. (** General notations **) (* Substitution of var 0 *) Notation "s '··'" := (scons s ids) (at level 1, format "s ··") : subst_scope. -(* scope and context *) -(* A is of scope m *) -Notation "Γ ⊢ A ∷ m" := (scoping (sc Γ) A m) - (at level 80, Γ constr, A, m at next level, format "Γ ⊢ A ∷ m"). -Notation "Γ ⊨ A ∷ m" := (scoping Γ A m) - (at level 80, A, m at next level, format "Γ ⊨ A ∷ m"). + +Declare Scope term_scope. +Open Scope term_scope. (*Mode abreviations*) -Notation ℙ := mProp. -Notation 𝔾 := mGhost. -Notation 𝕋 := mType. -Notation 𝕂 := mKind. +Notation "'ℙ'" := mProp : term_scope. +Notation "'𝔾'" := mGhost : term_scope. +Notation "'𝕋'" := mType : term_scope. +Notation "'𝕂'" := mKind : term_scope. -Notation "⊥" := bot. -Notation "⊤" := top. -Notation "⋆" := (lam ℙ ⊥ (var 0)). +(*Term abreviations*) +Notation "⊥" := bot : term_scope. +Notation "⊤" := (Pi 0 0 ℙ ℙ ⊥ ⊥) : term_scope. +Notation "⋆" := (lam ℙ ⊥ (var 0)) : term_scope. + + + +(* A is of scope m *) +Notation "Γ '⊢' A ∷ m" := (scoping (sc Γ) A m) + (at level 80, A, m at next level, format "Γ ⊢ A ∷ m") : term_scope. +Notation "Γ '⊨' A ∷ m" := (scoping Γ A m) + (at level 80, A, m at next level, format "Γ ⊨ A ∷ m") : term_scope. (** Notation for reductions **) +Declare Scope red_scope. +Open Scope red_scope. (* Multi-step reduction *) Reserved Notation "Γ ⊨ u ⇶ v" @@ -37,4 +52,25 @@ Reserved Notation "Γ ⊨ u ε⇶ v" (* Step by step reduction *) Reserved Notation "u ↣ v" (at level 80, v at next level, format "u ↣ v"). + + +(** Notation for wrapping **) +Declare Scope wrap_scope. +Open Scope wrap_scope. +Reserved Notation "□¹_term" (at level 0). +Reserved Notation "□²_term" (at level 0). +Reserved Notation "□_term" (at level 0). +Reserved Notation "C '[□/' t ']'" (at level 70). +Reserved Notation "C '[□¹/' t ']'" (at level 70). +Reserved Notation "C '[□²/' t ']'" (at level 70). + +Reserved Notation "Γ ⇜~ C" (at level 70). +Reserved Notation "Γ ⇜ C" (at level 70). + +Reserved Notation "Γ ⇜ C ⊢ t : A" + (at level 80, t, A at next level, format "Γ ⇜ C ⊢ t : A", only printing). +Reserved Notation "Γ ⇜ C ⊢ t ∷ m" + (at level 80, t, m at next level, format "Γ ⇜ C ⊢ t ∷ m", only printing). +Reserved Notation "Γ ⇜ C ⊨ t ∷ m" + (at level 80, t, m at next level, format "Γ ⇜ C ⊨ t ∷ m", only printing). diff --git a/theories/reduction/onestep/Properties.v b/theories/reduction/onestep/Properties.v index 9939999..30db354 100644 --- a/theories/reduction/onestep/Properties.v +++ b/theories/reduction/onestep/Properties.v @@ -2,7 +2,7 @@ (* and proof that the system is confluent *) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. -From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping Typing BasicMetaTheory. From GhostTT.reduction Require Import Notations Injectivity. From GhostTT.reduction.wrapping Require Import Properties. From GhostTT.reduction.onestep Require Export Reduction. diff --git a/theories/reduction/onestep/Reduction.v b/theories/reduction/onestep/Reduction.v index 068dbdd..a3e2eaa 100644 --- a/theories/reduction/onestep/Reduction.v +++ b/theories/reduction/onestep/Reduction.v @@ -2,11 +2,10 @@ (* and proof that the system is confluent *) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST. -From GhostTT Require Import SubstNotations. +From GhostTT Require Import SubstNotations Typing. From GhostTT.reduction Require Export Notations Utils. -From GhostTT.reduction.wrapping Require Export Core. +From GhostTT.reduction.wrapping Require Export Core Notations. -Import ListNotations. Set Default Goal Selector "!". (* ------------------------------------------------------------------------- *) @@ -53,7 +52,7 @@ Inductive reduction : term → term → Prop := | red_subterm : ∀ u u' C, u ↣ u' → - C □ u ↣ C □ u' + C[□/u] ↣ C[□/u'] where "u ↣ v" := (reduction u v). diff --git a/theories/reduction/onestep/SubjectReduction.v b/theories/reduction/onestep/SubjectReduction.v index 2960f44..f5b490a 100644 --- a/theories/reduction/onestep/SubjectReduction.v +++ b/theories/reduction/onestep/SubjectReduction.v @@ -1,7 +1,7 @@ From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval - TermMode Scoping BasicMetaTheory. + TermMode Scoping Typing BasicMetaTheory. From GhostTT.reduction Require Import Injectivity Model. From GhostTT.reduction.wrapping Require Import Core Properties. From GhostTT.reduction.onestep Require Export Reduction. @@ -22,6 +22,7 @@ Section subjet_reduction. |_ => idtac end in destruct_conj h. + Open Scope subst_scope. Local Notation 𝝢 := (Erased tnat). Local Notation "P ⁺¹" := (S ⋅ P) (at level 1). Local Notation "P ⁺²" := (S ⋅ S ⋅ P) (at level 1). diff --git a/theories/reduction/wrapping/Core.v b/theories/reduction/wrapping/Core.v index f3282cb..943e40b 100644 --- a/theories/reduction/wrapping/Core.v +++ b/theories/reduction/wrapping/Core.v @@ -1,13 +1,8 @@ -(* Definition of reduction rules which corresponds to the congruence relation *) -(* and proof that the system is confluent *) -From Coq Require Import Utf8 List. -From GhostTT.autosubst Require Import GAST unscoped. -From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping. -From GhostTT Require Export Univ TermNotations Typing. -From GhostTT.reduction Require Export Notations Utils. - -Import ListNotations. -Set Default Goal Selector "!". +(* Definition of partial terms and of a wrapping to complete it*) +From Coq Require Import Utf8. +From GhostTT.autosubst Require Import GAST. +From GhostTT Require Import BasicAST ContextDecl. +From GhostTT.reduction Require Export Notations. Inductive box_term_1 : Set := @@ -71,7 +66,7 @@ Inductive box_term : Set := | Box_1 : box_term_1 → box_term | Box_2 : box_term_2 → box_term. -Definition eval_box_term (C : box_term) (t0: term) : term := +Definition wrapping (C : box_term) (t0: term) : term := match C with | Box_1 (Box_Pi_1 i j m mx B) => Pi i j m mx t0 B | Box_2 (Box_Pi_2 i j m mx A) => Pi i j m mx A t0 @@ -128,34 +123,17 @@ Definition eval_box_term (C : box_term) (t0: term) : term := | Box_1 (Box_bot_elim_2 m A) => bot_elim m A t0 end. -Definition context_box_term (Γ: context) (C : box_term) : context := +Definition wrapping_context (Γ: context) (C : box_term) : context := match C with | Box_1 _ => Γ | Box_2 (Box_Pi_2 i j m mx A) => (Γ,, (mx,A)) | Box_2 (Box_lam_2 m A) => (Γ,, (m,A)) end. -Definition scope_box_term (Γ: scope) (C : box_term) : scope := +Definition wrapping_scope (Γ: scope) (C : box_term) : scope := match C with | Box_1 _ => Γ | Box_2 (Box_Pi_2 i j m mx A) => mx::Γ | Box_2 (Box_lam_2 m A) => m::Γ end. -Notation "□¹_term" := box_term_1. -Notation "□²_term" := box_term_2. -Notation "□_term" := box_term. -Notation "C □ t" := (eval_box_term C t) (at level 70). -Notation "C □¹ t" := (eval_box_term (Box_1 C) t) (at level 70). -Notation "C □² t" := (eval_box_term (Box_2 C) t) (at level 70). - -Notation "[ Γ ] □ C" := (context_box_term Γ C) (at level 70). -Notation "[| Γ |] □ C" := (scope_box_term Γ C) (at level 70). - -Notation "Γ □ C ⊢ t : A" := (typing (context_box_term Γ C) t A) - (at level 80, t, A at next level, format "Γ □ C ⊢ t : A", only printing). -Notation "Γ □ C ⊢ t ∷ m" := (scoping (sc (context_box_term Γ C)) t m) - (at level 80, t, m at next level, format "Γ □ C ⊢ t ∷ m", only printing). -Notation "Γ □ C ⊨ t ∷ m" := (scoping (scope_box_term Γ C) t m) - (at level 80, t, m at next level, format "Γ □ C ⊨ t ∷ m", only printing). - diff --git a/theories/reduction/wrapping/Notations.v b/theories/reduction/wrapping/Notations.v new file mode 100644 index 0000000..84a1747 --- /dev/null +++ b/theories/reduction/wrapping/Notations.v @@ -0,0 +1,19 @@ +From Coq Require Import Utf8. +From GhostTT Require Import Scoping ContextDecl Typing. +From GhostTT.reduction Require Export Notations. +From GhostTT.reduction.wrapping Require Import Core. + +Notation "□¹_term" := box_term_1 : wrap_scope. +Notation "□²_term" := box_term_2 : wrap_scope. +Notation "□_term" := box_term : wrap_scope. +Notation "C [□/ t ]" := (wrapping C t) : wrap_scope. +Notation "C [□¹/ t ]" := (wrapping (Box_1 C) t) : wrap_scope. +Notation "C [□²/ t ]" := (wrapping (Box_2 C) t) : wrap_scope. + +Notation "Γ ⇜~ C" := (wrapping_context Γ C) : wrap_scope. +Notation "Γ ⇜ C" := (wrapping_scope Γ C) : wrap_scope. + +Notation "Γ ⇜ C ⊢ t : A" := (typing (wrapping_context Γ C) t A) : wrap_scope. +Notation "Γ ⇜ C ⊢ t ∷ m" := (scoping (sc (wrapping_context Γ C)) t m) : wrap_scope. +Notation "Γ ⇜ C ⊨ t ∷ m" := (scoping (wrapping_scope Γ C) t m) : wrap_scope. + diff --git a/theories/reduction/wrapping/Properties.v b/theories/reduction/wrapping/Properties.v index 8a470df..2382808 100644 --- a/theories/reduction/wrapping/Properties.v +++ b/theories/reduction/wrapping/Properties.v @@ -2,10 +2,9 @@ (* and proof that the system is confluent *) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. -From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. -From GhostTT.reduction.wrapping Require Export Core. +From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping Typing BasicMetaTheory. +From GhostTT.reduction.wrapping Require Export Core Notations. -Import ListNotations. Set Default Goal Selector "!". Ltac ttinv_destruct h HN:= @@ -18,7 +17,7 @@ Ltac ttinv_destruct h HN:= in destruct_conj HN. Lemma scoping_box {Γ : scope} {m : mode } {u : term} {C: □_term} : - Γ ⊨ C □ u∷m → [|Γ|] □ C ⊨ u∷md ([|Γ|] □ C) u. + Γ ⊨ C[□/u] ∷ m → Γ⇜C ⊨ u∷md (Γ⇜C) u. Proof. intros scope_Cu. destruct C as [C|C]. @@ -29,7 +28,7 @@ Proof. Qed. Lemma scoping_change_box {Γ : scope} {m : mode } {u u': term} {C: □_term} : - Γ ⊨ C □ u∷m → [|Γ|] □ C ⊨ u'∷md ([|Γ|] □ C) u → Γ ⊨ C □ u'∷m. + Γ ⊨ C[□/u]∷m → Γ⇜C ⊨ u'∷md (Γ⇜C) u → Γ ⊨ C[□/u'] ∷ m. Proof. intros scope_Cu scope_u'. destruct C as [C|C]. @@ -39,7 +38,7 @@ Proof. Qed. Lemma typing_box {Γ : context} {u A: term} {C: □_term} : - Γ ⊢ C □ u : A → ∃ B, [Γ] □ C ⊢ u : B. + Γ ⊢ C[□/u] : A → ∃ B, Γ⇜~C ⊢ u : B. Proof. intro type_Cu. destruct C as [C|C]. @@ -49,7 +48,7 @@ Proof. Qed. Lemma sc_scope_box_term {Γ : context} {C: □_term} : - [| sc Γ |] □ C = sc ([ Γ ] □ C). + (sc Γ) ⇜ C = sc (Γ ⇜~ C). Proof. destruct C as [C|[C|C]]; reflexivity. Qed. From fd645ef853ebc8373eb9407a7a4d378694acd73b Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Mon, 15 Jul 2024 13:15:54 +0200 Subject: [PATCH 14/15] documentation --- theories/reduction/Injectivity.v | 1 + theories/reduction/Notations.v | 2 +- theories/reduction/Utils.v | 2 +- theories/reduction/multisteps/Confluence.v | 3 +-- theories/reduction/multisteps/Properties.v | 3 +-- theories/reduction/multisteps/ReductionToCongruence.v | 2 ++ theories/reduction/multisteps/Rho.v | 3 +-- theories/reduction/multisteps/Transitivity.v | 1 + theories/reduction/wrapping/Properties.v | 3 +-- 9 files changed, 10 insertions(+), 10 deletions(-) diff --git a/theories/reduction/Injectivity.v b/theories/reduction/Injectivity.v index 4aff5bc..a15aa6c 100644 --- a/theories/reduction/Injectivity.v +++ b/theories/reduction/Injectivity.v @@ -1,3 +1,4 @@ +(* injectivity of the constructors of terms using the multistep reduction *) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval diff --git a/theories/reduction/Notations.v b/theories/reduction/Notations.v index e8384e4..506a4f2 100644 --- a/theories/reduction/Notations.v +++ b/theories/reduction/Notations.v @@ -1,6 +1,6 @@ (* All global notations defined in the reduction folder *) (* Some notations used in this folder come from other parts of the repository - and aren't listed here (for example from GhostTT.SubstNotations or Coq.Utf8) *) + and aren't listed here (for example those from GhostTT.SubstNotations or Coq.Utf8) *) (* List of special caracters used : · ℙ 𝔾 𝕋 𝕂 ⊥ ⊤ ⋆ ⊢ ⊨ ∷ ⇶ ε □ ¹ ² ↣ ⇜ *) diff --git a/theories/reduction/Utils.v b/theories/reduction/Utils.v index bd8e061..cfed74f 100644 --- a/theories/reduction/Utils.v +++ b/theories/reduction/Utils.v @@ -1,4 +1,4 @@ -(* Definition of multistep reduction rules which corresponds to the congruence relation *) +(* Definition of a few functions and lemma for the reductions*) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl diff --git a/theories/reduction/multisteps/Confluence.v b/theories/reduction/multisteps/Confluence.v index 8c85861..d802c5a 100644 --- a/theories/reduction/multisteps/Confluence.v +++ b/theories/reduction/multisteps/Confluence.v @@ -1,5 +1,4 @@ -(* Definition of reduction rules which corresponds to the congruence relation *) -(* and proof that the system is confluent *) +(* Proof that the multistep reduction is confluent *) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. diff --git a/theories/reduction/multisteps/Properties.v b/theories/reduction/multisteps/Properties.v index cf2be11..ec45ba9 100644 --- a/theories/reduction/multisteps/Properties.v +++ b/theories/reduction/multisteps/Properties.v @@ -1,5 +1,4 @@ -(* Definition of reduction rules which corresponds to the congruence relation *) -(* and proof that the system is confluent *) +(* Definition of the main properties and and inversion lemmas of the multistep reduction*) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. diff --git a/theories/reduction/multisteps/ReductionToCongruence.v b/theories/reduction/multisteps/ReductionToCongruence.v index ffb21a9..930c676 100644 --- a/theories/reduction/multisteps/ReductionToCongruence.v +++ b/theories/reduction/multisteps/ReductionToCongruence.v @@ -1,3 +1,5 @@ +(* Proof that the conversion is the symetric ans transitive closure + of the multistep reduction *) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval diff --git a/theories/reduction/multisteps/Rho.v b/theories/reduction/multisteps/Rho.v index be59b0c..847c261 100644 --- a/theories/reduction/multisteps/Rho.v +++ b/theories/reduction/multisteps/Rho.v @@ -1,5 +1,4 @@ -(* Definition of reduction rules which corresponds to the congruence relation *) -(* and proof that the system is confluent *) +(*Triangle property via a function rho used in the proof that the multistep reduction system is confluent *) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping BasicMetaTheory. diff --git a/theories/reduction/multisteps/Transitivity.v b/theories/reduction/multisteps/Transitivity.v index e994627..ca9ab43 100644 --- a/theories/reduction/multisteps/Transitivity.v +++ b/theories/reduction/multisteps/Transitivity.v @@ -1,3 +1,4 @@ +(* Definition of ⇶* , the transitive closure of ⇶ and inversion lemmas of ⇶* *) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval diff --git a/theories/reduction/wrapping/Properties.v b/theories/reduction/wrapping/Properties.v index 2382808..918e139 100644 --- a/theories/reduction/wrapping/Properties.v +++ b/theories/reduction/wrapping/Properties.v @@ -1,5 +1,4 @@ -(* Definition of reduction rules which corresponds to the congruence relation *) -(* and proof that the system is confluent *) +(* basic properties of wrapping *) From Coq Require Import Utf8 List. From GhostTT.autosubst Require Import GAST unscoped. From GhostTT Require Import Util BasicAST SubstNotations ContextDecl CastRemoval TermMode Scoping Typing BasicMetaTheory. From 8a6d2a73f35e760de3d60358ae796596086f0a49 Mon Sep 17 00:00:00 2001 From: Ewen Broudin-Caradec Date: Thu, 18 Jul 2024 14:43:36 +0200 Subject: [PATCH 15/15] autosubst Makefile --- theories/autosubst/Makefile | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/theories/autosubst/Makefile b/theories/autosubst/Makefile index c5e9187..c5cd16e 100644 --- a/theories/autosubst/Makefile +++ b/theories/autosubst/Makefile @@ -1,3 +1,11 @@ syntax: autosubst -f -s ucoq -v ge813 -allfv GAST.sig -o GAST.v + sed -i '1s/^/From GhostTT Require Import BasicAST.\n/' GAST.v + sed -i '2s/^/From GhostTT.autosubst /' GAST.v autosubst -f -s ucoq -v ge813 -allfv CCAST.sig -o CCAST.v + sed -i '1s/^/From GhostTT Require Import BasicAST.\n/' CCAST.v + sed -i '2s/^/From GhostTT.autosubst /' CCAST.v + sed -i '10s/^/From GhostTT.autosubst /' unscoped.v + +clean: + rm *.v