From 55a9ad9d9396063bd3e6b497df19247b76865343 Mon Sep 17 00:00:00 2001 From: Junyoung/Clare Jang Date: Wed, 4 Sep 2024 18:20:40 -0400 Subject: [PATCH] Use consistent naming --- theories/Algorithmic/Subtyping/Definitions.v | 25 +- theories/Algorithmic/Subtyping/Lemmas.v | 64 ++-- theories/Algorithmic/Typing/Lemmas.v | 33 +- theories/Core/Completeness/ContextCases.v | 4 +- theories/Core/Completeness/FunctionCases.v | 36 +- .../LogicalRelation/Definitions.v | 20 +- .../Completeness/LogicalRelation/Lemmas.v | 16 +- theories/Core/Completeness/NatCases.v | 222 +++++++----- .../Core/Completeness/SubstitutionCases.v | 4 +- theories/Core/Completeness/SubtypingCases.v | 23 +- .../Core/Completeness/TermStructureCases.v | 24 +- theories/Core/Completeness/UniverseCases.v | 8 +- theories/Core/Semantic/Domain.v | 20 +- .../Core/Semantic/Evaluation/Definitions.v | 93 ++--- theories/Core/Semantic/Evaluation/Lemmas.v | 102 +++--- theories/Core/Semantic/Evaluation/Tactics.v | 6 +- theories/Core/Semantic/NbE.v | 30 +- theories/Core/Semantic/PER/Definitions.v | 61 ++-- theories/Core/Semantic/PER/Lemmas.v | 225 ++++++------ theories/Core/Semantic/Readback/Definitions.v | 22 +- theories/Core/Semantic/Readback/Lemmas.v | 70 ++-- theories/Core/Semantic/Realizability.v | 26 +- theories/Core/Soundness/FunctionCases.v | 47 +-- theories/Core/Soundness/LogicalRelation.v | 2 +- .../Core/Soundness/LogicalRelation/Core.v | 2 +- .../Soundness/LogicalRelation/CoreLemmas.v | 325 +++++++++--------- .../Soundness/LogicalRelation/CoreTactics.v | 13 + .../Soundness/LogicalRelation/Definitions.v | 255 +++++++------- .../Core/Soundness/LogicalRelation/Lemmas.v | 213 ++++++------ theories/Core/Soundness/NatCases.v | 89 +++-- theories/Core/Soundness/Realizability.v | 159 ++++----- theories/Core/Soundness/SubstitutionCases.v | 7 +- theories/Core/Soundness/TermStructureCases.v | 3 - .../Core/Soundness/Weakening/Definition.v | 2 +- theories/Core/Soundness/Weakening/Lemmas.v | 6 +- theories/Extraction/TypeCheck.v | 38 +- 36 files changed, 1191 insertions(+), 1104 deletions(-) diff --git a/theories/Algorithmic/Subtyping/Definitions.v b/theories/Algorithmic/Subtyping/Definitions.v index 3013865f..78262e0b 100644 --- a/theories/Algorithmic/Subtyping/Definitions.v +++ b/theories/Algorithmic/Subtyping/Definitions.v @@ -13,10 +13,10 @@ Definition not_univ_pi (A : nf) : Prop := end. Inductive alg_subtyping_nf : nf -> nf -> Prop := -| asnf_refl : forall M N, - not_univ_pi M -> - M = N -> - {{ ⊢anf M ⊆ N }} +| asnf_refl : forall A A', + not_univ_pi A -> + A = A' -> + {{ ⊢anf A ⊆ A' }} | asnf_univ : forall i j, i <= j -> {{ ⊢anf Type@i ⊆ Type@j }} @@ -24,16 +24,15 @@ Inductive alg_subtyping_nf : nf -> nf -> Prop := A = A' -> {{ ⊢anf B ⊆ B' }} -> {{ ⊢anf Π A B ⊆ Π A' B' }} -where "⊢anf M ⊆ N" := (alg_subtyping_nf M N) (in custom judg) : type_scope. - +where "⊢anf A ⊆ A'" := (alg_subtyping_nf A A') (in custom judg) : type_scope. Inductive alg_subtyping : ctx -> typ -> typ -> Prop := -| alg_subtyp_run : forall Γ M N A B, - nbe_ty Γ M A -> - nbe_ty Γ N B -> - {{ ⊢anf A ⊆ B }} -> - {{ Γ ⊢a M ⊆ N }} -where "Γ ⊢a M ⊆ N" := (alg_subtyping Γ M N) (in custom judg) : type_scope. +| alg_subtyp_run : forall Γ A B A' B', + nbe_ty Γ A A' -> + nbe_ty Γ B B' -> + {{ ⊢anf A' ⊆ B' }} -> + {{ Γ ⊢a A ⊆ B }} +where "Γ ⊢a A ⊆ B" := (alg_subtyping Γ A B) (in custom judg) : type_scope. #[export] - Hint Constructors alg_subtyping_nf alg_subtyping: mcltt. +Hint Constructors alg_subtyping_nf alg_subtyping: mcltt. diff --git a/theories/Algorithmic/Subtyping/Lemmas.v b/theories/Algorithmic/Subtyping/Lemmas.v index 421a4d65..fb223de8 100644 --- a/theories/Algorithmic/Subtyping/Lemmas.v +++ b/theories/Algorithmic/Subtyping/Lemmas.v @@ -8,17 +8,17 @@ Import Syntax_Notations. #[local] Ltac apply_subtyping := repeat match goal with - | H : {{ ~?Γ ⊢ ~?A : ~?M }}, - H1 : {{ ~?Γ ⊢ ~?M ⊆ ~?N }} |- _ => - assert {{ Γ ⊢ A : N }} by mauto; clear H + | H : {{ ~?Γ ⊢ ~?M : ~?A }}, + H1 : {{ ~?Γ ⊢ ~?A ⊆ ~?B }} |- _ => + assert {{ Γ ⊢ M : B }} by mauto; clear H end. -Lemma alg_subtyping_nf_sound : forall M N, - {{ ⊢anf M ⊆ N }} -> +Lemma alg_subtyping_nf_sound : forall A B, + {{ ⊢anf A ⊆ B }} -> forall Γ i, - {{ Γ ⊢ M : Type@i }} -> - {{ Γ ⊢ N : Type@i }} -> - {{ Γ ⊢ M ⊆ N }}. + {{ Γ ⊢ A : Type@i }} -> + {{ Γ ⊢ B : Type@i }} -> + {{ Γ ⊢ A ⊆ B }}. Proof. induction 1; intros; subst; simpl in *. - eapply wf_subtyp_refl'; mauto. @@ -27,7 +27,7 @@ Proof. destruct_all. gen_presups. repeat match goal with - | H : {{ ~?Γ ⊢ ~?M ⊆ ~?N }}, H1: {{ ⊢ ~?Γ , ~_ }} |- _ => + | H : {{ ~?Γ ⊢ ~?A ⊆ ~?B }}, H1: {{ ⊢ ~?Γ , ~_ }} |- _ => pose proof (wf_subtyp_univ_weaken _ _ _ _ H H1); fail_if_dup end. @@ -38,12 +38,12 @@ Proof. mauto 3. Qed. -Lemma alg_subtyping_nf_trans : forall M1 M0 M2, - {{ ⊢anf M0 ⊆ M1 }} -> - {{ ⊢anf M1 ⊆ M2 }} -> - {{ ⊢anf M0 ⊆ M2 }}. +Lemma alg_subtyping_nf_trans : forall A0 A1 A2, + {{ ⊢anf A0 ⊆ A1 }} -> + {{ ⊢anf A1 ⊆ A2 }} -> + {{ ⊢anf A0 ⊆ A2 }}. Proof. - intros * H1; gen M2. + intros * H1; gen A2. induction H1; subst; intros ? H2; dependent destruction H2; simpl in *; @@ -53,20 +53,20 @@ Proof. constructor; lia. Qed. -Lemma alg_subtyping_nf_refl : forall M, - {{ ⊢anf M ⊆ M }}. +Lemma alg_subtyping_nf_refl : forall A, + {{ ⊢anf A ⊆ A }}. Proof. - intro M; induction M; + induction A; solve [constructor; simpl; trivial]. Qed. #[local] Hint Resolve alg_subtyping_nf_trans alg_subtyping_nf_refl : mcltt. -Lemma alg_subtyping_trans : forall Γ M0 M1 M2, - {{ Γ ⊢a M0 ⊆ M1 }} -> - {{ Γ ⊢a M1 ⊆ M2 }} -> - {{ Γ ⊢a M0 ⊆ M2 }}. +Lemma alg_subtyping_trans : forall Γ A0 A1 A2, + {{ Γ ⊢a A0 ⊆ A1 }} -> + {{ Γ ⊢a A1 ⊆ A2 }} -> + {{ Γ ⊢a A0 ⊆ A2 }}. Proof. intros. progressive_inversion. functional_nbe_rewrite_clear. @@ -76,9 +76,9 @@ Qed. #[local] Hint Resolve alg_subtyping_trans : mcltt. -Lemma alg_subtyping_complete : forall Γ M N, - {{ Γ ⊢ M ⊆ N }} -> - {{ Γ ⊢a M ⊆ N }}. +Lemma alg_subtyping_complete : forall Γ A B, + {{ Γ ⊢ A ⊆ B }} -> + {{ Γ ⊢a A ⊆ B }}. Proof. induction 1; mauto. - apply completeness in H0 as [W [? ?]]. @@ -106,11 +106,11 @@ Proof. mauto 2. Qed. -Lemma alg_subtyping_sound : forall Γ M N i, - {{ Γ ⊢a M ⊆ N }} -> - {{ Γ ⊢ M : Type@i }} -> - {{ Γ ⊢ N : Type@i }} -> - {{ Γ ⊢ M ⊆ N }}. +Lemma alg_subtyping_sound : forall Γ A B i, + {{ Γ ⊢a A ⊆ B }} -> + {{ Γ ⊢ A : Type@i }} -> + {{ Γ ⊢ B : Type@i }} -> + {{ Γ ⊢ A ⊆ B }}. Proof. intros. destruct H. on_all_hyp: fun H => apply soundness in H. @@ -118,8 +118,8 @@ Proof. on_all_hyp: fun H => apply nbe_type_to_nbe_ty in H. functional_nbe_rewrite_clear. gen_presups. - assert {{ Γ ⊢ A ⊆ B }} by mauto 3 using alg_subtyping_nf_sound. - transitivity A; [mauto |]. - transitivity B; [eassumption |]. + assert {{ Γ ⊢ A' ⊆ B' }} by mauto 3 using alg_subtyping_nf_sound. + transitivity A'; [mauto |]. + transitivity B'; [eassumption |]. mauto. Qed. diff --git a/theories/Algorithmic/Typing/Lemmas.v b/theories/Algorithmic/Typing/Lemmas.v index d96e4a68..14cd716f 100644 --- a/theories/Algorithmic/Typing/Lemmas.v +++ b/theories/Algorithmic/Typing/Lemmas.v @@ -125,7 +125,7 @@ Proof with (f_equiv; mautosolve 4). simplify_evals. dir_inversion_by_head read_typ; subst. functional_initial_env_rewrite_clear. - assert (initial_env {{{ Γ, ~(C : exp) }}} d{{{ p ↦ ⇑! a (length Γ) }}}) by mauto 3. + assert (initial_env {{{ Γ, ~(C : exp) }}} d{{{ ρ ↦ ⇑! a (length Γ) }}}) by mauto 3. assert (nbe_ty Γ A C) by mauto 3. assert (nbe_ty Γ C A0) by mauto 3. replace A0 with C by mauto 3. @@ -148,9 +148,9 @@ Lemma alg_type_check_typ_implies_alg_type_infer_typ : forall {Γ A i}, exists j, {{ Γ ⊢a A ⟹ Type@j }} /\ j <= i. Proof. intros * ? Hcheck. - inversion Hcheck as [? ? ? ? Hinfer Hsub]; subst. - inversion Hsub as [? ? ? ? ? Hnbe1 Hnbe2 Hnfsub]; subst. - replace A0 with A1 in * by (symmetry; mauto 3). + inversion Hcheck as [? A' ? ? Hinfer Hsub]; subst. + inversion Hsub as [? ? ? A'' ? Hnbe1 Hnbe2 Hnfsub]; subst. + replace A' with A'' in * by (symmetry; mauto 3). inversion Hnbe2; subst. dir_inversion_by_head eval_exp; subst. dir_inversion_by_head read_typ; subst. @@ -169,9 +169,9 @@ Lemma alg_type_check_pi_implies_alg_type_infer_pi : forall {Γ M A B i}, Proof. intros * ? ? Hcheck. assert ({{ Γ ⊢ A : Type@i }} /\ {{ Γ, A ⊢ B : Type@i }}) as [] by mauto 3. - inversion Hcheck as [? ? ? ? Hinfer Hsub]; subst. - inversion Hsub as [? ? ? ? C Hnbe1 Hnbe2 Hnfsub]; subst. - replace A0 with A1 in * by (symmetry; mauto 3). + inversion Hcheck as [? A' ? ? Hinfer Hsub]; subst. + inversion Hsub as [? ? ? A'' C Hnbe1 Hnbe2 Hnfsub]; subst. + replace A' with A'' in * by (symmetry; mauto 3). inversion Hnbe2; subst. dir_inversion_by_head eval_exp; subst. dir_inversion_by_head read_typ; subst. @@ -181,18 +181,18 @@ Proof. dir_inversion_by_head eval_exp; subst. dir_inversion_by_head read_typ; subst. simpl in *. - assert {{ Γ ⊢ M : ~n{{{ Π A2 B0 }}} }} by mauto 3 using alg_type_infer_sound. - assert (exists j, {{ Γ ⊢ Π A2 B0 : Type@j }}) as [j] by (gen_presups; eauto 2). - assert ({{ Γ ⊢ A2 : Type@j }} /\ {{ Γ, ~(A2 : exp) ⊢ B0 : Type@j }}) as [] by mauto 3. + assert {{ Γ ⊢ M : ~n{{{ Π A0 B0 }}} }} by mauto 3 using alg_type_infer_sound. + assert (exists j, {{ Γ ⊢ Π A0 B0 : Type@j }}) as [j] by (gen_presups; eauto 2). + assert ({{ Γ ⊢ A0 : Type@j }} /\ {{ Γ, ~(A0 : exp) ⊢ B0 : Type@j }}) as [] by mauto 3. do 2 eexists. - assert (nbe_ty Γ A A2) by mauto 3. - assert {{ Γ ⊢ A ≈ A2 : Type@i }} by mauto 3 using soundness_ty'. + assert (nbe_ty Γ A A0) by mauto 3. + assert {{ Γ ⊢ A ≈ A0 : Type@i }} by mauto 3 using soundness_ty'. repeat split; mauto 3. - assert (initial_env {{{ Γ, A }}} d{{{ p ↦ ⇑! a (length Γ) }}}) by mauto 3. + assert (initial_env {{{ Γ, A }}} d{{{ ρ ↦ ⇑! a (length Γ) }}}) by mauto 3. assert (nbe_ty {{{ Γ, A }}} B B') by mauto 3. - assert (initial_env {{{ Γ, ~(A2 : exp) }}} d{{{ p ↦ ⇑! a0 (length Γ) }}}) by mauto 3. - assert (nbe_ty {{{ Γ, ~(A2 : exp) }}} B0 B0) by mauto 3. - assert {{ ⊢ Γ, A ≈ Γ, ~(A2 : exp) }} by mauto 3. + assert (initial_env {{{ Γ, ~(A0 : exp) }}} d{{{ ρ ↦ ⇑! a0 (length Γ) }}}) by mauto 3. + assert (nbe_ty {{{ Γ, ~(A0 : exp) }}} B0 B0) by mauto 3. + assert {{ ⊢ Γ, A ≈ Γ, ~(A0 : exp) }} by mauto 3. assert (nbe_ty {{{ Γ, A }}} B0 B0) by mauto 4 using ctxeq_nbe_ty_eq'. mauto 3. Qed. @@ -312,4 +312,3 @@ Qed. #[export] Hint Resolve alg_type_infer_pi_complete : mcltt. - diff --git a/theories/Core/Completeness/ContextCases.v b/theories/Core/Completeness/ContextCases.v index c83daf65..aa2a338c 100644 --- a/theories/Core/Completeness/ContextCases.v +++ b/theories/Core/Completeness/ContextCases.v @@ -25,9 +25,9 @@ Proof with intuition. handle_per_ctx_env_irrel. eexists. per_ctx_env_econstructor; eauto. - - instantiate (1 := fun p p' (equiv_p_p' : env_relΓ p p') m m' => + - instantiate (1 := fun ρ ρ' (equiv_ρ_ρ' : env_relΓ ρ ρ') m m' => forall i R, - rel_typ i A p A' p' R -> + rel_typ i A ρ A' ρ' R -> R m m'). intros. (on_all_hyp: destruct_rel_by_assumption env_relΓ). diff --git a/theories/Core/Completeness/FunctionCases.v b/theories/Core/Completeness/FunctionCases.v index 3fa0ea4b..757729be 100644 --- a/theories/Core/Completeness/FunctionCases.v +++ b/theories/Core/Completeness/FunctionCases.v @@ -7,11 +7,11 @@ Import Domain_Notations. Lemma rel_exp_of_pi_inversion : forall {Γ M M' A B}, {{ Γ ⊨ M ≈ M' : Π A B }} -> exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}) i, - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), exists in_rel out_rel, - rel_typ i A p A p' in_rel /\ - (forall c c' (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_typ i B d{{{ p ↦ c }}} B d{{{ p' ↦ c' }}} (out_rel c c' equiv_c_c')) /\ - rel_exp M p M' p' + rel_typ i A ρ A ρ' in_rel /\ + (forall c c' (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_typ i B d{{{ ρ ↦ c }}} B d{{{ ρ' ↦ c' }}} (out_rel c c' equiv_c_c')) /\ + rel_exp M ρ M' ρ' (fun f f' : domain => forall (c c' : domain) (equiv_c_c' : in_rel c c'), rel_mod_app f c f' c' (out_rel c c' equiv_c_c')). Proof. intros * [env_relΓ]. @@ -26,11 +26,11 @@ Qed. Lemma rel_exp_of_pi : forall {Γ M M' A B}, (exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}) i j, - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), exists in_rel out_rel, - rel_typ i A p A p' in_rel /\ - (forall c c' (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_typ j B d{{{ p ↦ c }}} B d{{{ p' ↦ c' }}} (out_rel c c' equiv_c_c')) /\ - rel_exp M p M' p' + rel_typ i A ρ A ρ' in_rel /\ + (forall c c' (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_typ j B d{{{ ρ ↦ c }}} B d{{{ ρ' ↦ c' }}} (out_rel c c' equiv_c_c')) /\ + rel_exp M ρ M' ρ' (fun f f' : domain => forall (c c' : domain) (equiv_c_c' : in_rel c c'), rel_mod_app f c f' c' (out_rel c c' equiv_c_c'))) -> {{ Γ ⊨ M ≈ M' : Π A B }}. Proof. @@ -57,9 +57,9 @@ Ltac eexists_rel_exp_of_pi := eexists. #[local] -Ltac extract_output_info_with p c p' c' env_rel := +Ltac extract_output_info_with ρ c ρ' c' env_rel := let Hequiv := fresh "equiv" in - (assert (Hequiv : {{ Dom p ↦ c ≈ p' ↦ c' ∈ env_rel }}) by (apply_relation_equivalence; mauto 4); + (assert (Hequiv : {{ Dom ρ ↦ c ≈ ρ' ↦ c' ∈ env_rel }}) by (apply_relation_equivalence; mauto 4); apply_relation_equivalence; (on_all_hyp: fun H => destruct (H _ _ Hequiv)); destruct_conjs; @@ -130,7 +130,7 @@ Proof with mautosolve. eexists_rel_exp_of_typ. intros. (on_all_hyp: destruct_rel_by_assumption env_relΓ). - assert {{ Dom o' ≈ o' ∈ env_relΔ }} by (etransitivity; [symmetry |]; eassumption). + assert {{ Dom ρ'σ' ≈ ρ'σ' ∈ env_relΔ }} by (etransitivity; [symmetry |]; eassumption). (on_all_hyp: destruct_rel_by_assumption env_relΔ). destruct_by_head per_univ. handle_per_univ_elem_irrel. @@ -139,7 +139,7 @@ Proof with mautosolve. per_univ_elem_econstructor; eauto. - eapply rel_exp_pi_core; eauto; try reflexivity. intros. - extract_output_info_with o c o' c' env_relΔA... + extract_output_info_with ρσ c ρ'σ' c' env_relΔA... - (* `reflexivity` does not work as (simple) unification fails for some unknown reason. *) apply Equivalence_Reflexive. Qed. @@ -166,11 +166,11 @@ Proof with mautosolve. repeat split; [econstructor | | econstructor]; mauto. - eapply rel_exp_pi_core; eauto; try reflexivity. intros. - extract_output_info_with p c p' c' env_relΓA. + extract_output_info_with ρ c ρ' c' env_relΓA. econstructor; eauto. eexists... - intros. - extract_output_info_with p c p' c' env_relΓA. + extract_output_info_with ρ c ρ' c' env_relΓA. econstructor; mauto. intros. destruct_by_head rel_typ. @@ -202,12 +202,12 @@ Proof with mautosolve. clear dependent c. clear dependent c'. intros. - extract_output_info_with o c o' c' env_relΔA. + extract_output_info_with ρσ c ρ'σ' c' env_relΔA. econstructor; eauto. eexists. eapply per_univ_elem_cumu_max_left... - intros ? **. - extract_output_info_with o c o' c' env_relΔA. + extract_output_info_with ρσ c ρ'σ' c' env_relΔA. econstructor; mauto. intros. destruct_by_head rel_typ. @@ -228,7 +228,7 @@ Proof with intuition. handle_per_ctx_env_irrel. eexists_rel_exp. intros. - assert (equiv_p'_p' : env_relΓ p' p') by (etransitivity; [symmetry |]; eassumption). + assert (equiv_p'_p' : env_relΓ ρ' ρ') by (etransitivity; [symmetry |]; eassumption). (on_all_hyp: destruct_rel_by_assumption env_relΓ). rename x2 into in_rel. destruct_by_head rel_typ. @@ -290,7 +290,7 @@ Proof with mautosolve. destruct_by_head rel_exp. rename m into n. rename m' into n'. - extract_output_info_with p n p' n' env_relΓA. + extract_output_info_with ρ n ρ' n' env_relΓA. eexists. split; econstructor... Qed. diff --git a/theories/Core/Completeness/LogicalRelation/Definitions.v b/theories/Core/Completeness/LogicalRelation/Definitions.v index cf09a2cb..7ddc4b6e 100644 --- a/theories/Core/Completeness/LogicalRelation/Definitions.v +++ b/theories/Core/Completeness/LogicalRelation/Definitions.v @@ -3,8 +3,8 @@ From Mcltt Require Import Base. From Mcltt.Core Require Export PER. Import Domain_Notations. -Inductive rel_exp M p M' p' (R : relation domain) : Prop := -| mk_rel_exp : forall m m', {{ ⟦ M ⟧ p ↘ m }} -> {{ ⟦ M' ⟧ p' ↘ m' }} -> {{ Dom m ≈ m' ∈ R }} -> rel_exp M p M' p' R. +Inductive rel_exp M ρ M' ρ' (R : relation domain) : Prop := +| mk_rel_exp : forall m m', {{ ⟦ M ⟧ ρ ↘ m }} -> {{ ⟦ M' ⟧ ρ' ↘ m' }} -> {{ Dom m ≈ m' ∈ R }} -> rel_exp M ρ M' ρ' R. #[global] Arguments mk_rel_exp {_ _ _ _ _}. #[export] @@ -12,9 +12,9 @@ Hint Constructors rel_exp : mcltt. Definition rel_exp_under_ctx Γ A M M' := exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}) i, - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), exists (elem_rel : relation domain), - rel_typ i A p A p' elem_rel /\ rel_exp M p M' p' elem_rel. + rel_typ i A ρ A ρ' elem_rel /\ rel_exp M ρ M' ρ' elem_rel. Definition valid_exp_under_ctx Γ A M := rel_exp_under_ctx Γ A M M. #[global] @@ -24,8 +24,8 @@ Hint Transparent valid_exp_under_ctx : mcltt. #[export] Hint Unfold valid_exp_under_ctx : mcltt. -Inductive rel_sub σ p σ' p' (R : relation env) : Prop := -| mk_rel_sub : forall o o', {{ ⟦ σ ⟧s p ↘ o }} -> {{ ⟦ σ' ⟧s p' ↘ o' }} -> {{ Dom o ≈ o' ∈ R }} -> rel_sub σ p σ' p' R. +Inductive rel_sub σ ρ σ' ρ' (R : relation env) : Prop := +| mk_rel_sub : forall ρσ ρ'σ', {{ ⟦ σ ⟧s ρ ↘ ρσ }} -> {{ ⟦ σ' ⟧s ρ' ↘ ρ'σ' }} -> {{ Dom ρσ ≈ ρ'σ' ∈ R }} -> rel_sub σ ρ σ' ρ' R. #[global] Arguments mk_rel_sub {_ _ _ _ _ _}. #[export] @@ -34,8 +34,8 @@ Hint Constructors rel_sub : mcltt. Definition rel_sub_under_ctx Γ Δ σ σ' := exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}) env_rel' (_ : {{ EF Δ ≈ Δ ∈ per_ctx_env ↘ env_rel' }}), - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), - rel_sub σ p σ' p' env_rel'. + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), + rel_sub σ ρ σ' ρ' env_rel'. Definition valid_sub_under_ctx Γ Δ σ := rel_sub_under_ctx Γ Δ σ σ. #[global] @@ -47,8 +47,8 @@ Hint Unfold valid_sub_under_ctx : mcltt. Definition subtyp_under_ctx Γ M M' := exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}) i, - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), - exists R R', rel_typ i M p M p' R /\ rel_typ i M' p M' p' R' /\ rel_exp M p M' p' (per_subtyp i). + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), + exists R R', rel_typ i M ρ M ρ' R /\ rel_typ i M' ρ M' ρ' R' /\ rel_exp M ρ M' ρ' (per_subtyp i). Notation "⊨ Γ ≈ Γ'" := (per_ctx Γ Γ') (in custom judg at level 80, Γ custom exp, Γ' custom exp). Notation "⊨ Γ" := (valid_ctx Γ) (in custom judg at level 80, Γ custom exp). diff --git a/theories/Core/Completeness/LogicalRelation/Lemmas.v b/theories/Core/Completeness/LogicalRelation/Lemmas.v index 4e3d5bca..d35c2f69 100644 --- a/theories/Core/Completeness/LogicalRelation/Lemmas.v +++ b/theories/Core/Completeness/LogicalRelation/Lemmas.v @@ -3,23 +3,23 @@ From Mcltt Require Import Base LibTactics. From Mcltt.Core.Completeness Require Import LogicalRelation.Definitions LogicalRelation.Tactics. Import Domain_Notations. -Add Parametric Morphism M p M' p' : (rel_exp M p M' p') +Add Parametric Morphism M ρ M' ρ' : (rel_exp M ρ M' ρ') with signature (@relation_equivalence domain) ==> iff as rel_exp_morphism. Proof. intros R R' HRR'. split; intros []; econstructor; intuition. Qed. -Add Parametric Morphism σ p σ' p' : (rel_sub σ p σ' p') +Add Parametric Morphism σ ρ σ' ρ' : (rel_sub σ ρ σ' ρ') with signature (@relation_equivalence env) ==> iff as rel_sub_morphism. Proof. intros R R' HRR'. split; intros []; econstructor; intuition. Qed. -Lemma rel_exp_implies_rel_typ : forall {i A p A' p'}, - rel_exp A p A' p' (per_univ i) -> - exists R, rel_typ i A p A' p' R. +Lemma rel_exp_implies_rel_typ : forall {i A ρ A' ρ'}, + rel_exp A ρ A' ρ' (per_univ i) -> + exists R, rel_typ i A ρ A' ρ' R. Proof. intros. destruct_by_head rel_exp. @@ -30,9 +30,9 @@ Qed. #[export] Hint Resolve rel_exp_implies_rel_typ : mcltt. -Lemma rel_typ_implies_rel_exp : forall {i A p A' p' R}, - rel_typ i A p A' p' R -> - rel_exp A p A' p' (per_univ i). +Lemma rel_typ_implies_rel_exp : forall {i A ρ A' ρ' R}, + rel_typ i A ρ A' ρ' R -> + rel_exp A ρ A' ρ' (per_univ i). Proof. intros. destruct_by_head rel_typ. diff --git a/theories/Core/Completeness/NatCases.v b/theories/Core/Completeness/NatCases.v index ec429a5d..fc8ed9b6 100644 --- a/theories/Core/Completeness/NatCases.v +++ b/theories/Core/Completeness/NatCases.v @@ -8,8 +8,8 @@ Import Domain_Notations. Lemma rel_exp_of_nat_inversion : forall {Γ M M'}, {{ Γ ⊨ M ≈ M' : ℕ }} -> exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}), - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), - rel_exp M p M' p' per_nat. + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), + rel_exp M ρ M' ρ' per_nat. Proof. intros * [env_relΓ]. destruct_conjs. @@ -24,8 +24,8 @@ Qed. Lemma rel_exp_of_nat : forall {Γ M M'}, (exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}), - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), - rel_exp M p M' p' per_nat) -> + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), + rel_exp M ρ M' ρ' per_nat) -> {{ Γ ⊨ M ≈ M' : ℕ }}. Proof. intros * [env_relΓ]. @@ -145,8 +145,8 @@ Hint Resolve rel_exp_succ_cong : mcltt. Lemma rel_exp_of_sub_id_zero_inversion : forall {Γ M M' A}, {{ Γ ⊨ M ≈ M' : A[Id,,zero] }} -> exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}) i, - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), - exists elem_rel, rel_typ i A d{{{ p ↦ zero }}} A d{{{ p' ↦ zero }}} elem_rel /\ rel_exp M p M' p' elem_rel. + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), + exists elem_rel, rel_typ i A d{{{ ρ ↦ zero }}} A d{{{ ρ' ↦ zero }}} elem_rel /\ rel_exp M ρ M' ρ' elem_rel. Proof. intros * [env_relΓ]. destruct_conjs. @@ -160,8 +160,8 @@ Qed. Lemma rel_exp_of_sub_id_zero : forall {Γ M M' A}, (exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}) i, - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), - exists elem_rel, rel_typ i A d{{{ p ↦ zero }}} A d{{{ p' ↦ zero }}} elem_rel /\ rel_exp M p M' p' elem_rel) -> + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), + exists elem_rel, rel_typ i A d{{{ ρ ↦ zero }}} A d{{{ ρ' ↦ zero }}} elem_rel /\ rel_exp M ρ M' ρ' elem_rel) -> {{ Γ ⊨ M ≈ M' : A[Id,,zero] }}. Proof. intros * [env_relΓ]. @@ -182,8 +182,8 @@ Ltac eexists_rel_exp_of_sub_id_zero := Lemma rel_exp_of_sub_wkwk_succ_var1_inversion : forall {Γ M M' A}, {{ Γ, ℕ, A ⊨ M ≈ M' : A[Wk∘Wk,,succ(#1)] }} -> exists env_rel (_ : {{ EF Γ, ℕ, A ≈ Γ, ℕ, A ∈ per_ctx_env ↘ env_rel }}) i, - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), - exists elem_rel, rel_typ i A d{{{ p ↯ ↯ ↦ succ ~(p 1) }}} A d{{{ p' ↯ ↯ ↦ succ ~(p' 1) }}} elem_rel /\ rel_exp M p M' p' elem_rel. + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), + exists elem_rel, rel_typ i A d{{{ ρ ↯ ↯ ↦ succ ~(ρ 1) }}} A d{{{ ρ' ↯ ↯ ↦ succ ~(ρ' 1) }}} elem_rel /\ rel_exp M ρ M' ρ' elem_rel. Proof. intros * [env_relΓℕA]. destruct_conjs. @@ -198,11 +198,11 @@ Qed. Lemma rel_exp_of_sub_id_N : forall {Γ M M' N A}, {{ Γ ⊨ N : ℕ }} -> (exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}) i, - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), exists n n', - {{ ⟦ N ⟧ p ↘ n }} /\ - {{ ⟦ N ⟧ p' ↘ n' }} /\ - exists elem_rel, rel_typ i A d{{{ p ↦ n }}} A d{{{ p' ↦ n' }}} elem_rel /\ rel_exp M p M' p' elem_rel) -> + {{ ⟦ N ⟧ ρ ↘ n }} /\ + {{ ⟦ N ⟧ ρ' ↘ n' }} /\ + exists elem_rel, rel_typ i A d{{{ ρ ↦ n }}} A d{{{ ρ' ↦ n' }}} elem_rel /\ rel_exp M ρ M' ρ' elem_rel) -> {{ Γ ⊨ M ≈ M' : A[Id,,N] }}. Proof. intros * []%rel_exp_of_nat_inversion [env_relΓ]. @@ -230,13 +230,13 @@ Lemma eval_natrec_sub_neut : forall {Γ env_relΓ σ Δ env_relΔ MZ MZ' MS MS' {{ Δ ⊨ MZ ≈ MZ' : A[Id,,zero] }} -> {{ Δ, ℕ, A ⊨ MS ≈ MS' : A[Wk∘Wk,,succ(#1)] }} -> {{ Dom m ≈ m' ∈ per_bot }} -> - (forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_relΓ }}) o o' mz mz', - {{ ⟦ σ ⟧s p ↘ o }} -> - {{ ⟦ σ ⟧s p' ↘ o' }} -> - {{ Dom o ≈ o' ∈ env_relΔ }} -> - {{ ⟦ MZ ⟧ o ↘ mz }} -> - {{ ⟦ MZ' ⟧ o' ↘ mz' }} -> - {{ Dom rec m under o return A | zero -> mz | succ -> MS end ≈ rec m' under p' return A'[q σ] | zero -> mz' | succ -> MS'[q (q σ)] end ∈ per_bot }}). + (forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_relΓ }}) ρσ ρ'σ' mz mz', + {{ ⟦ σ ⟧s ρ ↘ ρσ }} -> + {{ ⟦ σ ⟧s ρ' ↘ ρ'σ' }} -> + {{ Dom ρσ ≈ ρ'σ' ∈ env_relΔ }} -> + {{ ⟦ MZ ⟧ ρσ ↘ mz }} -> + {{ ⟦ MZ' ⟧ ρ'σ' ↘ mz' }} -> + {{ Dom rec m under ρσ return A | zero -> mz | succ -> MS end ≈ rec m' under ρ' return A'[q σ] | zero -> mz' | succ -> MS'[q (q σ)] end ∈ per_bot }}). Proof. intros * equiv_Γ_Γ equiv_Δ_Δ [env_relΔℕ]%rel_exp_of_typ_inversion @@ -258,17 +258,23 @@ Proof. destruct_by_head rel_exp. destruct_conjs. functional_eval_rewrite_clear. + match goal with + | _: {{ ⟦ σ ⟧s ρ ↘ ~?ρ1 }}, + _: {{ ⟦ σ ⟧s ρ' ↘ ~?ρ2 }} |- _ => + rename ρ1 into ρσ; + rename ρ2 into ρ'σ + end. intro s. assert {{ Dom ⇑! ℕ s ≈ ⇑! ℕ s ∈ per_nat }} by mauto. - assert {{ Dom o ↦ ⇑! ℕ s ≈ o' ↦ ⇑! ℕ s ∈ env_relΔℕ }} as HinΔℕs by (apply_relation_equivalence; mauto). + assert {{ Dom ρσ ↦ ⇑! ℕ s ≈ ρ'σ ↦ ⇑! ℕ s ∈ env_relΔℕ }} as HinΔℕs by (apply_relation_equivalence; mauto). (on_all_hyp: fun H => destruct (H _ _ HinΔℕs)). assert {{ Dom succ (⇑! ℕ s) ≈ succ (⇑! ℕ s) ∈ per_nat }} by mauto. - assert {{ Dom o ↦ succ (⇑! ℕ s) ≈ o' ↦ succ (⇑! ℕ s) ∈ env_relΔℕ }} as HinΔℕsuccs by (apply_relation_equivalence; mauto). + assert {{ Dom ρσ ↦ succ (⇑! ℕ s) ≈ ρ'σ ↦ succ (⇑! ℕ s) ∈ env_relΔℕ }} as HinΔℕsuccs by (apply_relation_equivalence; mauto). (on_all_hyp: fun H => destruct (H _ _ HinΔℕsuccs)). - assert {{ Dom o ↦ ⇑! ℕ s ≈ o' ↦ ⇑! ℕ s ∈ env_relΔℕ }} as HinΔℕs' by (apply_relation_equivalence; mauto). - assert {{ Dom o ↦ succ (⇑! ℕ s) ≈ o' ↦ succ (⇑! ℕ s) ∈ env_relΔℕ }} as HinΔℕsuccs' by (apply_relation_equivalence; mauto). + assert {{ Dom ρσ ↦ ⇑! ℕ s ≈ ρ'σ ↦ ⇑! ℕ s ∈ env_relΔℕ }} as HinΔℕs' by (apply_relation_equivalence; mauto). + assert {{ Dom ρσ ↦ succ (⇑! ℕ s) ≈ ρ'σ ↦ succ (⇑! ℕ s) ∈ env_relΔℕ }} as HinΔℕsuccs' by (apply_relation_equivalence; mauto). assert {{ Dom zero ≈ zero ∈ per_nat }} by econstructor. - assert {{ Dom o ↦ zero ≈ o' ↦ zero ∈ env_relΔℕ }} as HinΔℕz by (apply_relation_equivalence; mauto). + assert {{ Dom ρσ ↦ zero ≈ ρ'σ ↦ zero ∈ env_relΔℕ }} as HinΔℕz by (apply_relation_equivalence; mauto). apply_relation_equivalence. (on_all_hyp: fun H => destruct (H _ _ HinΔℕs')). (on_all_hyp: fun H => destruct (H _ _ HinΔℕsuccs')). @@ -280,12 +286,12 @@ Proof. rename m'0 into a'. rename a1 into asucc. rename m'1 into asucc'. - assert {{ Dom o ↦ ⇑! ℕ s ↦ ⇑! a (S s) ≈ o' ↦ ⇑! ℕ s ↦ ⇑! a' (S s) ∈ env_relΔℕA }} as HinΔℕA. + assert {{ Dom ρσ ↦ ⇑! ℕ s ↦ ⇑! a (S s) ≈ ρ'σ ↦ ⇑! ℕ s ↦ ⇑! a' (S s) ∈ env_relΔℕA }} as HinΔℕA. { apply_relation_equivalence; eexists; eauto. unfold drop_env. - repeat change (fun n => d{{{ ~?p ↦ ~?x ↦ ~?y }}} (S n)) with (fun n => d{{{ p ↦ x }}} n). - repeat change (d{{{ ~?p ↦ ~?x ↦ ~?y }}} 0) with y. + repeat change (fun n => d{{{ ~?ρσ ↦ ~?x ↦ ~?y }}} (S n)) with (fun n => d{{{ ρ ↦ x }}} n). + repeat change (d{{{ ~?ρσ ↦ ~?x ↦ ~?y }}} 0) with y. eapply per_bot_then_per_elem; mauto. } apply_relation_equivalence. @@ -310,13 +316,13 @@ Corollary eval_natrec_neut : forall {Γ env_relΓ MZ MZ' MS MS' A A' i m m'}, {{ Γ ⊨ MZ ≈ MZ' : A[Id,,zero] }} -> {{ Γ, ℕ, A ⊨ MS ≈ MS' : A[Wk∘Wk,,succ(#1)] }} -> {{ Dom m ≈ m' ∈ per_bot }} -> - (forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_relΓ }}) mz mz', - {{ ⟦ MZ ⟧ p ↘ mz }} -> - {{ ⟦ MZ' ⟧ p' ↘ mz' }} -> - {{ Dom rec m under p return A | zero -> mz | succ -> MS end ≈ rec m' under p' return A' | zero -> mz' | succ -> MS' end ∈ per_bot }}). + (forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_relΓ }}) mz mz', + {{ ⟦ MZ ⟧ ρ ↘ mz }} -> + {{ ⟦ MZ' ⟧ ρ' ↘ mz' }} -> + {{ Dom rec m under ρ return A | zero -> mz | succ -> MS end ≈ rec m' under ρ' return A' | zero -> mz' | succ -> MS' end ∈ per_bot }}). Proof. intros. - assert {{ Dom rec m under p return A | zero -> mz | succ -> MS end ≈ rec m' under p' return A'[q Id] | zero -> mz' | succ -> MS'[q (q Id)] end ∈ per_bot }} by (mauto using eval_natrec_sub_neut). + assert {{ Dom rec m under ρ return A | zero -> mz | succ -> MS end ≈ rec m' under ρ' return A'[q Id] | zero -> mz' | succ -> MS'[q (q Id)] end ∈ per_bot }} by (mauto using eval_natrec_sub_neut). etransitivity; [eassumption |]. intros s. match_by_head per_bot ltac:(fun H => specialize (H s) as [? []]). @@ -332,12 +338,12 @@ Lemma eval_natrec_rel : forall {Γ env_relΓ MZ MZ' MS MS' A A' i m m'}, {{ Γ ⊨ MZ ≈ MZ' : A[Id,,zero] }} -> {{ Γ, ℕ, A ⊨ MS ≈ MS' : A[Wk∘Wk,,succ(#1)] }} -> {{ Dom m ≈ m' ∈ per_nat }} -> - (forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_relΓ }}), + (forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_relΓ }}), forall elem_rel, - rel_typ i A d{{{ p ↦ m }}} A d{{{ p' ↦ m' }}} elem_rel -> + rel_typ i A d{{{ ρ ↦ m }}} A d{{{ ρ' ↦ m' }}} elem_rel -> exists r r', - {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} /\ - {{ rec m' ⟦return A' | zero -> MZ' | succ -> MS' end⟧ p' ↘ r' }} /\ + {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} /\ + {{ rec m' ⟦return A' | zero -> MZ' | succ -> MS' end⟧ ρ' ↘ r' }} /\ {{ Dom r ≈ r' ∈ elem_rel }}). Proof. intros * equiv_Γ_Γ HA HMZ HMS equiv_m_m'. @@ -367,17 +373,25 @@ Proof. (on_all_hyp_rev: destruct_rel_by_assumption env_relΓ). destruct_by_head rel_typ. invert_rel_typ_body. - assert {{ Dom p ↦ m ≈ p' ↦ m' ∈ env_relΓℕ }} by (apply_relation_equivalence; mauto). + match goal with + | _: env_relΓ ρ ?ρ0 |- _ => + rename ρ0 into ρ' + end. + assert {{ Dom ρ ↦ m ≈ ρ' ↦ m' ∈ env_relΓℕ }} by (apply_relation_equivalence; mauto). (on_all_hyp: destruct_rel_by_assumption env_relΓℕ). - assert {{ Dom p ↦ m ≈ p' ↦ m' ∈ env_relΓℕ }} as HinΓℕ by (apply_relation_equivalence; mauto). + assert {{ Dom ρ ↦ m ≈ ρ' ↦ m' ∈ env_relΓℕ }} as HinΓℕ by (apply_relation_equivalence; mauto). apply_relation_equivalence. (on_all_hyp: fun H => directed destruct (H _ _ HinΓℕ)). destruct_by_head per_univ. - unshelve epose proof (IHequiv_m_m' _ _ equiv_p_p' _ _) as [? [? [? []]]]; shelve_unifiable; [solve [mauto] |]. + unshelve epose proof (IHequiv_m_m' _ _ equiv_ρ_ρ' _ _) as [? [? [? []]]]; shelve_unifiable; [solve [mauto] |]. handle_per_univ_elem_irrel. - rename x5 into r1. - rename x6 into r1'. - assert {{ Dom p ↦ m ↦ r1 ≈ p' ↦ m' ↦ r1' ∈ env_relΓℕA }} as HinΓℕA by (apply_relation_equivalence; mauto). + match goal with + | _: {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ ~?r0 }}, + _: {{ rec m' ⟦return A' | zero -> MZ' | succ -> MS' end⟧ ρ' ↘ ~?r0' }} |- _ => + rename r0 into rm; + rename r0' into rm' + end. + assert {{ Dom ρ ↦ m ↦ rm ≈ ρ' ↦ m' ↦ rm' ∈ env_relΓℕA }} as HinΓℕA by (apply_relation_equivalence; mauto). apply_relation_equivalence. (on_all_hyp: fun H => directed destruct (H _ _ HinΓℕA)). destruct_conjs. @@ -385,15 +399,22 @@ Proof. destruct_by_head rel_exp. handle_per_univ_elem_irrel. do 2 eexists; mauto. - - rename n into m'. + - match goal with + | _: per_bot m ?n |- _ => + rename n into m' + end. handle_per_ctx_env_irrel. match_by_head (per_ctx_env env_relΓℕ) invert_per_ctx_env. handle_per_ctx_env_irrel. (on_all_hyp: destruct_rel_by_assumption env_relΓ). (on_all_hyp_rev: destruct_rel_by_assumption env_relΓ). invert_rel_typ_body. + match goal with + | _: env_relΓ ρ ?ρ0 |- _ => + rename ρ0 into ρ' + end. assert {{ Dom ⇑ ℕ m ≈ ⇑ ℕ m' ∈ per_nat }} by (econstructor; eassumption). - assert {{ Dom p ↦ ⇑ ℕ m ≈ p' ↦ ⇑ ℕ m' ∈ env_relΓℕ }} as HinΓℕ by (apply_relation_equivalence; mauto). + assert {{ Dom ρ ↦ ⇑ ℕ m ≈ ρ' ↦ ⇑ ℕ m' ∈ env_relΓℕ }} as HinΓℕ by (apply_relation_equivalence; mauto). apply_relation_equivalence. (on_all_hyp: fun H => directed destruct (H _ _ HinΓℕ)). destruct_by_head per_univ. @@ -416,11 +437,11 @@ Lemma rel_exp_natrec_cong_rel_typ: forall {Γ A A' i M M' env_relΓ}, {{ DF Γ ≈ Γ ∈ per_ctx_env ↘ env_relΓ }} -> {{ Γ, ℕ ⊨ A ≈ A' : Type@i }} -> {{ Γ ⊨ M ≈ M' : ℕ }} -> - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_relΓ }}) n n', - {{ ⟦ M ⟧ p ↘ n }} -> - {{ ⟦ M' ⟧ p' ↘ n' }} -> + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_relΓ }}) n n', + {{ ⟦ M ⟧ ρ ↘ n }} -> + {{ ⟦ M' ⟧ ρ' ↘ n' }} -> exists elem_rel, - rel_typ i A d{{{ p ↦ n }}} A d{{{ p' ↦ n' }}} elem_rel. + rel_typ i A d{{{ ρ ↦ n }}} A d{{{ ρ' ↦ n' }}} elem_rel. Proof. intros. assert {{ ⊨ Γ }} by (eexists; eauto). @@ -457,9 +478,9 @@ Proof. intros. (on_all_hyp: destruct_rel_by_assumption env_relΓ). functional_eval_rewrite_clear. - assert (exists elem_rel, rel_typ i A d{{{ p ↦ m }}} A d{{{ p' ↦ m' }}} elem_rel) as [elem_rel] + assert (exists elem_rel, rel_typ i A d{{{ ρ ↦ m }}} A d{{{ ρ' ↦ m' }}} elem_rel) as [elem_rel] by mauto using rel_exp_natrec_cong_rel_typ. - assert (exists elem_rel, rel_typ i A d{{{ p ↦ m }}} A d{{{ p' ↦ m'0 }}} elem_rel) as [] + assert (exists elem_rel, rel_typ i A d{{{ ρ ↦ m }}} A d{{{ ρ' ↦ m'0 }}} elem_rel) as [] by mauto using rel_exp_natrec_cong_rel_typ. do 2 eexists. repeat split; [eassumption | eassumption |]. @@ -468,8 +489,8 @@ Proof. destruct_by_head rel_typ. handle_per_univ_elem_irrel. assert (exists r r', - {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} /\ - {{ rec m'0 ⟦return A' | zero -> MZ' | succ -> MS' end⟧ p' ↘ r' }} /\ + {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} /\ + {{ rec m'0 ⟦return A' | zero -> MZ' | succ -> MS' end⟧ ρ' ↘ r' }} /\ {{ Dom r ≈ r' ∈ elem_rel }}) by mauto using eval_natrec_rel. destruct_conjs. @@ -486,16 +507,16 @@ Lemma eval_natrec_sub_rel : forall {Γ env_relΓ σ Δ env_relΔ MZ MZ' MS MS' A {{ Δ ⊨ MZ ≈ MZ' : A[Id,,zero] }} -> {{ Δ, ℕ, A ⊨ MS ≈ MS' : A[Wk∘Wk,,succ(#1)] }} -> {{ Dom m ≈ m' ∈ per_nat }} -> - (forall p p' - (equiv_p_p' : {{ Dom p ≈ p' ∈ env_relΓ }}) + (forall ρ ρ' + (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_relΓ }}) o o' elem_rel, - {{ ⟦ σ ⟧s p ↘ o }} -> - {{ ⟦ σ ⟧s p' ↘ o' }} -> + {{ ⟦ σ ⟧s ρ ↘ o }} -> + {{ ⟦ σ ⟧s ρ' ↘ o' }} -> {{ Dom o ≈ o' ∈ env_relΔ }} -> rel_typ i A d{{{ o ↦ m }}} A d{{{ o' ↦ m' }}} elem_rel -> exists r r', {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ o ↘ r }} /\ - {{ rec m' ⟦return A'[q σ] | zero -> MZ'[σ] | succ -> MS'[q (q σ)] end⟧ p' ↘ r' }} /\ + {{ rec m' ⟦return A'[q σ] | zero -> MZ'[σ] | succ -> MS'[q (q σ)] end⟧ ρ' ↘ r' }} /\ {{ Dom r ≈ r' ∈ elem_rel }}). Proof. intros * equiv_Γ_Γ equiv_Δ_Δ HA HMZ HMS equiv_m_m'. @@ -512,7 +533,10 @@ Proof. destruct_by_head rel_exp. handle_per_univ_elem_irrel. do 2 eexists; repeat split; only 1-2: econstructor; mauto. - - rename n into m'. + - match goal with + | _: per_nat m ?n |- _ => + rename n into m' + end. assert {{ Δ, ℕ ⊨ A ≈ A : Type@i }} as []%rel_exp_of_typ_inversion by (etransitivity; mauto). apply rel_exp_of_sub_wkwk_succ_var1_inversion in HMS as [env_relΔℕA]. destruct_conjs. @@ -525,17 +549,27 @@ Proof. (on_all_hyp_rev: destruct_rel_by_assumption env_relΔ). destruct_by_head rel_typ. invert_rel_typ_body. - assert {{ Dom o ↦ m ≈ o' ↦ m' ∈ env_relΔℕ }} by (apply_relation_equivalence; mauto). + match goal with + | _: {{ ⟦ σ ⟧s ρ ↘ ~?ρ1 }}, + _: {{ ⟦ σ ⟧s ρ' ↘ ~?ρ2 }} |- _ => + rename ρ1 into ρσ; + rename ρ2 into ρ'σ + end. + assert {{ Dom ρσ ↦ m ≈ ρ'σ ↦ m' ∈ env_relΔℕ }} by (apply_relation_equivalence; mauto). (on_all_hyp: destruct_rel_by_assumption env_relΔℕ). - assert {{ Dom o ↦ m ≈ o' ↦ m' ∈ env_relΔℕ }} as HinΔℕ by (apply_relation_equivalence; mauto). + assert {{ Dom ρσ ↦ m ≈ ρ'σ ↦ m' ∈ env_relΔℕ }} as HinΔℕ by (apply_relation_equivalence; mauto). apply_relation_equivalence. (on_all_hyp: fun H => directed destruct (H _ _ HinΔℕ)). destruct_by_head per_univ. - unshelve epose proof (IHequiv_m_m' _ _ equiv_p_p' _ _) as [? [? [? []]]]; shelve_unifiable; only 4: solve [mauto]; eauto. + unshelve epose proof (IHequiv_m_m' _ _ equiv_ρ_ρ' _ _) as [? [? [? []]]]; shelve_unifiable; only 4: solve [mauto]; eauto. handle_per_univ_elem_irrel. - rename x5 into r1. - rename x6 into r1'. - assert {{ Dom o ↦ m ↦ r1 ≈ o' ↦ m' ↦ r1' ∈ env_relΔℕA }} as HinΔℕA by (apply_relation_equivalence; mauto). + match goal with + | _: {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ~_ ↘ ~?r0 }}, + _: {{ rec m' ⟦return A'[q σ] | zero -> MZ'[σ] | succ -> MS'[q (q σ)] end⟧ ~_ ↘ ~?r0' }} |- _ => + rename r0 into rm; + rename r0' into rm' + end. + assert {{ Dom ρσ ↦ m ↦ rm ≈ ρ'σ ↦ m' ↦ rm' ∈ env_relΔℕA }} as HinΔℕA by (apply_relation_equivalence; mauto). apply_relation_equivalence. (on_all_hyp: fun H => directed destruct (H _ _ HinΔℕA)). destruct_conjs. @@ -543,15 +577,24 @@ Proof. destruct_by_head rel_exp. handle_per_univ_elem_irrel. do 2 eexists; repeat split; only 1-2: repeat econstructor; mauto. - - rename n into m'. + - match goal with + | _: per_bot m ?n |- _ => + rename n into m' + end. handle_per_ctx_env_irrel. match_by_head (per_ctx_env env_relΔℕ) invert_per_ctx_env. handle_per_ctx_env_irrel. (on_all_hyp: destruct_rel_by_assumption env_relΔ). (on_all_hyp_rev: destruct_rel_by_assumption env_relΔ). invert_rel_typ_body. + match goal with + | _: {{ ⟦ σ ⟧s ρ ↘ ~?ρ1 }}, + _: {{ ⟦ σ ⟧s ρ' ↘ ~?ρ2 }} |- _ => + rename ρ1 into ρσ; + rename ρ2 into ρ'σ + end. assert {{ Dom ⇑ ℕ m ≈ ⇑ ℕ m' ∈ per_nat }} by (econstructor; eassumption). - assert {{ Dom o ↦ ⇑ ℕ m ≈ o' ↦ ⇑ ℕ m' ∈ env_relΔℕ }} as HinΔℕ by (apply_relation_equivalence; mauto). + assert {{ Dom ρσ ↦ ⇑ ℕ m ≈ ρ'σ ↦ ⇑ ℕ m' ∈ env_relΔℕ }} as HinΔℕ by (apply_relation_equivalence; mauto). apply_relation_equivalence. (on_all_hyp: fun H => directed destruct (H _ _ HinΔℕ)). destruct_by_head per_univ. @@ -575,9 +618,9 @@ Lemma rel_exp_natrec_sub_rel_typ: forall {Γ σ Δ A i M env_relΓ}, {{ Γ ⊨s σ : Δ }} -> {{ Δ, ℕ ⊨ A : Type@i }} -> {{ Δ ⊨ M : ℕ }} -> - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_relΓ }}), + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_relΓ }}), exists elem_rel, - rel_typ i {{{ A[σ,,M[σ]] }}} p {{{ A[σ,,M[σ]] }}} p' elem_rel. + rel_typ i {{{ A[σ,,M[σ]] }}} ρ {{{ A[σ,,M[σ]] }}} ρ' elem_rel. Proof. intros. assert {{ ⊨ Γ }} by (eexists; eauto). @@ -605,7 +648,7 @@ Proof. handle_per_ctx_env_irrel. eexists_rel_exp. intros. - assert (exists elem_rel, rel_typ i {{{ A[σ,,M[σ]] }}} p {{{ A[σ,,M[σ]] }}} p' elem_rel) as [elem_rel] + assert (exists elem_rel, rel_typ i {{{ A[σ,,M[σ]] }}} ρ {{{ A[σ,,M[σ]] }}} ρ' elem_rel) as [elem_rel] by (eapply rel_exp_natrec_sub_rel_typ; only 5: eassumption; mauto; eexists; mauto). eexists. split; [eassumption |]. @@ -613,9 +656,20 @@ Proof. (on_all_hyp: destruct_rel_by_assumption env_relΔ). destruct_by_head rel_typ. invert_rel_typ_body. + match goal with + | _: {{ ⟦ σ ⟧s ~?ρ0 ↘ ~?ρσ0 }}, + _: {{ ⟦ A ⟧ ρσ ↦ ~?m0 ↘ ~?a0 }}, + _: {{ ⟦ A ⟧ ~?ρσ0 ↦ ~?m0' ↘ ~?a0' }} |- _ => + rename ρ0 into ρ'; + rename ρσ0 into ρ'σ; + rename a0 into a; + rename m0 into m; + rename a0' into a'; + rename m0' into m' + end. enough (exists r r', - {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ o ↘ r }} /\ - {{ rec m' ⟦return A[q σ] | zero -> MZ[σ] | succ -> MS[q (q σ)] end⟧ p' ↘ r' }} /\ + {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ρσ ↘ r }} /\ + {{ rec m' ⟦return A[q σ] | zero -> MZ[σ] | succ -> MS[q (q σ)] end⟧ ρ' ↘ r' }} /\ {{ Dom r ≈ r' ∈ elem_rel }}) by (destruct_conjs; econstructor; mauto). mauto 4 using eval_natrec_sub_rel. @@ -644,8 +698,12 @@ Proof. destruct_by_head rel_typ. invert_rel_typ_body. destruct_by_head rel_exp. + match goal with + | _: env_relΓ ρ ?ρ0 |- _ => + rename ρ0 into ρ' + end. assert {{ Dom zero ≈ zero ∈ per_nat }} by econstructor. - assert {{ Dom p ↦ zero ≈ p' ↦ zero ∈ env_relΓℕ }} by (apply_relation_equivalence; eauto). + assert {{ Dom ρ ↦ zero ≈ ρ' ↦ zero ∈ env_relΓℕ }} by (apply_relation_equivalence; eauto). (on_all_hyp: destruct_rel_by_assumption env_relΓℕ). handle_per_univ_elem_irrel. eexists. @@ -659,9 +717,9 @@ Lemma rel_exp_nat_beta_succ_rel_typ : forall {Γ env_relΓ A i M}, {{ DF Γ ≈ Γ ∈ per_ctx_env ↘ env_relΓ }} -> {{ Γ, ℕ ⊨ A : Type@i }} -> {{ Γ ⊨ M : ℕ }} -> - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_relΓ }}), + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_relΓ }}), exists elem_rel, - rel_typ i {{{ A[Id,,succ M] }}} p {{{ A[Id,,succ M] }}} p' elem_rel. + rel_typ i {{{ A[Id,,succ M] }}} ρ {{{ A[Id,,succ M] }}} ρ' elem_rel. Proof. intros. assert {{ ⊨ Γ }} by (eexists; eauto). @@ -690,16 +748,18 @@ Proof. intros. (on_all_hyp: destruct_rel_by_assumption env_relΓ). assert (exists elem_rel, - rel_typ i {{{ A[Id,,succ M] }}} p {{{ A[Id,,succ M] }}} p' elem_rel) as [elem_rel] + rel_typ i {{{ A[Id,,succ M] }}} ρ {{{ A[Id,,succ M] }}} ρ' elem_rel) as [elem_rel] by (eapply rel_exp_nat_beta_succ_rel_typ; mauto). eexists; split; [eassumption |]. destruct_by_head rel_typ. invert_rel_typ_body. - rename p'2 into p. - rename p'1 into p'. + match goal with + | _: env_relΓ ρ ?ρ0 |- _ => + rename ρ0 into ρ' + end. assert (exists r r', - {{ rec succ m ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} /\ - {{ rec succ m' ⟦return A | zero -> MZ | succ -> MS end⟧ p' ↘ r' }} /\ + {{ rec succ m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} /\ + {{ rec succ m' ⟦return A | zero -> MZ | succ -> MS end⟧ ρ' ↘ r' }} /\ {{ Dom r ≈ r' ∈ elem_rel }}) by (eapply eval_natrec_rel; mauto). destruct_conjs. diff --git a/theories/Core/Completeness/SubstitutionCases.v b/theories/Core/Completeness/SubstitutionCases.v index 46d1ea2b..8dfcb085 100644 --- a/theories/Core/Completeness/SubstitutionCases.v +++ b/theories/Core/Completeness/SubstitutionCases.v @@ -205,7 +205,7 @@ Proof with mautosolve. destruct_conjs. eexists_rel_sub. intros. - assert (env_relΓ p' p) by (symmetry; eassumption). + assert (env_relΓ ρ' ρ) by (symmetry; eassumption). (on_all_hyp: destruct_rel_by_assumption env_relΓ). econstructor; mauto; symmetry... Qed. @@ -224,7 +224,7 @@ Proof with mautosolve. handle_per_ctx_env_irrel. eexists_rel_sub. intros. - assert (env_relΓ p' p') by (etransitivity; [symmetry |]; eassumption). + assert (env_relΓ ρ' ρ') by (etransitivity; [symmetry |]; eassumption). (on_all_hyp: destruct_rel_by_assumption env_relΓ). functional_eval_rewrite_clear. econstructor; mauto; etransitivity... diff --git a/theories/Core/Completeness/SubtypingCases.v b/theories/Core/Completeness/SubtypingCases.v index 1ae31917..df63a3d4 100644 --- a/theories/Core/Completeness/SubtypingCases.v +++ b/theories/Core/Completeness/SubtypingCases.v @@ -89,38 +89,43 @@ Proof. destruct_by_head rel_exp. destruct_conjs. handle_per_univ_elem_irrel. + match goal with + | _: env_relΓ ρ ?ρ0 |- _ => + assert_fails (unify ρ ρ0); + rename ρ0 into ρ' + end. rename x0 into head_rel. - assert (forall c c', head_rel p p' equiv_p_p' c c' -> env_relΓA' d{{{ p ↦ c }}} d{{{ p' ↦ c' }}}) as HΓA' + assert (forall c c', head_rel ρ ρ' equiv_ρ_ρ' c c' -> env_relΓA' d{{{ ρ ↦ c }}} d{{{ ρ' ↦ c' }}}) as HΓA' by (intros; apply_relation_equivalence; unshelve eexists; eassumption). (* The proofs for the next two assertions are basically the same *) exvar (relation domain) - ltac:(fun R => assert ({{ DF Π m0 p B ≈ Π m1 p' B ∈ per_univ_elem (Nat.max i k) ↘ R }})). + ltac:(fun R => assert ({{ DF Π m0 ρ B ≈ Π m1 ρ' B ∈ per_univ_elem (Nat.max i k) ↘ R }})). { intros. per_univ_elem_econstructor; [| | solve_refl]. - etransitivity; [| symmetry]; mauto using per_univ_elem_cumu_max_left. - eapply rel_exp_pi_core; [| reflexivity]. intros. - assert (env_relΓA' d{{{ p ↦ c }}} d{{{ p' ↦ c' }}}) as equiv_pc_p'c' by (apply HΓA'; intuition). + assert (env_relΓA' d{{{ ρ ↦ c }}} d{{{ ρ' ↦ c' }}}) as equiv_ρc_ρ'c' by (apply HΓA'; intuition). apply_relation_equivalence. - (on_all_hyp: fun H => destruct (H _ _ equiv_pc_p'c')). + (on_all_hyp: fun H => destruct (H _ _ equiv_ρc_ρ'c')). destruct_conjs. destruct_by_head rel_typ. econstructor; mauto using per_univ_elem_cumu_max_right. } exvar (relation domain) - ltac:(fun R => assert ({{ DF Π a0 p B' ≈ Π a p' B' ∈ per_univ_elem (Nat.max i k) ↘ R }})). + ltac:(fun R => assert ({{ DF Π a0 ρ B' ≈ Π a ρ' B' ∈ per_univ_elem (Nat.max i k) ↘ R }})). { intros. per_univ_elem_econstructor; [| | solve_refl]. - etransitivity; [symmetry |]; mauto using per_univ_elem_cumu_max_left. - eapply rel_exp_pi_core; [| reflexivity]. intros. - assert (env_relΓA' d{{{ p ↦ c }}} d{{{ p' ↦ c' }}}) as equiv_pc_p'c' by (apply HΓA'; intuition). + assert (env_relΓA' d{{{ ρ ↦ c }}} d{{{ ρ' ↦ c' }}}) as equiv_ρc_ρ'c' by (apply HΓA'; intuition). apply_relation_equivalence. - (on_all_hyp: fun H => destruct (H _ _ equiv_pc_p'c')). + (on_all_hyp: fun H => destruct (H _ _ equiv_ρc_ρ'c')). destruct_conjs. destruct_by_head rel_typ. econstructor; mauto using per_univ_elem_cumu_max_right. @@ -131,9 +136,9 @@ Proof. econstructor; only 3-4: try (saturate_refl; mautosolve 2). - eauto using per_univ_elem_cumu_max_left. - intros. - assert (env_relΓA' d{{{ p ↦ c }}} d{{{ p' ↦ c' }}}) as equiv_pc_p'c' by (apply HΓA'; intuition). + assert (env_relΓA' d{{{ ρ ↦ c }}} d{{{ ρ' ↦ c' }}}) as equiv_ρc_ρ'c' by (apply HΓA'; intuition). apply_relation_equivalence. - (on_all_hyp: fun H => destruct (H _ _ equiv_pc_p'c')). + (on_all_hyp: fun H => destruct (H _ _ equiv_ρc_ρ'c')). destruct_conjs. destruct_by_head rel_exp. simplify_evals. diff --git a/theories/Core/Completeness/TermStructureCases.v b/theories/Core/Completeness/TermStructureCases.v index 3d6f5ba6..1b4ab845 100644 --- a/theories/Core/Completeness/TermStructureCases.v +++ b/theories/Core/Completeness/TermStructureCases.v @@ -15,11 +15,17 @@ Proof with mautosolve. handle_per_ctx_env_irrel. eexists_rel_exp. intros. - assert (env_relΓ p' p) by (symmetry; eassumption). - assert (env_relΓ p p) by (etransitivity; eassumption). + assert (env_relΓ ρ' ρ) by (symmetry; eassumption). + assert (env_relΓ ρ ρ) by (etransitivity; eassumption). (on_all_hyp: destruct_rel_by_assumption env_relΓ). handle_per_univ_elem_irrel. - assert (env_relΔ o o0) by (etransitivity; [|symmetry; eassumption]; eassumption). + match goal with + | _: {{ ⟦ σ ⟧s ρ ↘ ~?ρ0 }}, + _: {{ ⟦ σ ⟧s ρ' ↘ ~?ρ'0 }} |- _ => + rename ρ0 into ρσ; + rename ρ'0 into ρ'σ + end. + assert (env_relΔ ρσ ρ'σ) by (etransitivity; [|symmetry; eassumption]; eassumption). (on_all_hyp: destruct_rel_by_assumption env_relΔ). destruct_by_head rel_typ. destruct_by_head rel_exp. @@ -62,8 +68,8 @@ Proof with mautosolve. handle_per_ctx_env_irrel. eexists_rel_exp. intros. - assert (env_relΓ p' p) by (symmetry; eassumption). - assert (env_relΓ p p) by (etransitivity; eassumption). + assert (env_relΓ ρ' ρ) by (symmetry; eassumption). + assert (env_relΓ ρ ρ) by (etransitivity; eassumption). (on_all_hyp: destruct_rel_by_assumption env_relΓ). handle_per_univ_elem_irrel. (on_all_hyp: destruct_rel_by_assumption env_relΓ'). @@ -90,7 +96,7 @@ Proof with mautosolve. handle_per_ctx_env_irrel. eexists_rel_exp. intros. - assert (env_relΓ p p) by (etransitivity; [| symmetry]; eassumption). + assert (env_relΓ ρ ρ) by (etransitivity; [| symmetry]; eassumption). (on_all_hyp: destruct_rel_by_assumption env_relΓ). destruct_by_head per_univ. destruct_by_head rel_typ. @@ -112,7 +118,7 @@ Proof with mautosolve. destruct_conjs. eexists_rel_exp. intros. - assert (env_relΓ p' p) by (symmetry; eauto). + assert (env_relΓ ρ' ρ) by (symmetry; eauto). (on_all_hyp: destruct_rel_by_assumption env_relΓ); destruct_conjs. destruct_by_head rel_typ. destruct_by_head rel_exp. @@ -136,8 +142,8 @@ Proof with mautosolve. handle_per_ctx_env_irrel. eexists_rel_exp. intros. - assert (env_relΓ p' p) by (symmetry; eauto). - assert (env_relΓ p' p') by (etransitivity; eauto). + assert (env_relΓ ρ' ρ) by (symmetry; eauto). + assert (env_relΓ ρ' ρ') by (etransitivity; eauto). (on_all_hyp: destruct_rel_by_assumption env_relΓ); destruct_conjs. destruct_by_head rel_typ. destruct_by_head rel_exp. diff --git a/theories/Core/Completeness/UniverseCases.v b/theories/Core/Completeness/UniverseCases.v index e0662c43..5835e22a 100644 --- a/theories/Core/Completeness/UniverseCases.v +++ b/theories/Core/Completeness/UniverseCases.v @@ -7,8 +7,8 @@ Import Domain_Notations. Lemma rel_exp_of_typ_inversion : forall {Γ A A' i}, {{ Γ ⊨ A ≈ A' : Type@i }} -> exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}), - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), - rel_exp A p A' p' (per_univ i). + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), + rel_exp A ρ A' ρ' (per_univ i). Proof. intros * [env_relΓ]. destruct_conjs. @@ -23,8 +23,8 @@ Qed. Lemma rel_exp_of_typ : forall {Γ A A' i}, (exists env_rel (_ : {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}), - forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_rel }}), - rel_exp A p A' p' (per_univ i)) -> + forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_rel }}), + rel_exp A ρ A' ρ' (per_univ i)) -> {{ Γ ⊨ A ≈ A' : Type@i }}. Proof. intros * [env_relΓ]. diff --git a/theories/Core/Semantic/Domain.v b/theories/Core/Semantic/Domain.v index 4d2e19d7..5ee3ac7f 100644 --- a/theories/Core/Semantic/Domain.v +++ b/theories/Core/Semantic/Domain.v @@ -28,16 +28,16 @@ Derive NoConfusion for domain domain_ne domain_nf. Definition empty_env : env := fun x => d_zero. Arguments empty_env _ /. -Definition extend_env (p : env) (d : domain) : env := +Definition extend_env (ρ : env) (d : domain) : env := fun n => match n with | 0 => d - | S n' => p n' + | S n' => ρ n' end. Arguments extend_env _ _ _ /. Transparent extend_env. -Definition drop_env (p : env) : env := fun n => p (S n). +Definition drop_env (ρ : env) : env := fun n => ρ (S n). Arguments drop_env _ _ /. Transparent drop_env. @@ -51,27 +51,27 @@ Module Domain_Notations. Notation "( x )" := x (in custom domain at level 0, x custom domain at level 60) : mcltt_scope. Notation "~ x" := x (in custom domain at level 0, x constr at level 0) : mcltt_scope. Notation "x" := x (in custom domain at level 0, x global) : mcltt_scope. - Notation "'λ' p M" := (d_fn p M) (in custom domain at level 0, p custom domain at level 30, M custom exp at level 30) : mcltt_scope. + Notation "'λ' ρ M" := (d_fn ρ M) (in custom domain at level 0, ρ custom domain at level 30, M custom exp at level 30) : mcltt_scope. Notation "f x .. y" := (d_app .. (d_app f x) .. y) (in custom domain at level 40, f custom domain, x custom domain at next level, y custom domain at next level) : mcltt_scope. Notation "'ℕ'" := d_nat (in custom domain) : mcltt_scope. Notation "'𝕌' @ n" := (d_univ n) (in custom domain at level 0, n constr at level 0) : mcltt_scope. - Notation "'Π' a p B" := (d_pi a p B) (in custom domain at level 0, a custom domain at level 30, p custom domain at level 0, B custom exp at level 30) : mcltt_scope. + Notation "'Π' a ρ B" := (d_pi a ρ B) (in custom domain at level 0, a custom domain at level 30, ρ custom domain at level 0, B custom exp at level 30) : mcltt_scope. Notation "'zero'" := d_zero (in custom domain at level 0) : mcltt_scope. Notation "'succ' m" := (d_succ m) (in custom domain at level 30, m custom domain at level 30) : mcltt_scope. - Notation "'rec' m 'under' p 'return' P | 'zero' -> mz | 'succ' -> MS 'end'" := (d_natrec p P mz MS m) (in custom domain at level 0, P custom exp at level 60, mz custom domain at level 60, MS custom exp at level 60, p custom domain at level 60, m custom domain at level 60) : mcltt_scope. + Notation "'rec' m 'under' ρ 'return' P | 'zero' -> mz | 'succ' -> MS 'end'" := (d_natrec ρ P mz MS m) (in custom domain at level 0, P custom exp at level 60, mz custom domain at level 60, MS custom exp at level 60, ρ custom domain at level 60, m custom domain at level 60) : mcltt_scope. Notation "'!' n" := (d_var n) (in custom domain at level 0, n constr at level 0) : mcltt_scope. Notation "'⇑' a m" := (d_neut a m) (in custom domain at level 0, a custom domain at level 30, m custom domain at level 30) : mcltt_scope. Notation "'⇓' a m" := (d_dom a m) (in custom domain at level 0, a custom domain at level 30, m custom domain at level 30) : mcltt_scope. Notation "'⇑!' a n" := (d_neut a (d_var n)) (in custom domain at level 0, a custom domain at level 30, n constr at level 0) : mcltt_scope. - Notation "p ↦ m" := (extend_env p m) (in custom domain at level 20, left associativity) : mcltt_scope. - Notation "p '↯'" := (drop_env p) (in custom domain at level 10, p custom domain) : mcltt_scope. + Notation "ρ ↦ m" := (extend_env ρ m) (in custom domain at level 20, left associativity) : mcltt_scope. + Notation "ρ '↯'" := (drop_env ρ) (in custom domain at level 10, ρ custom domain) : mcltt_scope. End Domain_Notations. Import Domain_Notations. -Proposition drop_env_extend_env_cancel : forall p a, - d{{{ (p ↦ a) ↯ }}} = p. +Proposition drop_env_extend_env_cancel : forall ρ a, + d{{{ (ρ ↦ a) ↯ }}} = ρ. Proof. reflexivity. Qed. diff --git a/theories/Core/Semantic/Evaluation/Definitions.v b/theories/Core/Semantic/Evaluation/Definitions.v index 79d3cd22..87779978 100644 --- a/theories/Core/Semantic/Evaluation/Definitions.v +++ b/theories/Core/Semantic/Evaluation/Definitions.v @@ -2,85 +2,90 @@ From Mcltt Require Import Base. From Mcltt Require Export Domain. Import Domain_Notations. -Reserved Notation "'⟦' M '⟧' p '↘' r" (in custom judg at level 80, M custom exp at level 99, p custom domain at level 99, r custom domain at level 99). -Reserved Notation "'rec' m '⟦return' A | 'zero' -> MZ | 'succ' -> MS 'end⟧' p '↘' r" (in custom judg at level 80, m custom domain at level 99, A custom exp at level 99, MZ custom exp at level 99, MS custom exp at level 99, p custom domain at level 99, r custom domain at level 99). +Reserved Notation "'⟦' M '⟧' ρ '↘' r" (in custom judg at level 80, M custom exp at level 99, ρ custom domain at level 99, r custom domain at level 99). +Reserved Notation "'rec' m '⟦return' A | 'zero' -> MZ | 'succ' -> MS 'end⟧' ρ '↘' r" (in custom judg at level 80, m custom domain at level 99, A custom exp at level 99, MZ custom exp at level 99, MS custom exp at level 99, ρ custom domain at level 99, r custom domain at level 99). Reserved Notation "'$|' m '&' n '|↘' r" (in custom judg at level 80, m custom domain at level 99, n custom domain at level 99, r custom domain at level 99). -Reserved Notation "'⟦' σ '⟧s' p '↘' p'" (in custom judg at level 80, σ custom exp at level 99, p custom domain at level 99, p' custom domain at level 99). +Reserved Notation "'⟦' σ '⟧s' ρ '↘' ρσ" (in custom judg at level 80, σ custom exp at level 99, ρ custom domain at level 99, ρσ custom domain at level 99). Generalizable All Variables. Inductive eval_exp : exp -> env -> domain -> Prop := | eval_exp_typ : - `( {{ ⟦ Type@i ⟧ p ↘ 𝕌@i }} ) + `( {{ ⟦ Type@i ⟧ ρ ↘ 𝕌@i }} ) | eval_exp_var : - `( {{ ⟦ # x ⟧ p ↘ ~(p x) }} ) + `( {{ ⟦ # x ⟧ ρ ↘ ~(ρ x) }} ) | eval_exp_nat : - `( {{ ⟦ ℕ ⟧ p ↘ ℕ }} ) + `( {{ ⟦ ℕ ⟧ ρ ↘ ℕ }} ) | eval_exp_zero : - `( {{ ⟦ zero ⟧ p ↘ zero }} ) + `( {{ ⟦ zero ⟧ ρ ↘ zero }} ) | eval_exp_succ : - `( {{ ⟦ M ⟧ p ↘ m }} -> - {{ ⟦ succ M ⟧ p ↘ succ m }} ) + `( {{ ⟦ M ⟧ ρ ↘ m }} -> + {{ ⟦ succ M ⟧ ρ ↘ succ m }} ) | eval_exp_natrec : - `( {{ ⟦ M ⟧ p ↘ m }} -> - {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} -> - {{ ⟦ rec M return A | zero -> MZ | succ -> MS end ⟧ p ↘ r }} ) + `( {{ ⟦ M ⟧ ρ ↘ m }} -> + {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} -> + {{ ⟦ rec M return A | zero -> MZ | succ -> MS end ⟧ ρ ↘ r }} ) | eval_exp_pi : - `( {{ ⟦ A ⟧ p ↘ a }} -> - {{ ⟦ Π A B ⟧ p ↘ Π a p B }} ) + `( {{ ⟦ A ⟧ ρ ↘ a }} -> + {{ ⟦ Π A B ⟧ ρ ↘ Π a ρ B }} ) | eval_exp_fn : - `( {{ ⟦ λ A M ⟧ p ↘ λ p M }} ) + `( {{ ⟦ λ A M ⟧ ρ ↘ λ ρ M }} ) | eval_exp_app : - `( {{ ⟦ M ⟧ p ↘ m }} -> - {{ ⟦ N ⟧ p ↘ n }} -> + `( {{ ⟦ M ⟧ ρ ↘ m }} -> + {{ ⟦ N ⟧ ρ ↘ n }} -> {{ $| m & n |↘ r }} -> - {{ ⟦ M N ⟧ p ↘ r }} ) + {{ ⟦ M N ⟧ ρ ↘ r }} ) | eval_exp_sub : - `( {{ ⟦ σ ⟧s p ↘ p' }} -> - {{ ⟦ M ⟧ p' ↘ m }} -> - {{ ⟦ M[σ] ⟧ p ↘ m }} ) -where "'⟦' e '⟧' p '↘' r" := (eval_exp e p r) (in custom judg) + `( {{ ⟦ σ ⟧s ρ ↘ ρ' }} -> + {{ ⟦ M ⟧ ρ' ↘ m }} -> + {{ ⟦ M[σ] ⟧ ρ ↘ m }} ) +where "'⟦' e '⟧' ρ '↘' r" := (eval_exp e ρ r) (in custom judg) with eval_natrec : exp -> exp -> exp -> domain -> env -> domain -> Prop := | eval_natrec_zero : - `( {{ ⟦ MZ ⟧ p ↘ mz }} -> - {{ rec zero ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ mz }} ) + `( {{ ⟦ MZ ⟧ ρ ↘ mz }} -> + {{ rec zero ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ mz }} ) | eval_natrec_succ : - `( {{ rec b ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} -> - {{ ⟦ MS ⟧ p ↦ b ↦ r ↘ ms }} -> - {{ rec succ b ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ ms }} ) + `( {{ rec b ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} -> + {{ ⟦ MS ⟧ ρ ↦ b ↦ r ↘ ms }} -> + {{ rec succ b ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ ms }} ) | eval_natrec_neut : - `( {{ ⟦ MZ ⟧ p ↘ mz }} -> - {{ ⟦ A ⟧ p ↦ ⇑ ℕ m ↘ a }} -> - {{ rec ⇑ ℕ m ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ ⇑ a (rec m under p return A | zero -> mz | succ -> MS end) }} ) -where "'rec' m '⟦return' A | 'zero' -> MZ | 'succ' -> MS 'end⟧' p '↘' r" := (eval_natrec A MZ MS m p r) (in custom judg) + `( {{ ⟦ MZ ⟧ ρ ↘ mz }} -> + {{ ⟦ A ⟧ ρ ↦ ⇑ ℕ m ↘ a }} -> + {{ rec ⇑ ℕ m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ ⇑ a (rec m under ρ return A | zero -> mz | succ -> MS end) }} ) +where "'rec' m '⟦return' A | 'zero' -> MZ | 'succ' -> MS 'end⟧' ρ '↘' r" := (eval_natrec A MZ MS m ρ r) (in custom judg) with eval_app : domain -> domain -> domain -> Prop := | eval_app_fn : - `( {{ ⟦ M ⟧ p ↦ n ↘ m }} -> - {{ $| λ p M & n |↘ m }} ) + `( {{ ⟦ M ⟧ ρ ↦ n ↘ m }} -> + {{ $| λ ρ M & n |↘ m }} ) | eval_app_neut : - `( {{ ⟦ B ⟧ p ↦ n ↘ b }} -> - {{ $| ⇑ (Π a p B) m & n |↘ ⇑ b (m (⇓ a n)) }} ) + `( {{ ⟦ B ⟧ ρ ↦ n ↘ b }} -> + {{ $| ⇑ (Π a ρ B) m & n |↘ ⇑ b (m (⇓ a n)) }} ) where "'$|' m '&' n '|↘' r" := (eval_app m n r) (in custom judg) with eval_sub : sub -> env -> env -> Prop := | eval_sub_id : - `( {{ ⟦ Id ⟧s p ↘ p }} ) + `( {{ ⟦ Id ⟧s ρ ↘ ρ }} ) | eval_sub_weaken : - `( {{ ⟦ Wk ⟧s p ↘ p↯ }} ) + `( {{ ⟦ Wk ⟧s ρ ↘ ρ ↯ }} ) | eval_sub_extend : - `( {{ ⟦ σ ⟧s p ↘ p' }} -> - {{ ⟦ M ⟧ p ↘ m }} -> - {{ ⟦ σ ,, M ⟧s p ↘ p' ↦ m }} ) + `( {{ ⟦ σ ⟧s ρ ↘ ρσ }} -> + {{ ⟦ M ⟧ ρ ↘ m }} -> + {{ ⟦ σ ,, M ⟧s ρ ↘ ρσ ↦ m }} ) | eval_sub_compose : - `( {{ ⟦ τ ⟧s p ↘ p' }} -> - {{ ⟦ σ ⟧s p' ↘ p'' }} -> - {{ ⟦ σ ∘ τ ⟧s p ↘ p'' }} ) -where "'⟦' σ '⟧s' p '↘' p'" := (eval_sub σ p p') (in custom judg) + `( {{ ⟦ τ ⟧s ρ ↘ ρτ }} -> + {{ ⟦ σ ⟧s ρτ ↘ ρτσ }} -> + {{ ⟦ σ ∘ τ ⟧s ρ ↘ ρτσ }} ) +where "'⟦' σ '⟧s' ρ '↘' ρσ" := (eval_sub σ ρ ρσ) (in custom judg) . Scheme eval_exp_mut_ind := Induction for eval_exp Sort Prop with eval_natrec_mut_ind := Induction for eval_natrec Sort Prop with eval_app_mut_ind := Induction for eval_app Sort Prop with eval_sub_mut_ind := Induction for eval_sub Sort Prop. +Combined Scheme eval_mut_ind from + eval_exp_mut_ind, + eval_natrec_mut_ind, + eval_app_mut_ind, + eval_sub_mut_ind. #[export] Hint Constructors eval_exp eval_natrec eval_app eval_sub : mcltt. diff --git a/theories/Core/Semantic/Evaluation/Lemmas.v b/theories/Core/Semantic/Evaluation/Lemmas.v index a9586819..c7490d66 100644 --- a/theories/Core/Semantic/Evaluation/Lemmas.v +++ b/theories/Core/Semantic/Evaluation/Lemmas.v @@ -3,58 +3,62 @@ From Mcltt Require Import Base LibTactics Evaluation.Definitions. Import Domain_Notations. Section functional_eval. - Let functional_eval_exp_prop M p m1 (_ : {{ ⟦ M ⟧ p ↘ m1 }}) : Prop := forall m2 (Heval2: {{ ⟦ M ⟧ p ↘ m2 }}), m1 = m2. - Let functional_eval_natrec_prop A MZ MS m p r1 (_ : {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r1 }}) : Prop := forall r2 (Heval2 : {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r2 }}), r1 = r2. - Let functional_eval_app_prop m n r1 (_ : {{ $| m & n |↘ r1 }}) : Prop := forall r2 (Heval2 : {{ $| m & n |↘ r2 }}), r1 = r2. - Let functional_eval_sub_prop σ p p1 (_ : {{ ⟦ σ ⟧s p ↘ p1 }}) : Prop := forall p2 (Heval2 : {{ ⟦ σ ⟧s p ↘ p2 }}), p1 = p2. - - #[local] - Ltac unfold_functional_eval := - unfold functional_eval_exp_prop, functional_eval_natrec_prop, functional_eval_app_prop, functional_eval_sub_prop in *; - clear functional_eval_exp_prop functional_eval_natrec_prop functional_eval_app_prop functional_eval_sub_prop. - - Lemma functional_eval_exp : forall M p m1 (Heval1 : {{ ⟦ M ⟧ p ↘ m1 }}), functional_eval_exp_prop M p m1 Heval1. + Lemma functional_eval : + (forall M ρ m1, + {{ ⟦ M ⟧ ρ ↘ m1 }} -> + forall m2, + {{ ⟦ M ⟧ ρ ↘ m2 }} -> + m1 = m2) /\ + (forall A MZ MS m ρ r1, + {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r1 }} -> + forall r2, + {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r2 }} -> + r1 = r2) /\ + (forall m n r1, + {{ $| m & n |↘ r1 }} -> + forall r2, + {{ $| m & n |↘ r2 }} -> + r1 = r2) /\ + (forall σ ρ ρσ1, + {{ ⟦ σ ⟧s ρ ↘ ρσ1 }} -> + forall ρσ2, + {{ ⟦ σ ⟧s ρ ↘ ρσ2 }} -> + ρσ1 = ρσ2). Proof with ((on_all_hyp: fun H => erewrite H in *; eauto); solve [eauto]) using. - induction Heval1 - using eval_exp_mut_ind - with (P0 := functional_eval_natrec_prop) - (P1 := functional_eval_app_prop) - (P2 := functional_eval_sub_prop); - unfold_functional_eval; - inversion_clear 1; do 2 f_equal... + apply eval_mut_ind; intros; + progressive_inversion; do 2 f_equal; try reflexivity... Qed. - Lemma functional_eval_natrec : forall A MZ MS m p r1 (Heval1 : {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r1 }}), functional_eval_natrec_prop A MZ MS m p r1 Heval1. - Proof with ((on_all_hyp: fun H => erewrite H in *; eauto); solve [eauto]) using. - induction Heval1 - using eval_natrec_mut_ind - with (P := functional_eval_exp_prop) - (P1 := functional_eval_app_prop) - (P2 := functional_eval_sub_prop); - unfold_functional_eval; - inversion_clear 1; do 2 f_equal... + Corollary functional_eval_exp : forall M ρ m1 m2, + {{ ⟦ M ⟧ ρ ↘ m1 }} -> + {{ ⟦ M ⟧ ρ ↘ m2 }} -> + m1 = m2. + Proof. + pose proof functional_eval; firstorder. Qed. - Lemma functional_eval_app : forall m n r1 (Heval1 : {{ $| m & n |↘ r1 }}), functional_eval_app_prop m n r1 Heval1. - Proof with ((on_all_hyp: fun H => erewrite H in *; eauto); solve [eauto]) using. - induction Heval1 - using eval_app_mut_ind - with (P := functional_eval_exp_prop) - (P0 := functional_eval_natrec_prop) - (P2 := functional_eval_sub_prop); - unfold_functional_eval; - inversion_clear 1; do 2 f_equal... + Corollary functional_eval_natrec : forall A MZ MS m ρ r1 r2, + {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r1 }} -> + {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r2 }} -> + r1 = r2. + Proof. + pose proof functional_eval; intuition. Qed. - Lemma functional_eval_sub : forall σ p p1 (Heval1 : {{ ⟦ σ ⟧s p ↘ p1 }}), functional_eval_sub_prop σ p p1 Heval1. - Proof with ((on_all_hyp: fun H => erewrite H in *; eauto); solve [eauto]) using. - induction Heval1 - using eval_sub_mut_ind - with (P := functional_eval_exp_prop) - (P0 := functional_eval_natrec_prop) - (P1 := functional_eval_app_prop); - unfold_functional_eval; - inversion_clear 1; do 2 f_equal... + Corollary functional_eval_app : forall m n r1 r2, + {{ $| m & n |↘ r1 }} -> + {{ $| m & n |↘ r2 }} -> + r1 = r2. + Proof. + pose proof functional_eval; intuition. + Qed. + + Corollary functional_eval_sub : forall σ ρ ρσ1 ρσ2, + {{ ⟦ σ ⟧s ρ ↘ ρσ1 }} -> + {{ ⟦ σ ⟧s ρ ↘ ρσ2 }} -> + ρσ1 = ρσ2. + Proof. + pose proof functional_eval; firstorder. Qed. End functional_eval. @@ -64,13 +68,13 @@ Hint Resolve functional_eval_exp functional_eval_natrec functional_eval_app func Ltac functional_eval_rewrite_clear1 := let tactic_error o1 o2 := fail 3 "functional_eval equality between" o1 "and" o2 "cannot be solved by mauto" in match goal with - | H1 : {{ ⟦ ~?M ⟧ ~?p ↘ ~?m1 }}, H2 : {{ ⟦ ~?M ⟧ ~?p ↘ ~?m2 }} |- _ => + | H1 : {{ ⟦ ~?M ⟧ ~?ρ ↘ ~?m1 }}, H2 : {{ ⟦ ~?M ⟧ ~?ρ ↘ ~?m2 }} |- _ => clean replace m2 with m1 by first [solve [mauto 2] | tactic_error m2 m1]; clear H2 | H1 : {{ $| ~?m & ~?n |↘ ~?r1 }}, H2 : {{ $| ~?m & ~?n |↘ ~?r2 }} |- _ => clean replace r2 with r1 by first [solve [mauto 2] | tactic_error r2 r1]; clear H2 - | H1 : {{ rec ~?m ⟦return ~?A | zero -> ~?MZ | succ -> ~?MS end⟧ ~?p ↘ ~?r1 }}, H2 : {{ rec ~?m ⟦return ~?A | zero -> ~?MZ | succ -> ~?MS end⟧ ~?p ↘ ~?r2 }} |- _ => + | H1 : {{ rec ~?m ⟦return ~?A | zero -> ~?MZ | succ -> ~?MS end⟧ ~?ρ ↘ ~?r1 }}, H2 : {{ rec ~?m ⟦return ~?A | zero -> ~?MZ | succ -> ~?MS end⟧ ~?ρ ↘ ~?r2 }} |- _ => clean replace r2 with r1 by first [solve [mauto 2] | tactic_error r2 r1]; clear H2 - | H1 : {{ ⟦ ~?σ ⟧s ~?p ↘ ~?p1 }}, H2 : {{ ⟦ ~?σ ⟧s ~?p ↘ ~?p2 }} |- _ => - clean replace p2 with p1 by first [solve [mauto 2] | tactic_error p2 p1]; clear H2 + | H1 : {{ ⟦ ~?σ ⟧s ~?ρ ↘ ~?ρσ1 }}, H2 : {{ ⟦ ~?σ ⟧s ~?ρ ↘ ~?ρσ2 }} |- _ => + clean replace ρσ2 with ρσ1 by first [solve [mauto 2] | tactic_error ρσ2 ρσ1]; clear H2 end. Ltac functional_eval_rewrite_clear := repeat functional_eval_rewrite_clear1. diff --git a/theories/Core/Semantic/Evaluation/Tactics.v b/theories/Core/Semantic/Evaluation/Tactics.v index 1a1694f5..5669f3b2 100644 --- a/theories/Core/Semantic/Evaluation/Tactics.v +++ b/theories/Core/Semantic/Evaluation/Tactics.v @@ -4,7 +4,9 @@ Import Domain_Notations. Ltac simplify_evals := functional_eval_rewrite_clear; clear_dups; - repeat (match_by_head eval_exp ltac:(fun H => directed inversion H; subst; clear H) - || match_by_head eval_sub ltac:(fun H => directed inversion H; subst; clear H)); + repeat (match_by_head eval_exp ltac:(fun H => directed dependent destruction H) + || match_by_head eval_app ltac:(fun H => directed dependent destruction H) + || match_by_head eval_natrec ltac:(fun H => directed dependent destruction H) + || match_by_head eval_sub ltac:(fun H => directed dependent destruction H)); functional_eval_rewrite_clear; clear_dups. diff --git a/theories/Core/Semantic/NbE.v b/theories/Core/Semantic/NbE.v index 87c3cc14..16b5ced3 100644 --- a/theories/Core/Semantic/NbE.v +++ b/theories/Core/Semantic/NbE.v @@ -7,18 +7,18 @@ Generalizable All Variables. Inductive initial_env : ctx -> env -> Prop := | initial_env_nil : initial_env nil empty_env | initial_env_cons : - `( initial_env Γ p -> - {{ ⟦ A ⟧ p ↘ a }} -> - initial_env (A :: Γ) d{{{ p ↦ ⇑! a (length Γ) }}}). + `( initial_env Γ ρ -> + {{ ⟦ A ⟧ ρ ↘ a }} -> + initial_env (A :: Γ) d{{{ ρ ↦ ⇑! a (length Γ) }}}). #[export] Hint Constructors initial_env : mcltt. -Lemma functional_initial_env : forall Γ p, - initial_env Γ p -> - forall p', - initial_env Γ p' -> - p = p'. +Lemma functional_initial_env : forall Γ ρ, + initial_env Γ ρ -> + forall ρ', + initial_env Γ ρ' -> + ρ = ρ'. Proof. induction 1; intros ? Hother; inversion_clear Hother; eauto. erewrite IHinitial_env in *; try eassumption; @@ -32,16 +32,16 @@ Hint Resolve functional_initial_env : mcltt. Ltac functional_initial_env_rewrite_clear1 := let tactic_error o1 o2 := fail 3 "functional_initial_env equality between" o1 "and" o2 "cannot be solved by mauto" in match goal with - | H1 : initial_env ?G ?p, H2 : initial_env ?G ?p' |- _ => - clean replace p' with p by first [solve [mauto 2] | tactic_error p' p]; clear H2 + | H1 : initial_env ?G ?ρ, H2 : initial_env ?G ?ρ' |- _ => + clean replace ρ' with ρ by first [solve [mauto 2] | tactic_error ρ' ρ]; clear H2 end. Ltac functional_initial_env_rewrite_clear := repeat functional_initial_env_rewrite_clear1. Inductive nbe : ctx -> exp -> typ -> nf -> Prop := | nbe_run : - `( initial_env Γ p -> - {{ ⟦ A ⟧ p ↘ a }} -> - {{ ⟦ M ⟧ p ↘ m }} -> + `( initial_env Γ ρ -> + {{ ⟦ A ⟧ ρ ↘ a }} -> + {{ ⟦ M ⟧ ρ ↘ m }} -> {{ Rnf ⇓ a m in (length Γ) ↘ w }} -> nbe Γ M A w ). @@ -117,8 +117,8 @@ Hint Resolve functional_nbe_of_typ : mcltt. Inductive nbe_ty : ctx -> typ -> nf -> Prop := | nbe_ty_run : - `( initial_env Γ p -> - {{ ⟦ M ⟧ p ↘ m }} -> + `( initial_env Γ ρ -> + {{ ⟦ M ⟧ ρ ↘ m }} -> {{ Rtyp m in (length Γ) ↘ W }} -> nbe_ty Γ M W ). diff --git a/theories/Core/Semantic/PER/Definitions.v b/theories/Core/Semantic/PER/Definitions.v index 8b34c85b..3e38aafa 100644 --- a/theories/Core/Semantic/PER/Definitions.v +++ b/theories/Core/Semantic/PER/Definitions.v @@ -18,7 +18,7 @@ Generalizable All Variables. (** Helper Bundles *) (* Related modulo evaluation *) -Inductive rel_mod_eval (R : relation domain -> domain -> domain -> Prop) A p A' p' R' : Prop := mk_rel_mod_eval : forall a a', {{ ⟦ A ⟧ p ↘ a }} -> {{ ⟦ A' ⟧ p' ↘ a' }} -> {{ DF a ≈ a' ∈ R ↘ R' }} -> rel_mod_eval R A p A' p' R'. +Inductive rel_mod_eval (R : relation domain -> domain -> domain -> Prop) A ρ A' ρ' R' : Prop := mk_rel_mod_eval : forall a a', {{ ⟦ A ⟧ ρ ↘ a }} -> {{ ⟦ A' ⟧ ρ' ↘ a' }} -> {{ DF a ≈ a' ∈ R ↘ R' }} -> rel_mod_eval R A ρ A' ρ' R'. #[global] Arguments mk_rel_mod_eval {_ _ _ _ _ _}. #[export] @@ -102,9 +102,9 @@ Section Per_univ_elem_core_def. (equiv_a_a' : {{ DF a ≈ a' ∈ per_univ_elem_core ↘ in_rel}}), PER in_rel -> (forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), - rel_mod_eval per_univ_elem_core B d{{{ p ↦ c }}} B' d{{{ p' ↦ c' }}} (out_rel equiv_c_c')) -> + rel_mod_eval per_univ_elem_core B d{{{ ρ ↦ c }}} B' d{{{ ρ' ↦ c' }}} (out_rel equiv_c_c')) -> (elem_rel <~> fun f f' => forall c c' (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_mod_app f c f' c' (out_rel equiv_c_c')) -> - {{ DF Π a p B ≈ Π a' p' B' ∈ per_univ_elem_core ↘ elem_rel }} } + {{ DF Π a ρ B ≈ Π a' ρ' B' ∈ per_univ_elem_core ↘ elem_rel }} } | per_univ_elem_core_neut : `{ forall (elem_rel : relation domain), {{ Dom b ≈ b' ∈ per_bot }} -> @@ -122,16 +122,16 @@ Section Per_univ_elem_core_def. (elem_rel <~> per_nat) -> motive elem_rel d{{{ ℕ }}} d{{{ ℕ }}}) (case_Pi : - forall {A p B A' p' B' in_rel} + forall {a ρ B a' ρ' B' in_rel} (out_rel : forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), relation domain) {elem_rel}, - {{ DF A ≈ A' ∈ per_univ_elem_core ↘ in_rel }} -> - motive in_rel A A' -> + {{ DF a ≈ a' ∈ per_univ_elem_core ↘ in_rel }} -> + motive in_rel a a' -> PER in_rel -> (forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), - rel_mod_eval (fun R x y => {{ DF x ≈ y ∈ per_univ_elem_core ↘ R }} /\ motive R x y) B d{{{ p ↦ c }}} B' d{{{ p' ↦ c' }}} (out_rel equiv_c_c')) -> + rel_mod_eval (fun R x y => {{ DF x ≈ y ∈ per_univ_elem_core ↘ R }} /\ motive R x y) B d{{{ ρ ↦ c }}} B' d{{{ ρ' ↦ c' }}} (out_rel equiv_c_c')) -> (elem_rel <~> fun f f' => forall c c' (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_mod_app f c f' c' (out_rel equiv_c_c')) -> - motive elem_rel d{{{ Π A p B }}} d{{{ Π A' p' B' }}}) + motive elem_rel d{{{ Π a ρ B }}} d{{{ Π a' ρ' B' }}}) (case_ne : forall {a b a' b' elem_rel}, {{ Dom b ≈ b' ∈ per_bot }} -> (elem_rel <~> per_ne) -> @@ -193,16 +193,16 @@ Section Per_univ_elem_ind_def. (elem_rel <~> per_nat) -> motive i elem_rel d{{{ ℕ }}} d{{{ ℕ }}}) (case_Pi : - forall i {A p B A' p' B' in_rel} + forall i {a ρ B a' ρ' B' in_rel} (out_rel : forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), relation domain) {elem_rel}, - {{ DF A ≈ A' ∈ per_univ_elem i ↘ in_rel }} -> - motive i in_rel A A' -> + {{ DF a ≈ a' ∈ per_univ_elem i ↘ in_rel }} -> + motive i in_rel a a' -> PER in_rel -> (forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), - rel_mod_eval (fun R x y => {{ DF x ≈ y ∈ per_univ_elem i ↘ R }} /\ motive i R x y) B d{{{ p ↦ c }}} B' d{{{ p' ↦ c' }}} (out_rel equiv_c_c')) -> + rel_mod_eval (fun R x y => {{ DF x ≈ y ∈ per_univ_elem i ↘ R }} /\ motive i R x y) B d{{{ ρ ↦ c }}} B' d{{{ ρ' ↦ c' }}} (out_rel equiv_c_c')) -> (elem_rel <~> fun f f' => forall c c' (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_mod_app f c f' c' (out_rel equiv_c_c')) -> - motive i elem_rel d{{{ Π A p B }}} d{{{ Π A' p' B' }}}) + motive i elem_rel d{{{ Π a ρ B }}} d{{{ Π a' ρ' B' }}}) (case_ne : forall i {a b a' b' elem_rel}, {{ Dom b ≈ b' ∈ per_bot }} -> (elem_rel <~> per_ne) -> @@ -244,12 +244,12 @@ Inductive per_subtyp : nat -> domain -> domain -> Prop := {{ DF a ≈ a' ∈ per_univ_elem i ↘ in_rel }} -> (forall c c' b b', {{ Dom c ≈ c' ∈ in_rel }} -> - {{ ⟦ B ⟧ p ↦ c ↘ b }} -> - {{ ⟦ B' ⟧ p' ↦ c' ↘ b' }} -> + {{ ⟦ B ⟧ ρ ↦ c ↘ b }} -> + {{ ⟦ B' ⟧ ρ' ↦ c' ↘ b' }} -> {{ Sub b <: b' at i }}) -> - {{ DF Π a p B ≈ Π a p B ∈ per_univ_elem i ↘ elem_rel }} -> - {{ DF Π a' p' B' ≈ Π a' p' B' ∈ per_univ_elem i ↘ elem_rel' }} -> - {{ Sub Π a p B <: Π a' p' B' at i }}) + {{ DF Π a ρ B ≈ Π a ρ B ∈ per_univ_elem i ↘ elem_rel }} -> + {{ DF Π a' ρ' B' ≈ Π a' ρ' B' ∈ per_univ_elem i ↘ elem_rel' }} -> + {{ Sub Π a ρ B <: Π a' ρ' B' at i }}) where "'Sub' a <: b 'at' i" := (per_subtyp i a b) (in custom judg) : type_scope. #[export] @@ -257,7 +257,7 @@ where "'Sub' a <: b 'at' i" := (per_subtyp i a b) (in custom judg) : type_scope. (** Context/Environment PER *) -Definition rel_typ (i : nat) (A : typ) (p : env) (A' : typ) (p' : env) R' := rel_mod_eval (per_univ_elem i) A p A' p' R'. +Definition rel_typ i A ρ A' ρ' R' := rel_mod_eval (per_univ_elem i) A ρ A' ρ' R'. Arguments rel_typ _ _ _ _ _ _ /. #[export] Hint Transparent rel_typ : mcltt. @@ -267,19 +267,19 @@ Hint Unfold rel_typ : mcltt. Inductive per_ctx_env : relation env -> ctx -> ctx -> Prop := | per_ctx_env_nil : `{ forall env_rel, - (env_rel <~> fun p p' => True) -> + (env_rel <~> fun ρ ρ' => True) -> {{ EF ⋅ ≈ ⋅ ∈ per_ctx_env ↘ env_rel }} } | per_ctx_env_cons : `{ forall tail_rel - (head_rel : forall {p p'} (equiv_p_p' : {{ Dom p ≈ p' ∈ tail_rel }}), relation domain) + (head_rel : forall {ρ ρ'} (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ tail_rel }}), relation domain) env_rel (equiv_Γ_Γ' : {{ EF Γ ≈ Γ' ∈ per_ctx_env ↘ tail_rel }}), PER tail_rel -> - (forall {p p'} (equiv_p_p' : {{ Dom p ≈ p' ∈ tail_rel }}), - rel_typ i A p A' p' (head_rel equiv_p_p')) -> - (env_rel <~> fun p p' => - exists (equiv_p_drop_p'_drop : {{ Dom p ↯ ≈ p' ↯ ∈ tail_rel }}), - {{ Dom ~(p 0) ≈ ~(p' 0) ∈ head_rel equiv_p_drop_p'_drop }}) -> + (forall {ρ ρ'} (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ tail_rel }}), + rel_typ i A ρ A' ρ' (head_rel equiv_ρ_ρ')) -> + (env_rel <~> fun ρ ρ' => + exists (equiv_ρ_drop_ρ'_drop : {{ Dom ρ ↯ ≈ ρ' ↯ ∈ tail_rel }}), + {{ Dom ~(ρ 0) ≈ ~(ρ' 0) ∈ head_rel equiv_ρ_drop_ρ'_drop }}) -> {{ EF Γ, A ≈ Γ', A' ∈ per_ctx_env ↘ env_rel }} } . #[export] @@ -294,7 +294,6 @@ Hint Unfold valid_ctx : mcltt. Reserved Notation "'SubE' Γ <: Δ" (in custom judg at level 90, Γ custom exp, Δ custom exp). - Inductive per_ctx_subtyp : ctx -> ctx -> Prop := | per_ctx_subtyp_nil : {{ SubE ⋅ <: ⋅ }} @@ -302,10 +301,10 @@ Inductive per_ctx_subtyp : ctx -> ctx -> Prop := `{ forall tail_rel env_rel env_rel', {{ SubE Γ <: Γ' }} -> {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ tail_rel }} -> - (forall p p' a a' - (equiv_p_p' : {{ Dom p ≈ p' ∈ tail_rel }}), - {{ ⟦ A ⟧ p ↘ a }} -> - {{ ⟦ A' ⟧ p' ↘ a' }} -> + (forall ρ ρ' a a' + (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ tail_rel }}), + {{ ⟦ A ⟧ ρ ↘ a }} -> + {{ ⟦ A' ⟧ ρ' ↘ a' }} -> {{ Sub a <: a' at i }}) -> {{ EF Γ , A ≈ Γ , A ∈ per_ctx_env ↘ env_rel }} -> {{ EF Γ' , A' ≈ Γ' , A' ∈ per_ctx_env ↘ env_rel' }} -> diff --git a/theories/Core/Semantic/PER/Lemmas.v b/theories/Core/Semantic/PER/Lemmas.v index 5b8583d9..01ec4fed 100644 --- a/theories/Core/Semantic/PER/Lemmas.v +++ b/theories/Core/Semantic/PER/Lemmas.v @@ -4,7 +4,7 @@ From Mcltt Require Import Base LibTactics. From Mcltt.Core Require Import PER.Definitions PER.CoreTactics. Import Domain_Notations. -Add Parametric Morphism R0 `(R0_morphism : Proper _ ((@relation_equivalence domain) ==> (@relation_equivalence domain)) R0) A p A' p' : (rel_mod_eval R0 A p A' p') +Add Parametric Morphism R0 `(R0_morphism : Proper _ ((@relation_equivalence domain) ==> (@relation_equivalence domain)) R0) A ρ A' ρ' : (rel_mod_eval R0 A ρ A' ρ') with signature (@relation_equivalence domain) ==> iff as rel_mod_eval_morphism. Proof. split; intros []; econstructor; try eassumption; @@ -227,7 +227,7 @@ Proof with mautosolve. reflexivity. Qed. -Add Parametric Morphism i A p A' p' : (rel_typ i A p A' p') +Add Parametric Morphism i A ρ A' ρ' : (rel_typ i A ρ A' ρ') with signature (@relation_equivalence domain) ==> iff as rel_typ_morphism. Proof. intros * HRR'. @@ -401,24 +401,24 @@ Qed. Ltac do_per_univ_elem_irrel_assert1 := let tactic_error o1 o2 := fail 2 "per_univ_elem_irrel biconditional between" o1 "and" o2 "cannot be solved" in match goal with - | H1 : {{ DF ~?A ≈ ~_ ∈ per_univ_elem ?i ↘ ?R1 }}, - H2 : {{ DF ~?A ≈ ~_ ∈ per_univ_elem ?i' ↘ ?R2 }} |- _ => + | H1 : {{ DF ~?a ≈ ~_ ∈ per_univ_elem ?i ↘ ?R1 }}, + H2 : {{ DF ~?a ≈ ~_ ∈ per_univ_elem ?i' ↘ ?R2 }} |- _ => assert_fails (unify R1 R2); match goal with | H : R1 <~> R2 |- _ => fail 1 | H : R2 <~> R1 |- _ => fail 1 | _ => assert (R1 <~> R2) by (eapply per_univ_elem_right_irrel; [apply H1 | apply H2]) || tactic_error R1 R2 end - | H1 : {{ DF ~_ ≈ ~?B ∈ per_univ_elem ?i ↘ ?R1 }}, - H2 : {{ DF ~_ ≈ ~?B ∈ per_univ_elem ?i' ↘ ?R2 }} |- _ => + | H1 : {{ DF ~_ ≈ ~?b ∈ per_univ_elem ?i ↘ ?R1 }}, + H2 : {{ DF ~_ ≈ ~?b ∈ per_univ_elem ?i' ↘ ?R2 }} |- _ => assert_fails (unify R1 R2); match goal with | H : R1 <~> R2 |- _ => fail 1 | H : R2 <~> R1 |- _ => fail 1 | _ => assert (R1 <~> R2) by (eapply per_univ_elem_left_irrel; [apply H1 | apply H2]) || tactic_error R1 R2 end - | H1 : {{ DF ~?A ≈ ~_ ∈ per_univ_elem ?i ↘ ?R1 }}, - H2 : {{ DF ~_ ≈ ~?A ∈ per_univ_elem ?i' ↘ ?R2 }} |- _ => + | H1 : {{ DF ~?a ≈ ~_ ∈ per_univ_elem ?i ↘ ?R1 }}, + H2 : {{ DF ~_ ≈ ~?a ∈ per_univ_elem ?i' ↘ ?R2 }} |- _ => (* Order matters less here as H1 and H2 cannot be exchanged *) assert_fails (unify R1 R2); match goal with @@ -437,15 +437,15 @@ Ltac handle_per_univ_elem_irrel := apply_relation_equivalence; clear_dups. -Lemma per_univ_elem_trans : forall i R A1 A2, - per_univ_elem i R A1 A2 -> - (forall j A3, - per_univ_elem j R A2 A3 -> - per_univ_elem i R A1 A3) /\ - (forall a1 a2 a3, - R a1 a2 -> - R a2 a3 -> - R a1 a3). +Lemma per_univ_elem_trans : forall i R a1 a2, + per_univ_elem i R a1 a2 -> + (forall j a3, + per_univ_elem j R a2 a3 -> + per_univ_elem i R a1 a3) /\ + (forall m1 m2 m3, + R m1 m2 -> + R m2 m3 -> + R m1 m3). Proof with (basic_per_univ_elem_econstructor; mautosolve). induction 1 using per_univ_elem_ind; [> split; @@ -486,30 +486,30 @@ Proof with (basic_per_univ_elem_econstructor; mautosolve). idtac... Qed. -Corollary per_univ_trans : forall i j R A1 A2 A3, - per_univ_elem i R A1 A2 -> - per_univ_elem j R A2 A3 -> - per_univ_elem i R A1 A3. +Corollary per_univ_trans : forall i j R a1 a2 a3, + per_univ_elem i R a1 a2 -> + per_univ_elem j R a2 a3 -> + per_univ_elem i R a1 a3. Proof. intros * ?%per_univ_elem_trans. firstorder. Qed. -Corollary per_univ_trans' : forall i j A1 A2 A3, - {{ Dom A1 ≈ A2 ∈ per_univ i }} -> - {{ Dom A2 ≈ A3 ∈ per_univ j }} -> - {{ Dom A1 ≈ A3 ∈ per_univ i }}. +Corollary per_univ_trans' : forall i j a1 a2 a3, + {{ Dom a1 ≈ a2 ∈ per_univ i }} -> + {{ Dom a2 ≈ a3 ∈ per_univ j }} -> + {{ Dom a1 ≈ a3 ∈ per_univ i }}. Proof. intros * [? ?] [? ?]. handle_per_univ_elem_irrel. firstorder mauto using per_univ_trans. Qed. -Corollary per_elem_trans : forall i R A1 A2 a1 a2 a3, - per_univ_elem i R A1 A2 -> - R a1 a2 -> - R a2 a3 -> - R a1 a3. +Corollary per_elem_trans : forall i R a1 a2 m1 m2 m3, + per_univ_elem i R a1 a2 -> + R m1 m2 -> + R m2 m3 -> + R m1 m3. Proof. intros * ?% per_univ_elem_trans. firstorder. @@ -532,7 +532,7 @@ Proof. Qed. #[export] -Instance per_elem_PER {i R A B} `(H : per_univ_elem i R A B) : PER R. +Instance per_elem_PER {i R a b} `(H : per_univ_elem i R a b) : PER R. Proof. split. - eauto using (per_elem_sym _ _ _ _ _ _ H). @@ -541,15 +541,15 @@ Qed. (* This lemma gets rid of the unnecessary PER premise. *) Lemma per_univ_elem_pi' : - forall i a a' p B p' B' + forall i a a' ρ B ρ' B' (in_rel : relation domain) (out_rel : forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), relation domain) elem_rel, {{ DF a ≈ a' ∈ per_univ_elem i ↘ in_rel}} -> (forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), - rel_mod_eval (per_univ_elem i) B d{{{ p ↦ c }}} B' d{{{ p' ↦ c' }}} (out_rel equiv_c_c')) -> + rel_mod_eval (per_univ_elem i) B d{{{ ρ ↦ c }}} B' d{{{ ρ' ↦ c' }}} (out_rel equiv_c_c')) -> (elem_rel <~> fun f f' => forall c c' (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_mod_app f c f' c' (out_rel equiv_c_c')) -> - {{ DF Π a p B ≈ Π a' p' B' ∈ per_univ_elem i ↘ elem_rel }}. + {{ DF Π a ρ B ≈ Π a' ρ' B' ∈ per_univ_elem i ↘ elem_rel }}. Proof. intros. basic_per_univ_elem_econstructor; eauto. @@ -562,12 +562,12 @@ Ltac per_univ_elem_econstructor := #[export] Hint Resolve per_univ_elem_pi' : mcltt. -Lemma per_univ_elem_pi_clean_inversion : forall {i j a a' in_rel p p' B B' elem_rel}, +Lemma per_univ_elem_pi_clean_inversion : forall {i j a a' in_rel ρ ρ' B B' elem_rel}, {{ DF a ≈ a' ∈ per_univ_elem i ↘ in_rel }} -> - {{ DF Π a p B ≈ Π a' p' B' ∈ per_univ_elem j ↘ elem_rel }} -> + {{ DF Π a ρ B ≈ Π a' ρ' B' ∈ per_univ_elem j ↘ elem_rel }} -> exists (out_rel : forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), relation domain), (forall c c' (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), - rel_mod_eval (per_univ_elem j) B d{{{ p ↦ c }}} B' d{{{ p' ↦ c' }}} (out_rel equiv_c_c')) /\ + rel_mod_eval (per_univ_elem j) B d{{{ ρ ↦ c }}} B' d{{{ ρ' ↦ c' }}} (out_rel equiv_c_c')) /\ (elem_rel <~> fun f f' => forall c c' (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_mod_app f c f' c' (out_rel equiv_c_c')). Proof. intros * Ha HΠ. @@ -577,7 +577,7 @@ Proof. split. - instantiate (1 := fun c c' (equiv_c_c' : in_rel c c') m m' => forall R, - rel_typ j B d{{{ p ↦ c }}} B' d{{{ p' ↦ c' }}} R -> + rel_typ j B d{{{ ρ ↦ c }}} B' d{{{ ρ' ↦ c' }}} R -> R m m'). intros. assert (in_rel0 c c') by intuition. @@ -682,12 +682,12 @@ Proof. intuition. Qed. -Lemma per_elem_subtyping_gen : forall A B i A' B' R R' a b, - {{ Sub A <: B at i }} -> - {{ DF A ≈ A' ∈ per_univ_elem i ↘ R }} -> - {{ DF B ≈ B' ∈ per_univ_elem i ↘ R' }} -> - R a b -> - R' a b. +Lemma per_elem_subtyping_gen : forall a b i a' b' R R' m n, + {{ Sub a <: b at i }} -> + {{ DF a ≈ a' ∈ per_univ_elem i ↘ R }} -> + {{ DF b ≈ b' ∈ per_univ_elem i ↘ R' }} -> + R m n -> + R' m n. Proof. intros. eapply per_elem_subtyping; saturate_refl; try eassumption. @@ -701,7 +701,7 @@ Proof. subst; mauto; destruct_all. - assert ({{ DF Π A p B ≈ Π A' p' B' ∈ per_univ_elem i ↘ elem_rel }}) + assert ({{ DF Π a ρ B ≈ Π a' ρ' B' ∈ per_univ_elem i ↘ elem_rel }}) by (eapply per_univ_elem_pi'; eauto; intros; destruct_rel_mod_eval; mauto). saturate_refl. econstructor; eauto. @@ -726,11 +726,11 @@ Qed. #[export] Hint Resolve per_subtyp_refl2 : mcltt. -Lemma per_subtyp_trans : forall A1 A2 i, - {{ Sub A1 <: A2 at i }} -> - forall A3, - {{ Sub A2 <: A3 at i }} -> - {{ Sub A1 <: A3 at i }}. +Lemma per_subtyp_trans : forall a1 a2 i, + {{ Sub a1 <: a2 at i }} -> + forall a3, + {{ Sub a2 <: a3 at i }} -> + {{ Sub a1 <: a3 at i }}. Proof. induction 1; intros ? Hsub; simpl in *. 1-3: progressive_inversion; mauto. @@ -756,39 +756,39 @@ Proof. eauto using per_subtyp_trans. Qed. -Lemma per_subtyp_transp : forall A B i A' B' R R', - {{ Sub A <: B at i }} -> - {{ DF A ≈ A' ∈ per_univ_elem i ↘ R }} -> - {{ DF B ≈ B' ∈ per_univ_elem i ↘ R' }} -> - {{ Sub A' <: B' at i }}. +Lemma per_subtyp_transp : forall a b i a' b' R R', + {{ Sub a <: b at i }} -> + {{ DF a ≈ a' ∈ per_univ_elem i ↘ R }} -> + {{ DF b ≈ b' ∈ per_univ_elem i ↘ R' }} -> + {{ Sub a' <: b' at i }}. Proof. mauto using per_subtyp_refl1, per_subtyp_refl2. Qed. -Lemma per_subtyp_cumu : forall A1 A2 i, - {{ Sub A1 <: A2 at i }} -> +Lemma per_subtyp_cumu : forall a1 a2 i, + {{ Sub a1 <: a2 at i }} -> forall j, i <= j -> - {{ Sub A1 <: A2 at j }}. + {{ Sub a1 <: a2 at j }}. Proof. induction 1; intros; econstructor; mauto. lia. Qed. #[export] - Hint Resolve per_subtyp_cumu : mcltt. +Hint Resolve per_subtyp_cumu : mcltt. -Lemma per_subtyp_cumu_left : forall A1 A2 i j, - {{ Sub A1 <: A2 at i }} -> - {{ Sub A1 <: A2 at max i j }}. +Lemma per_subtyp_cumu_left : forall a1 a2 i j, + {{ Sub a1 <: a2 at i }} -> + {{ Sub a1 <: a2 at max i j }}. Proof. intros. eapply per_subtyp_cumu; try eassumption. lia. Qed. -Lemma per_subtyp_cumu_right : forall A1 A2 i j, - {{ Sub A1 <: A2 at i }} -> - {{ Sub A1 <: A2 at max j i }}. +Lemma per_subtyp_cumu_right : forall a1 a2 i j, + {{ Sub a1 <: a2 at i }} -> + {{ Sub a1 <: a2 at max j i }}. Proof. intros. eapply per_subtyp_cumu; try eassumption. lia. @@ -823,35 +823,34 @@ Proof with (destruct_rel_typ; handle_per_univ_elem_irrel; eexists; intuition). apply_relation_equivalence; try reflexivity. specialize (IHHorig _ _ equiv_Γ_Γ'0). - intros p p'. + intros ρ ρ'. split; intros []. - - assert (equiv0_p_drop_p'_drop : {{ Dom p ↯ ≈ p' ↯ ∈ tail_rel0 }}) by intuition... - - assert (equiv_p_drop_p'_drop : {{ Dom p ↯ ≈ p' ↯ ∈ tail_rel }}) by intuition... + - assert {{ Dom ρ ↯ ≈ ρ' ↯ ∈ tail_rel0 }} by intuition... + - assert {{ Dom ρ ↯ ≈ ρ' ↯ ∈ tail_rel }} by intuition... Qed. Lemma per_ctx_env_sym : forall Γ Δ R, {{ DF Γ ≈ Δ ∈ per_ctx_env ↘ R }} -> {{ DF Δ ≈ Γ ∈ per_ctx_env ↘ R }} /\ - (forall o p, - {{ Dom o ≈ p ∈ R }} -> - {{ Dom p ≈ o ∈ R }}). + (forall ρ ρ', + {{ Dom ρ ≈ ρ' ∈ R }} -> + {{ Dom ρ' ≈ ρ ∈ R }}). Proof with solve [intuition]. simpl. induction 1; split; simpl in *; destruct_conjs; try econstructor; intuition; pose proof (@relation_equivalence_pointwise env). - - assert (tail_rel p' p) by eauto. - assert (tail_rel p p) by (etransitivity; eassumption). + - assert (tail_rel ρ' ρ) by eauto. + assert (tail_rel ρ ρ) by (etransitivity; eassumption). destruct_rel_mod_eval. handle_per_univ_elem_irrel. econstructor; eauto. symmetry... - apply_relation_equivalence. destruct_conjs. - assert (tail_rel d{{{ p ↯ }}} d{{{ o ↯ }}}) by eauto. - assert (tail_rel d{{{ p ↯ }}} d{{{ p ↯ }}}) by (etransitivity; eassumption). + assert (tail_rel d{{{ ρ' ↯ }}} d{{{ ρ ↯ }}}) by eauto. + assert (tail_rel d{{{ ρ ↯ }}} d{{{ ρ ↯ }}}) by (etransitivity; eassumption). destruct_rel_mod_eval. - eexists. - handle_per_univ_elem_irrel; symmetry; intuition. + eexists; symmetry; handle_per_univ_elem_irrel; intuition. Qed. Corollary per_ctx_sym : forall Γ Δ R, @@ -862,10 +861,10 @@ Proof. firstorder. Qed. -Corollary per_env_sym : forall Γ Δ R o p, +Corollary per_env_sym : forall Γ Δ R ρ ρ', {{ DF Γ ≈ Δ ∈ per_ctx_env ↘ R }} -> - {{ Dom o ≈ p ∈ R }} -> - {{ Dom p ≈ o ∈ R }}. + {{ Dom ρ ≈ ρ' ∈ R }} -> + {{ Dom ρ' ≈ ρ ∈ R }}. Proof. intros * ?%per_ctx_env_sym. firstorder. @@ -933,10 +932,10 @@ Lemma per_ctx_env_trans : forall Γ1 Γ2 R, (forall Γ3, {{ DF Γ2 ≈ Γ3 ∈ per_ctx_env ↘ R }} -> {{ DF Γ1 ≈ Γ3 ∈ per_ctx_env ↘ R }}) /\ - (forall p1 p2 p3, - {{ Dom p1 ≈ p2 ∈ R }} -> - {{ Dom p2 ≈ p3 ∈ R }} -> - {{ Dom p1 ≈ p3 ∈ R }}). + (forall ρ1 ρ2 ρ3, + {{ Dom ρ1 ≈ ρ2 ∈ R }} -> + {{ Dom ρ2 ≈ ρ3 ∈ R }} -> + {{ Dom ρ1 ≈ ρ3 ∈ R }}). Proof with solve [eauto using per_univ_trans]. simpl. induction 1; subst; @@ -949,15 +948,15 @@ Proof with solve [eauto using per_univ_trans]. - econstructor; only 4: reflexivity; eauto. + apply_relation_equivalence. intuition. + intros. - assert (tail_rel p p) by intuition. - assert (tail_rel0 p p') by intuition. + assert (tail_rel ρ ρ) by intuition. + assert (tail_rel0 ρ ρ') by intuition. destruct_rel_typ. handle_per_univ_elem_irrel. econstructor; intuition. (* This one cannot be replaced with `etransitivity` as we need different `i`s. *) eapply per_univ_trans; [| eassumption]; eassumption. - destruct_conjs. - assert (tail_rel d{{{ p1 ↯ }}} d{{{ p3 ↯ }}}) by eauto. + assert (tail_rel d{{{ ρ1 ↯ }}} d{{{ ρ3 ↯ }}}) by eauto. destruct_rel_typ. handle_per_univ_elem_irrel. eexists. @@ -974,11 +973,11 @@ Proof. firstorder. Qed. -Corollary per_env_trans : forall Γ1 Γ2 R p1 p2 p3, +Corollary per_env_trans : forall Γ1 Γ2 R ρ1 ρ2 ρ3, {{ DF Γ1 ≈ Γ2 ∈ per_ctx_env ↘ R }} -> - {{ Dom p1 ≈ p2 ∈ R }} -> - {{ Dom p2 ≈ p3 ∈ R }} -> - {{ Dom p1 ≈ p3 ∈ R }}. + {{ Dom ρ1 ≈ ρ2 ∈ R }} -> + {{ Dom ρ2 ≈ ρ3 ∈ R }} -> + {{ Dom ρ1 ≈ ρ3 ∈ R }}. Proof. intros * ?% per_ctx_env_trans. firstorder. @@ -1002,14 +1001,14 @@ Qed. (* This lemma removes the PER argument *) Lemma per_ctx_env_cons' : forall {Γ Γ' i A A' tail_rel} - (head_rel : forall {p p'} (equiv_p_p' : {{ Dom p ≈ p' ∈ tail_rel }}), relation domain) + (head_rel : forall {ρ ρ'} (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ tail_rel }}), relation domain) env_rel, {{ EF Γ ≈ Γ' ∈ per_ctx_env ↘ tail_rel }} -> - (forall {p p'} (equiv_p_p' : {{ Dom p ≈ p' ∈ tail_rel }}), - rel_typ i A p A' p' (head_rel equiv_p_p')) -> - (env_rel <~> fun p p' => - exists (equiv_p_drop_p'_drop : {{ Dom p ↯ ≈ p' ↯ ∈ tail_rel }}), - {{ Dom ~(p 0) ≈ ~(p' 0) ∈ head_rel equiv_p_drop_p'_drop }}) -> + (forall {ρ ρ'} (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ tail_rel }}), + rel_typ i A ρ A' ρ' (head_rel equiv_ρ_ρ')) -> + (env_rel <~> fun ρ ρ' => + exists (equiv_ρ_drop_ρ'_drop : {{ Dom ρ ↯ ≈ ρ' ↯ ∈ tail_rel }}), + {{ Dom ~(ρ 0) ≈ ~(ρ' 0) ∈ head_rel equiv_ρ_drop_ρ'_drop }}) -> {{ EF Γ, A ≈ Γ', A' ∈ per_ctx_env ↘ env_rel }}. Proof. intros. @@ -1026,12 +1025,12 @@ Ltac per_ctx_env_econstructor := Lemma per_ctx_env_cons_clean_inversion : forall {Γ Γ' env_relΓ A A' env_relΓA}, {{ EF Γ ≈ Γ' ∈ per_ctx_env ↘ env_relΓ }} -> {{ EF Γ, A ≈ Γ', A' ∈ per_ctx_env ↘ env_relΓA }} -> - exists i (head_rel : forall {p p'} (equiv_p_p' : {{ Dom p ≈ p' ∈ env_relΓ }}), relation domain), - (forall p p' (equiv_p_p' : {{ Dom p ≈ p' ∈ env_relΓ }}), - rel_typ i A p A' p' (head_rel equiv_p_p')) /\ - (env_relΓA <~> fun p p' => - exists (equiv_p_drop_p'_drop : {{ Dom p ↯ ≈ p' ↯ ∈ env_relΓ }}), - {{ Dom ~(p 0) ≈ ~(p' 0) ∈ head_rel equiv_p_drop_p'_drop }}). + exists i (head_rel : forall {ρ ρ'} (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_relΓ }}), relation domain), + (forall ρ ρ' (equiv_ρ_ρ' : {{ Dom ρ ≈ ρ' ∈ env_relΓ }}), + rel_typ i A ρ A' ρ' (head_rel equiv_ρ_ρ')) /\ + (env_relΓA <~> fun ρ ρ' => + exists (equiv_ρ_drop_ρ'_drop : {{ Dom ρ ↯ ≈ ρ' ↯ ∈ env_relΓ }}), + {{ Dom ~(ρ 0) ≈ ~(ρ' 0) ∈ head_rel equiv_ρ_drop_ρ'_drop }}). Proof with intuition. intros * HΓ HΓA. inversion HΓA; subst. @@ -1039,20 +1038,20 @@ Proof with intuition. eexists. eexists. split; intros. - - instantiate (1 := fun p p' (equiv_p_p' : env_relΓ p p') m m' => + - instantiate (1 := fun ρ ρ' (equiv_ρ_ρ' : env_relΓ ρ ρ') m m' => forall R, - rel_typ i A p A' p' R -> + rel_typ i A ρ A' ρ' R -> {{ Dom m ≈ m' ∈ R }}). - assert (tail_rel p p') by intuition. + assert (tail_rel ρ ρ') by intuition. (on_all_hyp: destruct_rel_by_assumption tail_rel). econstructor; eauto. apply -> per_univ_elem_morphism_iff; eauto. split; intros... destruct_by_head rel_typ. handle_per_univ_elem_irrel... - - intros o o'. + - intros ρ ρ'. split; intros; destruct_conjs; - assert {{ Dom o ↯ ≈ o' ↯ ∈ tail_rel }} by intuition; + assert {{ Dom ρ ↯ ≈ ρ' ↯ ∈ tail_rel }} by intuition; (on_all_hyp: destruct_rel_by_assumption tail_rel); unshelve eexists; intros... destruct_by_head rel_typ. @@ -1084,11 +1083,11 @@ Qed. Lemma per_ctx_env_subtyping : forall Γ Δ, {{ SubE Γ <: Δ }} -> - forall R R' p p', + forall R R' ρ ρ', {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ R }} -> {{ EF Δ ≈ Δ ∈ per_ctx_env ↘ R' }} -> - R p p' -> - R' p p'. + R ρ ρ' -> + R' ρ ρ'. Proof. induction 1; intros; handle_per_ctx_env_irrel; @@ -1097,7 +1096,7 @@ Proof. trivial. destruct_all. - assert {{ Dom p ↯ ≈ p' ↯ ∈ tail_rel0 }} by intuition. + assert {{ Dom ρ ↯ ≈ ρ' ↯ ∈ tail_rel0 }} by intuition. unshelve eexists; [eassumption |]. destruct_rel_typ. eapply per_elem_subtyping with (i := max x (max i0 i)); try eassumption. @@ -1147,7 +1146,7 @@ Proof. - firstorder. - instantiate (1 := max i i0). intros. - assert {{ Dom p ≈ p' ∈ tail_rel0 }} by (eapply per_ctx_env_subtyping; revgoals; eassumption). + assert {{ Dom ρ ≈ ρ' ∈ tail_rel0 }} by (eapply per_ctx_env_subtyping; revgoals; eassumption). saturate_refl_for tail_rel. destruct_rel_typ. handle_per_univ_elem_irrel. diff --git a/theories/Core/Semantic/Readback/Definitions.v b/theories/Core/Semantic/Readback/Definitions.v index 35f51a7a..77ff2073 100644 --- a/theories/Core/Semantic/Readback/Definitions.v +++ b/theories/Core/Semantic/Readback/Definitions.v @@ -26,10 +26,10 @@ Inductive read_nf : nat -> domain_nf -> nf -> Prop := (* Nf of eta-expanded body *) {{ $| m & ⇑! a s |↘ m' }} -> - {{ ⟦ B ⟧ p ↦ ⇑! a s ↘ b }} -> + {{ ⟦ B ⟧ ρ ↦ ⇑! a s ↘ b }} -> {{ Rnf ⇓ b m' in S s ↘ M }} -> - {{ Rnf ⇓ (Π a p B) m in s ↘ λ A M }} ) + {{ Rnf ⇓ (Π a ρ B) m in s ↘ λ A M }} ) | read_nf_neut : `( {{ Rne m in s ↘ M }} -> {{ Rnf ⇓ (⇑ a b) (⇑ c m) in s ↘ ⇑ M }} ) @@ -43,22 +43,22 @@ with read_ne : nat -> domain_ne -> ne -> Prop := {{ Rne m n in s ↘ M N }} ) | read_ne_natrec : (* Nf of motive *) - `( {{ ⟦ B ⟧ p ↦ ⇑! ℕ s ↘ b }} -> + `( {{ ⟦ B ⟧ ρ ↦ ⇑! ℕ s ↘ b }} -> {{ Rtyp b in S s ↘ B' }} -> (* Nf of mz *) - {{ ⟦ B ⟧ p ↦ zero ↘ bz }} -> + {{ ⟦ B ⟧ ρ ↦ zero ↘ bz }} -> {{ Rnf ⇓ bz mz in s ↘ MZ }} -> (* Nf of MS *) - {{ ⟦ B ⟧ p ↦ succ (⇑! ℕ s) ↘ bs }} -> - {{ ⟦ MS ⟧ p ↦ ⇑! ℕ s ↦ ⇑! b (S s) ↘ ms }} -> + {{ ⟦ B ⟧ ρ ↦ succ (⇑! ℕ s) ↘ bs }} -> + {{ ⟦ MS ⟧ ρ ↦ ⇑! ℕ s ↦ ⇑! b (S s) ↘ ms }} -> {{ Rnf ⇓ bs ms in S (S s) ↘ MS' }} -> (* Ne of m *) {{ Rne m in s ↘ M }} -> - {{ Rne rec m under p return B | zero -> mz | succ -> MS end in s ↘ rec M return B' | zero -> MZ | succ -> MS' end }} ) + {{ Rne rec m under ρ return B | zero -> mz | succ -> MS end in s ↘ rec M return B' | zero -> MZ | succ -> MS' end }} ) where "'Rne' m 'in' s ↘ M" := (read_ne s m M) (in custom judg) : type_scope with read_typ : nat -> domain -> nf -> Prop := | read_typ_univ : @@ -70,10 +70,10 @@ with read_typ : nat -> domain -> nf -> Prop := `( {{ Rtyp a in s ↘ A }} -> (* Nf of ret type *) - {{ ⟦ B ⟧ p ↦ ⇑! a s ↘ b }} -> + {{ ⟦ B ⟧ ρ ↦ ⇑! a s ↘ b }} -> {{ Rtyp b in S s ↘ B' }} -> - {{ Rtyp Π a p B in s ↘ Π A B' }}) + {{ Rtyp Π a ρ B in s ↘ Π A B' }}) | read_typ_neut : `( {{ Rne b in s ↘ B }} -> {{ Rtyp ⇑ a b in s ↘ ⇑ B }}) @@ -83,6 +83,10 @@ where "'Rtyp' m 'in' s ↘ M" := (read_typ s m M) (in custom judg) : type_scope Scheme read_nf_mut_ind := Induction for read_nf Sort Prop with read_ne_mut_ind := Induction for read_ne Sort Prop with read_typ_mut_ind := Induction for read_typ Sort Prop. +Combined Scheme read_mut_ind from + read_nf_mut_ind, + read_ne_mut_ind, + read_typ_mut_ind. #[export] Hint Constructors read_nf read_ne read_typ : mcltt. diff --git a/theories/Core/Semantic/Readback/Lemmas.v b/theories/Core/Semantic/Readback/Lemmas.v index 4b13188a..58d146d4 100644 --- a/theories/Core/Semantic/Readback/Lemmas.v +++ b/theories/Core/Semantic/Readback/Lemmas.v @@ -4,44 +4,48 @@ From Mcltt.Core Require Import Evaluation Readback.Definitions. Import Domain_Notations. Section functional_read. - Let functional_read_nf_prop s m M1 (_ : {{ Rnf m in s ↘ M1 }}) : Prop := forall M2 (Hread2: {{ Rnf m in s ↘ M2 }}), M1 = M2. - Let functional_read_ne_prop s m M1 (_ : {{ Rne m in s ↘ M1 }}) : Prop := forall M2 (Hread2 : {{ Rne m in s ↘ M2 }}), M1 = M2. - Let functional_read_typ_prop s m M1 (_ : {{ Rtyp m in s ↘ M1 }}) : Prop := forall M2 (Hread2 : {{ Rtyp m in s ↘ M2 }}), M1 = M2. - - #[local] - Ltac unfold_functional_read := unfold functional_read_nf_prop, functional_read_ne_prop, functional_read_typ_prop in *. - - Lemma functional_read_nf : forall s m M1 (Hread1: {{ Rnf m in s ↘ M1 }}), functional_read_nf_prop s m M1 Hread1. + Lemma functional_read : + (forall s m M1, + {{ Rnf m in s ↘ M1 }} -> + forall M2, + {{ Rnf m in s ↘ M2 }} -> + M1 = M2) /\ + (forall s m M1, + {{ Rne m in s ↘ M1 }} -> + forall M2, + {{ Rne m in s ↘ M2 }} -> + M1 = M2) /\ + (forall s m M1, + {{ Rtyp m in s ↘ M1 }} -> + forall M2, + {{ Rtyp m in s ↘ M2 }} -> + M1 = M2). Proof with (functional_eval_rewrite_clear; f_equal; solve [eauto]) using. - intros *. - induction Hread1 - using read_nf_mut_ind - with (P0 := functional_read_ne_prop) - (P1 := functional_read_typ_prop); - unfold_functional_read; - inversion_clear 1... + apply read_mut_ind; intros; progressive_inversion... Qed. - Lemma functional_read_ne : forall s m M1 (Hread1 : {{ Rne m in s ↘ M1 }}), functional_read_ne_prop s m M1 Hread1. - Proof with (functional_eval_rewrite_clear; f_equal; solve [eauto]) using. - intros *. - induction Hread1 - using read_ne_mut_ind - with (P := functional_read_nf_prop) - (P1 := functional_read_typ_prop); - unfold_functional_read; - inversion_clear 1... + Corollary functional_read_nf : forall s m M1 M2, + {{ Rnf m in s ↘ M1 }} -> + {{ Rnf m in s ↘ M2 }} -> + M1 = M2. + Proof. + pose proof functional_read; firstorder. Qed. - Lemma functional_read_typ : forall s m M1 (Hread1 : {{ Rtyp m in s ↘ M1 }}), functional_read_typ_prop s m M1 Hread1. - Proof with (functional_eval_rewrite_clear; f_equal; solve [eauto]) using. - intros *. - induction Hread1 - using read_typ_mut_ind - with (P := functional_read_nf_prop) - (P0 := functional_read_ne_prop); - unfold_functional_read; - inversion_clear 1... + Lemma functional_read_ne : forall s m M1 M2, + {{ Rne m in s ↘ M1 }} -> + {{ Rne m in s ↘ M2 }} -> + M1 = M2. + Proof. + pose proof functional_read; firstorder. + Qed. + + Lemma functional_read_typ : forall s m M1 M2, + {{ Rtyp m in s ↘ M1 }} -> + {{ Rtyp m in s ↘ M2 }} -> + M1 = M2. + Proof. + pose proof functional_read; firstorder. Qed. End functional_read. diff --git a/theories/Core/Semantic/Realizability.v b/theories/Core/Semantic/Realizability.v index 40d8c753..f1b53320 100644 --- a/theories/Core/Semantic/Realizability.v +++ b/theories/Core/Semantic/Realizability.v @@ -37,7 +37,7 @@ Proof with (solve [try (try (eexists; split); econstructor); mauto]). specialize (H1 s) as [? []]... - destruct IHHunivelem as [? []]. intro s. - assert {{ Dom ⇑! A s ≈ ⇑! A' s ∈ in_rel }} by eauto using var_per_bot. + assert {{ Dom ⇑! a s ≈ ⇑! a' s ∈ in_rel }} by eauto using var_per_bot. destruct_rel_mod_eval. specialize (H9 (S s)) as [? []]. specialize (H2 s) as [? []]... @@ -45,16 +45,26 @@ Proof with (solve [try (try (eexists; split); econstructor); mauto]). destruct_conjs. destruct_rel_mod_eval. econstructor; try solve [econstructor; eauto]. - enough ({{ Dom c ⇓ A c0 ≈ c' ⇓ A' c0' ∈ per_bot }}) by eauto. + enough ({{ Dom c ⇓ a c0 ≈ c' ⇓ a' c0' ∈ per_bot }}) by eauto. intro s. specialize (H3 s) as [? []]. specialize (H5 _ _ equiv_c0_c0' s) as [? []]... - destruct_conjs. intro s. - assert {{ Dom ⇑! A s ≈ ⇑! A' s ∈ in_rel }} by eauto using var_per_bot. + assert {{ Dom ⇑! a s ≈ ⇑! a' s ∈ in_rel }} by eauto using var_per_bot. destruct_rel_mod_eval. destruct_rel_mod_app. - assert {{ Dom ⇓ a fa ≈ ⇓ a' f'a' ∈ per_top }} by eauto. + match goal with + | _: {{ $| ~?f0 & ⇑! a s |↘ ~_ }}, + _: {{ $| ~?f0' & ⇑! a' s |↘ ~_ }}, + _: {{ ⟦ B ⟧ ρ ↦ ⇑! a s ↘ ~?b0 }}, + _: {{ ⟦ B' ⟧ ρ' ↦ ⇑! a' s ↘ ~?b0' }} |- _ => + rename f0 into f; + rename f0' into f'; + rename b0 into b; + rename b0' into b' + end. + assert {{ Dom ⇓ b fa ≈ ⇓ b' f'a' ∈ per_top }} by eauto. specialize (H2 s) as [? []]. specialize (H16 (S s)) as [? []]... - intro s. @@ -97,7 +107,7 @@ Hint Resolve per_elem_then_per_top : mcltt. Lemma per_ctx_then_per_env_initial_env : forall {Γ Γ' env_rel}, {{ EF Γ ≈ Γ' ∈ per_ctx_env ↘ env_rel }} -> - exists p p', initial_env Γ p /\ initial_env Γ' p' /\ {{ Dom p ≈ p' ∈ env_rel }}. + exists ρ ρ', initial_env Γ ρ /\ initial_env Γ' ρ' /\ {{ Dom ρ ≈ ρ' ∈ env_rel }}. Proof. induction 1. - do 2 eexists; intuition. @@ -111,9 +121,9 @@ Proof. eexists; eauto. Qed. -Lemma var_per_elem : forall {A B i R} n, - {{ DF A ≈ B ∈ per_univ_elem i ↘ R }} -> - {{ Dom ⇑! A n ≈ ⇑! B n ∈ R }}. +Lemma var_per_elem : forall {a b i R} n, + {{ DF a ≈ b ∈ per_univ_elem i ↘ R }} -> + {{ Dom ⇑! a n ≈ ⇑! b n ∈ R }}. Proof. intros. eapply per_bot_then_per_elem; mauto. diff --git a/theories/Core/Soundness/FunctionCases.v b/theories/Core/Soundness/FunctionCases.v index c88d2884..6fe838af 100644 --- a/theories/Core/Soundness/FunctionCases.v +++ b/theories/Core/Soundness/FunctionCases.v @@ -77,18 +77,19 @@ Proof. handle_functional_glu_univ_elem. autorewrite with mcltt. eassumption. - - rename a0 into c. - rename equiv_a into equiv_c. - assert {{ Dom ρ ↦ c ≈ ρ ↦ c ∈ env_relΓA }} as HrelΓA by (apply_relation_equivalence; mautosolve 2). + - assert {{ Dom ρ ↦ m ≈ ρ ↦ m ∈ env_relΓA }} as HrelΓA by (apply_relation_equivalence; mautosolve 2). apply_relation_equivalence. (on_all_hyp: fun H => destruct (H _ _ HrelΓA)). destruct_by_head per_univ. functional_eval_rewrite_clear. - rename m0 into b. - assert {{ Δ' ⊢ m : A[σ][τ] }} by mauto 3 using glu_univ_elem_trm_escape. - assert {{ DG b ∈ glu_univ_elem i ↘ OP c equiv_c ↘ OEl c equiv_c }} by mauto 2. + match goal with + | _: {{ ⟦ B ⟧ ρ ↦ m ↘ ~?a }} |- _ => + rename a into b + end. + assert {{ Δ' ⊢ M : A[σ][τ] }} by mauto 3 using glu_univ_elem_trm_escape. + assert {{ DG b ∈ glu_univ_elem i ↘ OP m equiv_m ↘ OEl m equiv_m }} by mauto 2. erewrite <- @sub_decompose_q_typ; mauto 3. - assert {{ Δ' ⊢s (σ∘τ),,m ® ρ ↦ c ∈ cons_glu_sub_pred i Γ A SbΓ }} as Hconspred by mauto 2. + assert {{ Δ' ⊢s (σ∘τ),,M ® ρ ↦ m ∈ cons_glu_sub_pred i Γ A SbΓ }} as Hconspred by mauto 2. (on_all_hyp: fun H => destruct (H _ _ _ Hconspred)). simplify_evals. match_by_head glu_univ_elem ltac:(fun H => directed invert_glu_univ_elem H). @@ -184,23 +185,17 @@ Proof. handle_per_univ_elem_irrel. econstructor; mauto 3. - eapply glu_univ_elem_typ_monotone; mauto 3. - - match goal with - | H: {{ Dom ~?a ≈ ~?a ∈ in_rel }}, _: {{ ~_ ⊢ ~?M : ~_ ® ~?a ∈ Ela }} |- _ => - rename a into c; - rename H into equiv_c; - rename M into N - end. - assert {{ Dom ρ ↦ c ≈ ρ ↦ c ∈ env_relΓA }} as HrelΓA by (apply_relation_equivalence; mautosolve 2). + - assert {{ Dom ρ ↦ n ≈ ρ ↦ n ∈ env_relΓA }} as HrelΓA by (apply_relation_equivalence; mautosolve 2). destruct_rel_mod_eval. apply_relation_equivalence. (on_all_hyp: fun H => destruct (H _ _ HrelΓA) as [? [[] []]]). handle_per_univ_elem_irrel. eexists; split; mauto 3. match goal with - | _: {{ ⟦ B ⟧ ρ ↦ c ↘ ~?a }} |- _ => + | _: {{ ⟦ B ⟧ ρ ↦ n ↘ ~?a }} |- _ => rename a into b end. - assert {{ DG b ∈ glu_univ_elem i ↘ OP c equiv_c ↘ OEl c equiv_c }} by mauto 3. + assert {{ DG b ∈ glu_univ_elem i ↘ OP n equiv_n ↘ OEl n equiv_n }} by mauto 3. assert {{ Δ0 ⊢s σ0 : Δ }} by mauto 3. assert {{ Δ0 ⊢ N : A[σ][σ0] }} by mauto 2 using glu_univ_elem_trm_escape. erewrite <- @sub_decompose_q_typ; mauto 2. @@ -221,8 +216,8 @@ Proof. assert {{ Δ0 ⊢ B[q (σ∘σ0)∘(Id,,N)] ≈ B[σ∘σ0,,N] : Type@i }} as <- by mauto 2. transitivity {{{ M[q (σ∘σ0)∘(Id,,N)] }}}; mauto 3. } - assert {{ Δ0 ⊢ N : A[σ][σ0] ® c ∈ Ela }} by mauto 3. - assert {{ Δ0 ⊢s (σ∘σ0),,N ® ρ ↦ c ∈ SbΓA }} as HSbΓA by (unfold SbΓA; mauto 2). + assert {{ Δ0 ⊢ N : A[σ][σ0] ® n ∈ Ela }} by mauto 3. + assert {{ Δ0 ⊢s (σ∘σ0),,N ® ρ ↦ n ∈ SbΓA }} as HSbΓA by (unfold SbΓA; mauto 2). (on_all_hyp: fun H => destruct (H _ _ _ HSbΓA)). simplify_evals. match_by_head glu_univ_elem ltac:(fun H => directed invert_glu_univ_elem H). @@ -284,15 +279,23 @@ Proof. handle_functional_glu_univ_elem. match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). inversion_clear_by_head pi_glu_exp_pred. - rename m0 into a. - rename m1 into n. + match goal with + | _: {{ ⟦ A ⟧ ρ ↘ ~?a' }}, + _: {{ ⟦ N ⟧ ρ ↘ ~?n' }} |- _ => + rename a' into a; + rename n' into n + end. assert {{ Dom a ≈ a ∈ per_univ i }} as [] by mauto 3. handle_per_univ_elem_irrel. assert {{ Dom n ≈ n ∈ in_rel }} by (eapply glu_univ_elem_per_elem; revgoals; eassumption). (on_all_hyp: destruct_rel_by_assumption in_rel). simplify_evals. - rename a0 into b. - rename fa into mn. + match goal with + | _: {{ ⟦ B ⟧ ρ ↦ n ↘ ~?b' }}, + _: {{ $| m & n |↘ ~?mn' }} |- _ => + rename b' into b; + rename mn' into mn + end. eapply mk_glu_rel_exp_with_sub''; mauto 3. intros. match_by_head1 (in_rel n n) ltac:(fun H => rename H into equiv_n). diff --git a/theories/Core/Soundness/LogicalRelation.v b/theories/Core/Soundness/LogicalRelation.v index ffffd117..5665881b 100644 --- a/theories/Core/Soundness/LogicalRelation.v +++ b/theories/Core/Soundness/LogicalRelation.v @@ -1 +1 @@ -From Mcltt.Core.Soundness Require Export LogicalRelation.Core LogicalRelation.Lemmas. +From Mcltt.Core.Soundness.LogicalRelation Require Export Core Lemmas. diff --git a/theories/Core/Soundness/LogicalRelation/Core.v b/theories/Core/Soundness/LogicalRelation/Core.v index 96b8b8b5..6756faba 100644 --- a/theories/Core/Soundness/LogicalRelation/Core.v +++ b/theories/Core/Soundness/LogicalRelation/Core.v @@ -1 +1 @@ -From Mcltt.Core.Soundness Require Export LogicalRelation.Definitions LogicalRelation.CoreLemmas. +From Mcltt.Core.Soundness.LogicalRelation Require Export CoreLemmas CoreTactics Definitions. diff --git a/theories/Core/Soundness/LogicalRelation/CoreLemmas.v b/theories/Core/Soundness/LogicalRelation/CoreLemmas.v index 30f23e87..976d2247 100644 --- a/theories/Core/Soundness/LogicalRelation/CoreLemmas.v +++ b/theories/Core/Soundness/LogicalRelation/CoreLemmas.v @@ -6,8 +6,8 @@ From Mcltt.Core.Soundness Require Import LogicalRelation.Definitions LogicalRela From Mcltt.Core.Soundness Require Export Weakening.Lemmas. Import Domain_Notations. -Lemma glu_nat_per_nat : forall Γ m a, - glu_nat Γ m a -> +Lemma glu_nat_per_nat : forall Γ M a, + glu_nat Γ M a -> {{ Dom a ≈ a ∈ per_nat }}. Proof. induction 1; mauto. @@ -16,19 +16,19 @@ Qed. #[local] Hint Resolve glu_nat_per_nat : mcltt. -Lemma glu_nat_escape : forall Γ m a, - glu_nat Γ m a -> +Lemma glu_nat_escape : forall Γ M a, + glu_nat Γ M a -> {{ ⊢ Γ }} -> - {{ Γ ⊢ m : ℕ }}. + {{ Γ ⊢ M : ℕ }}. Proof. induction 1; intros; try match goal with | H : _ |- _ => solve [gen_presup H; mauto] end. assert {{ Γ ⊢w Id : Γ }} by mauto. - match_by_head (per_bot c c) ltac:(fun H => specialize (H (length Γ)) as [L []]). + match_by_head (per_bot m m) ltac:(fun H => specialize (H (length Γ)) as [M' []]). clear_dups. - assert {{ Γ ⊢ m[Id] ≈ L : ℕ }} by mauto. + assert {{ Γ ⊢ M[Id] ≈ M' : ℕ }} by mauto. gen_presups. mauto. Qed. @@ -36,10 +36,10 @@ Qed. #[export] Hint Resolve glu_nat_escape : mcltt. -Lemma glu_nat_resp_ctx_eq : forall Γ m a Δ, - glu_nat Γ m a -> +Lemma glu_nat_resp_ctx_eq : forall Γ M a Δ, + glu_nat Γ M a -> {{ ⊢ Γ ≈ Δ }} -> - glu_nat Δ m a. + glu_nat Δ M a. Proof. induction 1; intros; mauto. Qed. @@ -53,16 +53,16 @@ Proof. split; mauto using glu_nat_resp_ctx_eq. Qed. -Lemma glu_nat_resp_exp_eq : forall Γ m a, - glu_nat Γ m a -> - forall m', - {{ Γ ⊢ m ≈ m' : ℕ }} -> - glu_nat Γ m' a. +Lemma glu_nat_resp_exp_eq : forall Γ M a, + glu_nat Γ M a -> + forall M', + {{ Γ ⊢ M ≈ M' : ℕ }} -> + glu_nat Γ M' a. Proof. induction 1; intros; mauto. econstructor; trivial. intros. - transitivity {{{ m[σ] }}}; mauto. + transitivity {{{ M[σ] }}}; mauto. Qed. #[local] @@ -74,18 +74,18 @@ Proof. split; mauto using glu_nat_resp_exp_eq. Qed. -Lemma glu_nat_readback : forall Γ m a, - glu_nat Γ m a -> - forall Δ σ b, +Lemma glu_nat_readback : forall Γ M a, + glu_nat Γ M a -> + forall Δ σ M', {{ Δ ⊢w σ : Γ }} -> - {{ Rnf ⇓ ℕ a in length Δ ↘ b }} -> - {{ Δ ⊢ m [ σ ] ≈ b : ℕ }}. + {{ Rnf ⇓ ℕ a in length Δ ↘ M' }} -> + {{ Δ ⊢ M[σ] ≈ M' : ℕ }}. Proof. induction 1; intros; progressive_inversion; gen_presups. - transitivity {{{ zero[σ] }}}; mauto. - - assert {{ Δ ⊢ m'[σ] ≈ M : ℕ }} by mauto. - transitivity {{{ (succ m')[σ] }}}; mauto 3. - transitivity {{{ succ m'[σ] }}}; mauto 4. + - assert {{ Δ ⊢ M'[σ] ≈ M0 : ℕ }} by mauto. + transitivity {{{ (succ M')[σ] }}}; mauto 3. + transitivity {{{ succ M'[σ] }}}; mauto 4. - mauto. Qed. @@ -97,23 +97,23 @@ Ltac simpl_glu_rel := destruct_all; gen_presups. -Lemma glu_univ_elem_univ_lvl : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ T, - {{ Γ ⊢ T ® P }} -> - {{ Γ ⊢ T : Type@i }}. +Lemma glu_univ_elem_univ_lvl : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ A, + {{ Γ ⊢ A ® P }} -> + {{ Γ ⊢ A : Type@i }}. Proof. simpl. induction 1 using glu_univ_elem_ind; intros; simpl_glu_rel; trivial. Qed. -Lemma glu_univ_elem_typ_resp_exp_eq : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ T T', - {{ Γ ⊢ T ® P }} -> - {{ Γ ⊢ T ≈ T' : Type@i }} -> - {{ Γ ⊢ T' ® P }}. +Lemma glu_univ_elem_typ_resp_exp_eq : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ A A', + {{ Γ ⊢ A ® P }} -> + {{ Γ ⊢ A ≈ A' : Type@i }} -> + {{ Γ ⊢ A' ® P }}. Proof. simpl. induction 1 using glu_univ_elem_ind; intros; @@ -121,29 +121,31 @@ Proof. split; [trivial |]. intros. - assert {{ Δ ⊢ T[σ] ≈ V : Type@i }}; mauto. + transitivity {{{ A[σ] }}}; mauto 4. Qed. -Add Parametric Morphism i P El A (H : glu_univ_elem i P El A) Γ : (P Γ) - with signature wf_exp_eq Γ {{{Type@i}}} ==> iff as glu_univ_elem_typ_morphism_iff1. +Add Parametric Morphism i P El a (H : glu_univ_elem i P El a) Γ : (P Γ) + with signature wf_exp_eq Γ {{{ Type@i }}} ==> iff as glu_univ_elem_typ_morphism_iff1. Proof. split; intros; eapply glu_univ_elem_typ_resp_exp_eq; mauto 2. Qed. -Lemma glu_univ_elem_trm_resp_typ_exp_eq : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ t T a T', - {{ Γ ⊢ t : T ® a ∈ El }} -> - {{ Γ ⊢ T ≈ T' : Type@i }} -> - {{ Γ ⊢ t : T' ® a ∈ El }}. +Lemma glu_univ_elem_trm_resp_typ_exp_eq : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ M A m A', + {{ Γ ⊢ M : A ® m ∈ El }} -> + {{ Γ ⊢ A ≈ A' : Type@i }} -> + {{ Γ ⊢ M : A' ® m ∈ El }}. Proof. simpl. induction 1 using glu_univ_elem_ind; intros; simpl_glu_rel; repeat split; intros; mauto 3; - [firstorder | | assert {{ Δ ⊢ M[σ] ≈ V : Type@i }} by mauto 3 |]; mauto. + [firstorder | | transitivity {{{ A[σ] }}}; mauto 4 | assert {{ Δ ⊢ A[σ] ≈ A'[σ] : Type@i }}; mauto 3]. + + econstructor; mauto 3. Qed. -Add Parametric Morphism i P El A (H : glu_univ_elem i P El A) Γ : (El Γ) +Add Parametric Morphism i P El a (H : glu_univ_elem i P El a) Γ : (El Γ) with signature wf_exp_eq Γ {{{Type@i}}} ==> eq ==> eq ==> iff as glu_univ_elem_trm_morphism_iff1. Proof. split; intros; @@ -151,12 +153,12 @@ Proof. mauto 2. Qed. -Lemma glu_univ_elem_typ_resp_ctx_eq : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ T Δ, - {{ Γ ⊢ T ® P }} -> +Lemma glu_univ_elem_typ_resp_ctx_eq : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ A Δ, + {{ Γ ⊢ A ® P }} -> {{ ⊢ Γ ≈ Δ }} -> - {{ Δ ⊢ T ® P }}. + {{ Δ ⊢ A ® P }}. Proof. simpl. induction 1 using glu_univ_elem_ind; intros; @@ -164,7 +166,7 @@ Proof. econstructor; mauto 4. Qed. -Add Parametric Morphism i P El A (H : glu_univ_elem i P El A) : P +Add Parametric Morphism i P El a (H : glu_univ_elem i P El a) : P with signature wf_ctx_eq ==> eq ==> iff as glu_univ_elem_typ_morphism_iff2. Proof. intros. split; intros; @@ -172,12 +174,12 @@ Proof. mauto 2. Qed. -Lemma glu_univ_elem_trm_resp_ctx_eq : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ T M m Δ, - {{ Γ ⊢ M : T ® m ∈ El }} -> +Lemma glu_univ_elem_trm_resp_ctx_eq : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ A M m Δ, + {{ Γ ⊢ M : A ® m ∈ El }} -> {{ ⊢ Γ ≈ Δ }} -> - {{ Δ ⊢ M : T ® m ∈ El }}. + {{ Δ ⊢ M : A ® m ∈ El }}. Proof. simpl. induction 1 using glu_univ_elem_ind; intros; @@ -191,7 +193,7 @@ Proof. - split; mauto 4. Qed. -Add Parametric Morphism i P El A (H : glu_univ_elem i P El A) : El +Add Parametric Morphism i P El a (H : glu_univ_elem i P El a) : El with signature wf_ctx_eq ==> eq ==> eq ==> eq ==> iff as glu_univ_elem_trm_morphism_iff2. Proof. intros. split; intros; @@ -199,30 +201,31 @@ Proof. mauto 2. Qed. -Lemma glu_nat_resp_wk' : forall Γ m a, - glu_nat Γ m a -> +Lemma glu_nat_resp_wk' : forall Γ M a, + glu_nat Γ M a -> forall Δ σ, - {{ Γ ⊢ m : ℕ }} -> + {{ Γ ⊢ M : ℕ }} -> {{ Δ ⊢w σ : Γ }} -> - glu_nat Δ {{{ m[σ] }}} a. + glu_nat Δ {{{ M[σ] }}} a. Proof. induction 1; intros; gen_presups. - econstructor. transitivity {{{ zero[σ] }}}; mauto. - econstructor; [ |mauto]. - transitivity {{{ (succ m')[σ] }}}; mauto 4. + transitivity {{{ (succ M')[σ] }}}; mauto 4. - econstructor; trivial. - intros. gen_presups. - assert {{ Δ0 ⊢w σ ∘ σ0 : Γ }} by mauto. - assert {{ Δ0 ⊢ m[σ ∘ σ0] ≈ v : ℕ }} by mauto 4. - transitivity {{{ m[σ ∘ σ0] }}}; mauto 4. + intros Δ' τ M' **. + gen_presups. + assert {{ Δ' ⊢w σ∘τ : Γ }} by mauto. + assert {{ Δ' ⊢ M[σ∘τ] ≈ M' : ℕ }} by mauto 4. + transitivity {{{ M[σ∘τ] }}}; mauto 4. Qed. -Lemma glu_nat_resp_wk : forall Γ m a, - glu_nat Γ m a -> +Lemma glu_nat_resp_wk : forall Γ M a, + glu_nat Γ M a -> forall Δ σ, {{ Δ ⊢w σ : Γ }} -> - glu_nat Δ {{{ m[σ] }}} a. + glu_nat Δ {{{ M[σ] }}} a. Proof. mauto using glu_nat_resp_wk'. Qed. @@ -230,20 +233,20 @@ Qed. #[export] Hint Resolve glu_nat_resp_wk : mcltt. -Lemma glu_univ_elem_trm_escape : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ t T a, - {{ Γ ⊢ t : T ® a ∈ El }} -> - {{ Γ ⊢ t : T }}. +Lemma glu_univ_elem_trm_escape : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ M A m, + {{ Γ ⊢ M : A ® m ∈ El }} -> + {{ Γ ⊢ M : A }}. Proof. simpl. induction 1 using glu_univ_elem_ind; intros; simpl_glu_rel; mauto 4. Qed. -Lemma glu_univ_elem_per_univ : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - {{ Dom A ≈ A ∈ per_univ i }}. +Lemma glu_univ_elem_per_univ : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + {{ Dom a ≈ a ∈ per_univ i }}. Proof. simpl. induction 1 using glu_univ_elem_ind; intros; eexists; @@ -258,18 +261,18 @@ Qed. #[export] Hint Resolve glu_univ_elem_per_univ : mcltt. -Lemma glu_univ_elem_per_elem : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ t T a R, - {{ Γ ⊢ t : T ® a ∈ El }} -> - {{ DF A ≈ A ∈ per_univ_elem i ↘ R }} -> - {{ Dom a ≈ a ∈ R }}. +Lemma glu_univ_elem_per_elem : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ M A m R, + {{ Γ ⊢ M : A ® m ∈ El }} -> + {{ DF a ≈ a ∈ per_univ_elem i ↘ R }} -> + {{ Dom m ≈ m ∈ R }}. Proof. simpl. induction 1 using glu_univ_elem_ind; intros; try do 2 match_by_head1 per_univ_elem invert_per_univ_elem; simpl_glu_rel; - try fold (per_univ j a a); + try fold (per_univ j m m); mauto 4. intros. @@ -281,11 +284,11 @@ Proof. econstructor; firstorder eauto. Qed. -Lemma glu_univ_elem_trm_typ : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ t T a, - {{ Γ ⊢ t : T ® a ∈ El }} -> - {{ Γ ⊢ T ® P }}. +Lemma glu_univ_elem_trm_typ : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ M A m, + {{ Γ ⊢ M : A ® m ∈ El }} -> + {{ Γ ⊢ A ® P }}. Proof. simpl. induction 1 using glu_univ_elem_ind; intros; @@ -294,26 +297,26 @@ Proof. econstructor; eauto. match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). - intros. + intros ? ? N n ? ? equiv_n. destruct_rel_mod_eval. - edestruct H12 as [? []]; eauto. + enough (exists mn : domain, {{ $| m & n |↘ mn }} /\ {{ Δ ⊢ M[σ] N : OT[σ,,N] ® mn ∈ OEl n equiv_n }}) as [? []]; eauto 3. Qed. -Lemma glu_univ_elem_trm_univ_lvl : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ t T a, - {{ Γ ⊢ t : T ® a ∈ El }} -> - {{ Γ ⊢ T : Type@i }}. +Lemma glu_univ_elem_trm_univ_lvl : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ M A m, + {{ Γ ⊢ M : A ® m ∈ El }} -> + {{ Γ ⊢ A : Type@i }}. Proof. intros. eapply glu_univ_elem_univ_lvl; [| eapply glu_univ_elem_trm_typ]; eassumption. Qed. -Lemma glu_univ_elem_trm_resp_exp_eq : forall i P El A, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ t T a t', - {{ Γ ⊢ t : T ® a ∈ El }} -> - {{ Γ ⊢ t ≈ t' : T }} -> - {{ Γ ⊢ t' : T ® a ∈ El }}. +Lemma glu_univ_elem_trm_resp_exp_eq : forall i P El a, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ A M m M', + {{ Γ ⊢ M : A ® m ∈ El }} -> + {{ Γ ⊢ M ≈ M' : A }} -> + {{ Γ ⊢ M' : A ® m ∈ El }}. Proof. simpl. induction 1 using glu_univ_elem_ind; intros; @@ -324,26 +327,25 @@ Proof. eapply glu_univ_elem_typ_resp_exp_eq; mauto. - econstructor; eauto. - invert_per_univ_elem H3. + match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). intros. destruct_rel_mod_eval. - assert {{ Δ ⊢ m' : IT [σ]}} by eauto using glu_univ_elem_trm_escape. - edestruct H13 as [? []]; eauto. + assert {{ Δ ⊢ N : IT[σ]}} by eauto using glu_univ_elem_trm_escape. + assert (exists mn : domain, {{ $| m & n |↘ mn }} /\ {{ Δ ⊢ M[σ] N : OT[σ,,N] ® mn ∈ OEl n equiv_n }}) as [? []] by intuition. eexists; split; eauto. - enough {{ Δ ⊢ m[σ] m' ≈ t'[σ] m' : OT[σ,,m'] }} by eauto. - assert {{ Γ ⊢ m ≈ t' : Π IT OT }} as Hty by mauto. + enough {{ Δ ⊢ M[σ] N ≈ M'[σ] N : OT[σ,,N] }} by eauto. + assert {{ Γ ⊢ M ≈ M' : Π IT OT }} as Hty by mauto. eassert {{ Δ ⊢ IT[σ] : Type@_ }} by mauto 3. eapply wf_exp_eq_sub_cong with (Γ := Δ) in Hty; [| eapply sub_eq_refl; mauto 3]. autorewrite with mcltt in Hty. - eapply wf_exp_eq_app_cong with (N := m') (N' := m') in Hty; try pi_univ_level_tac; [|mauto 2]. + eapply wf_exp_eq_app_cong with (N := N) (N' := N) in Hty; try pi_univ_level_tac; [|mauto 2]. autorewrite with mcltt in Hty. eassumption. - intros. - assert {{ Δ ⊢ m[σ] ≈ t'[σ] : M[σ] }} by mauto 4. - mauto 4. + enough {{ Δ ⊢ M[σ] ≈ M'[σ] : A[σ] }}; mauto 4. Qed. -Add Parametric Morphism i P El A (H : glu_univ_elem i P El A) Γ T : (El Γ T) +Add Parametric Morphism i P El a (H : glu_univ_elem i P El a) Γ T : (El Γ T) with signature wf_exp_eq Γ T ==> eq ==> iff as glu_univ_elem_trm_morphism_iff3. Proof. split; intros; @@ -412,7 +414,7 @@ Ltac apply_predicate_equivalence := Add Parametric Morphism i : (glu_univ_elem i) with signature glu_typ_pred_equivalence ==> glu_exp_pred_equivalence ==> eq ==> iff as simple_glu_univ_elem_morphism_iff. Proof with mautosolve. - intros P P' HPP' El El' HElEl' a. + intros P P' ? El El'. split; intro Horig; [gen El' P' | gen El P]; induction Horig using glu_univ_elem_ind; unshelve glu_univ_elem_econstructor; try (etransitivity; [symmetry + idtac|]; eassumption); eauto. @@ -489,31 +491,28 @@ Proof. handle_per_univ_elem_irrel. split; [intros Γ C | intros Γ M C m]. - split; intros []; econstructor; intuition; - rename a0 into c; - [rename equiv_a into equiv_c; assert (equiv_c' : in_rel c c) by intuition - | rename equiv_a into equiv_c'; assert (equiv_c : in_rel0 c c) by intuition]; + [rename equiv_m into equiv0_m; assert (equiv_m : in_rel m m) by intuition + | assert (equiv0_m : in_rel0 m m) by intuition ]; destruct_rel_mod_eval; - rename a0 into b; - assert ((OP c equiv_c' <∙> OP0 c equiv_c) /\ (OEl c equiv_c' <∙> OEl0 c equiv_c)) as [] by mauto 3; + functional_eval_rewrite_clear; + assert ((OP m equiv_m <∙> OP0 m equiv0_m) /\ (OEl m equiv_m <∙> OEl0 m equiv0_m)) as [] by mauto 3; intuition. - split; intros []; econstructor; intuition; - clear m; - rename b into m; - [rename equiv_b into equiv_m; assert (equiv_m' : in_rel m m) by intuition - | rename equiv_b into equiv_m'; assert (equiv_m : in_rel0 m m) by intuition]; + [rename equiv_n into equiv0_n; assert (equiv_n : in_rel n n) by intuition + | assert (equiv0_n : in_rel0 n n) by intuition]; destruct_rel_mod_eval; - [assert (exists am : domain, {{ $| a0 & m |↘ am }} /\ {{ Δ ⊢ m0[σ] m' : OT[σ,,m'] ® am ∈ OEl m equiv_m' }}) by intuition - | assert (exists am : domain, {{ $| a0 & m |↘ am }} /\ {{ Δ ⊢ m0[σ] m' : OT[σ,,m'] ® am ∈ OEl0 m equiv_m }}) by intuition]; + [assert (exists m0n, {{ $| m0 & n |↘ m0n }} /\ {{ Δ ⊢ M0[σ] N : OT[σ,,N] ® m0n ∈ OEl n equiv_n }}) by intuition + | assert (exists m0n, {{ $| m0 & n |↘ m0n }} /\ {{ Δ ⊢ M0[σ] N : OT[σ,,N] ® m0n ∈ OEl0 n equiv0_n }}) by intuition]; destruct_conjs; - assert ((OP m equiv_m' <∙> OP0 m equiv_m) /\ (OEl m equiv_m' <∙> OEl0 m equiv_m)) as [] by mauto 3; + assert ((OP n equiv_n <∙> OP0 n equiv0_n) /\ (OEl n equiv_n <∙> OEl0 n equiv0_n)) as [] by mauto 3; eexists; split; intuition. Qed. Ltac apply_functional_glu_univ_elem1 := let tactic_error o1 o2 := fail 2 "functional_glu_univ_elem biconditional between" o1 "and" o2 "cannot be solved" in match goal with - | H1 : {{ DG ~?A ∈ glu_univ_elem ?i ↘ ?P1 ↘ ?El1 }}, - H2 : {{ DG ~?A ∈ glu_univ_elem ?i' ↘ ?P2 ↘ ?El2 }} |- _ => + | H1 : {{ DG ~?a ∈ glu_univ_elem ?i ↘ ?P1 ↘ ?El1 }}, + H2 : {{ DG ~?a ∈ glu_univ_elem ?i' ↘ ?P2 ↘ ?El2 }} |- _ => unify i i'; assert_fails (unify P1 P2; unify El1 El2); match goal with @@ -536,18 +535,18 @@ Ltac handle_functional_glu_univ_elem := apply_predicate_equivalence; clear_dups. -Lemma glu_univ_elem_pi_clean_inversion1 : forall {i a p B in_rel typ_rel el_rel}, +Lemma glu_univ_elem_pi_clean_inversion1 : forall {i a ρ B in_rel P El}, {{ DF a ≈ a ∈ per_univ_elem i ↘ in_rel }} -> - {{ DG Π a p B ∈ glu_univ_elem i ↘ typ_rel ↘ el_rel }} -> + {{ DG Π a ρ B ∈ glu_univ_elem i ↘ P ↘ El }} -> exists IP IEl (OP : forall c (equiv_c_c : {{ Dom c ≈ c ∈ in_rel }}), glu_typ_pred) (OEl : forall c (equiv_c_c : {{ Dom c ≈ c ∈ in_rel }}), glu_exp_pred) elem_rel, {{ DG a ∈ glu_univ_elem i ↘ IP ↘ IEl }} /\ (forall c (equiv_c : {{ Dom c ≈ c ∈ in_rel }}) b, - {{ ⟦ B ⟧ p ↦ c ↘ b }} -> + {{ ⟦ B ⟧ ρ ↦ c ↘ b }} -> {{ DG b ∈ glu_univ_elem i ↘ OP _ equiv_c ↘ OEl _ equiv_c }}) /\ - {{ DF Π a p B ≈ Π a p B ∈ per_univ_elem i ↘ elem_rel }} /\ - (typ_rel <∙> pi_glu_typ_pred i in_rel IP IEl OP) /\ - (el_rel <∙> pi_glu_exp_pred i in_rel IP IEl elem_rel OEl). + {{ DF Π a ρ B ≈ Π a ρ B ∈ per_univ_elem i ↘ elem_rel }} /\ + (P <∙> pi_glu_typ_pred i in_rel IP IEl OP) /\ + (El <∙> pi_glu_exp_pred i in_rel IP IEl elem_rel OEl). Proof. intros *. simpl. @@ -559,11 +558,11 @@ Proof. repeat split. 1,3: eassumption. 1: instantiate (1 := fun c equiv_c Γ A M m => forall (b : domain) Pb Elb, - {{ ⟦ B ⟧ p ↦ c ↘ b }} -> + {{ ⟦ B ⟧ ρ ↦ c ↘ b }} -> {{ DG b ∈ glu_univ_elem i ↘ Pb ↘ Elb }} -> {{ Γ ⊢ M : A ® m ∈ Elb }}). 1: instantiate (1 := fun c equiv_c Γ A => forall (b : domain) Pb Elb, - {{ ⟦ B ⟧ p ↦ c ↘ b }} -> + {{ ⟦ B ⟧ ρ ↦ c ↘ b }} -> {{ DG b ∈ glu_univ_elem i ↘ Pb ↘ Elb }} -> {{ Γ ⊢ A ® Pb }}). 2-5: intros []; econstructor; mauto. @@ -575,31 +574,27 @@ Proof. split; intros; handle_functional_glu_univ_elem; intuition. + intros ? ?. split; [intros; handle_functional_glu_univ_elem |]; intuition. - - rename a0 into c. - rename equiv_a into equiv_c. - assert {{ Dom c ≈ c ∈ in_rel0 }} as equiv0_c by intuition. - assert {{ DG b ∈ glu_univ_elem i ↘ OP c equiv0_c ↘ OEl c equiv0_c }} by mauto 3. + - assert {{ Dom m ≈ m ∈ in_rel0 }} as equiv0_m by intuition. + assert {{ DG b ∈ glu_univ_elem i ↘ OP m equiv0_m ↘ OEl m equiv0_m }} by mauto 3. handle_functional_glu_univ_elem. intuition. - match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). destruct_rel_mod_eval. intuition. - - rename b into c. - rename equiv_b into equiv_c. - assert {{ Dom c ≈ c ∈ in_rel0 }} as equiv0_c by intuition. - assert (exists ac : domain, {{ $| a0 & c |↘ ac }} /\ {{ Δ ⊢ m[σ] m' : OT[σ,,m'] ® ac ∈ OEl c equiv0_c }}) by mauto 3. + - assert {{ Dom n ≈ n ∈ in_rel0 }} as equiv0_n by intuition. + assert (exists mn : domain, {{ $| m & n |↘ mn }} /\ {{ Δ ⊢ M[σ] N : OT[σ,,N] ® mn ∈ OEl n equiv0_n }}) by mauto 3. destruct_conjs. eexists. intuition. - assert {{ DG b ∈ glu_univ_elem i ↘ OP c equiv0_c ↘ OEl c equiv0_c }} by mauto 3. + assert {{ DG b ∈ glu_univ_elem i ↘ OP n equiv0_n ↘ OEl n equiv0_n }} by mauto 3. handle_functional_glu_univ_elem. intuition. - - assert (exists ab : domain, - {{ $| a0 & b |↘ ab }} /\ - (forall (b0 : domain) (Pb : glu_typ_pred) (Elb : glu_exp_pred), - {{ ⟦ B ⟧ p ↦ b ↘ b0 }} -> - {{ DG b0 ∈ glu_univ_elem i ↘ Pb ↘ Elb }} -> - {{ Δ ⊢ m[σ] m' : OT[σ,,m'] ® ab ∈ Elb }})) by intuition. + - assert (exists mn : domain, + {{ $| m & n |↘ mn }} /\ + (forall (b : domain) (Pb : glu_typ_pred) (Elb : glu_exp_pred), + {{ ⟦ B ⟧ ρ ↦ n ↘ b }} -> + {{ DG b ∈ glu_univ_elem i ↘ Pb ↘ Elb }} -> + {{ Δ ⊢ M[σ] N : OT[σ,,N] ® mn ∈ Elb }})) by intuition. destruct_conjs. match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). handle_per_univ_elem_irrel. @@ -610,18 +605,18 @@ Proof. intuition. Qed. -Lemma glu_univ_elem_pi_clean_inversion2 : forall {i a p B in_rel IP IEl typ_rel el_rel}, +Lemma glu_univ_elem_pi_clean_inversion2 : forall {i a ρ B in_rel IP IEl P El}, {{ DF a ≈ a ∈ per_univ_elem i ↘ in_rel }} -> {{ DG a ∈ glu_univ_elem i ↘ IP ↘ IEl }} -> - {{ DG Π a p B ∈ glu_univ_elem i ↘ typ_rel ↘ el_rel }} -> + {{ DG Π a ρ B ∈ glu_univ_elem i ↘ P ↘ El }} -> exists (OP : forall c (equiv_c_c : {{ Dom c ≈ c ∈ in_rel }}), glu_typ_pred) (OEl : forall c (equiv_c_c : {{ Dom c ≈ c ∈ in_rel }}), glu_exp_pred) elem_rel, (forall c (equiv_c : {{ Dom c ≈ c ∈ in_rel }}) b, - {{ ⟦ B ⟧ p ↦ c ↘ b }} -> + {{ ⟦ B ⟧ ρ ↦ c ↘ b }} -> {{ DG b ∈ glu_univ_elem i ↘ OP _ equiv_c ↘ OEl _ equiv_c }}) /\ - {{ DF Π a p B ≈ Π a p B ∈ per_univ_elem i ↘ elem_rel }} /\ - (typ_rel <∙> pi_glu_typ_pred i in_rel IP IEl OP) /\ - (el_rel <∙> pi_glu_exp_pred i in_rel IP IEl elem_rel OEl). + {{ DF Π a ρ B ≈ Π a ρ B ∈ per_univ_elem i ↘ elem_rel }} /\ + (P <∙> pi_glu_typ_pred i in_rel IP IEl OP) /\ + (El <∙> pi_glu_exp_pred i in_rel IP IEl elem_rel OEl). Proof. intros *. simpl. @@ -680,7 +675,7 @@ Qed. Add Parametric Morphism i R : (glu_univ_elem i) with signature glu_typ_pred_equivalence ==> glu_exp_pred_equivalence ==> per_univ_elem i R ==> iff as glu_univ_elem_morphism_iff'. Proof with mautosolve. - intros P P' HPP' El El' HElEl' a a' Haa'. + intros P P' HPP' El El' HElEl' **. rewrite HPP', HElEl'. split; intros; eapply glu_univ_elem_resp_per_univ; mauto. symmetry; mauto. @@ -717,12 +712,12 @@ Proof. + saturate_refl; eassumption. + instantiate (1 := fun (c : domain) (equiv_c : in_rel c c) Γ A M m => forall b P El, - {{ ⟦ B' ⟧ p' ↦ c ↘ b }} -> + {{ ⟦ B' ⟧ ρ' ↦ c ↘ b }} -> glu_univ_elem i P El b -> {{ Γ ⊢ M : A ® m ∈ El }}). instantiate (1 := fun (c : domain) (equiv_c : in_rel c c) Γ A => forall b P El, - {{ ⟦ B' ⟧ p' ↦ c ↘ b }} -> + {{ ⟦ B' ⟧ ρ' ↦ c ↘ b }} -> glu_univ_elem i P El b -> {{ Γ ⊢ A ® P }}). intros. @@ -730,7 +725,7 @@ Proof. handle_per_univ_elem_irrel. rewrite simple_glu_univ_elem_morphism_iff; try (eassumption + reflexivity); split; intros; handle_functional_glu_univ_elem; intuition. - + enough {{ DF Π A p B ≈ Π A' p' B' ∈ per_univ_elem i ↘ elem_rel }} by (etransitivity; [symmetry |]; eassumption). + + enough {{ DF Π a ρ B ≈ Π a' ρ' B' ∈ per_univ_elem i ↘ elem_rel }} by (etransitivity; [symmetry |]; eassumption). per_univ_elem_econstructor; mauto. intros. (on_all_hyp: destruct_rel_by_assumption in_rel). @@ -833,23 +828,25 @@ Proof. deepexec H1 ltac:(fun H => pose proof H). autorewrite with mcltt in *. repeat eexists; eauto. - assert {{ Δ0 ⊢s σ0,, m' : Δ, ~ (a_sub IT σ) }}. { + assert {{ Δ0 ⊢s σ0,,N : Δ, ~ (a_sub IT σ) }}. + { econstructor; mauto 2. bulky_rewrite. } - assert {{Δ0 ⊢ (m [σ][σ0]) m' ≈ (m [σ ∘ σ0]) m' : OT [σ ∘ σ0,, m']}}. { + assert {{Δ0 ⊢ M[σ][σ0] N ≈ M[σ∘σ0] N : OT [σ∘σ0,,N]}}. + { rewrite <- @sub_eq_q_sigma_id_extend; mauto 4. rewrite <- @exp_eq_sub_compose_typ; mauto 3; [eapply wf_exp_eq_app_cong' |]; mauto 4. symmetry. bulky_rewrite_in H4. - assert {{ Δ0 ⊢ Π IT[σ ∘ σ0] (OT[q (σ ∘ σ0)]) ≈ (Π IT OT)[σ ∘ σ0] : Type@(S (max i4 i)) }} by mauto. + assert {{ Δ0 ⊢ Π IT[σ∘σ0] (OT[q (σ∘σ0)]) ≈ (Π IT OT)[σ∘σ0] : Type@(S (max i4 i)) }} by mauto. eapply wf_exp_eq_conv'; mauto 4. } bulky_rewrite. - edestruct H10 with (b := b) as [? []]; + edestruct H10 with (n := n) as [? []]; simplify_evals; [| | eassumption]; mauto. diff --git a/theories/Core/Soundness/LogicalRelation/CoreTactics.v b/theories/Core/Soundness/LogicalRelation/CoreTactics.v index 87da9f2c..2b85a547 100644 --- a/theories/Core/Soundness/LogicalRelation/CoreTactics.v +++ b/theories/Core/Soundness/LogicalRelation/CoreTactics.v @@ -1,5 +1,6 @@ From Equations Require Import Equations. +From Mcltt Require Import Base LibTactics. From Mcltt.Core.Soundness Require Import LogicalRelation.Definitions. Ltac basic_invert_glu_univ_elem H := @@ -12,3 +13,15 @@ Ltac basic_glu_univ_elem_econstructor := progress simp glu_univ_elem; econstructor; try rewrite <- glu_univ_elem_equation_1 in *. + +Ltac invert_glu_rel1 := + match goal with + | H : pi_glu_typ_pred _ _ _ _ _ _ _ |- _ => + progressive_invert H + | H : pi_glu_exp_pred _ _ _ _ _ _ _ _ _ _ |- _ => + progressive_invert H + | H : neut_glu_typ_pred _ _ _ _ |- _ => + progressive_invert H + | H : neut_glu_exp_pred _ _ _ _ _ _ |- _ => + progressive_invert H + end. diff --git a/theories/Core/Soundness/LogicalRelation/Definitions.v b/theories/Core/Soundness/LogicalRelation/Definitions.v index b5c44a16..9802d177 100644 --- a/theories/Core/Soundness/LogicalRelation/Definitions.v +++ b/theories/Core/Soundness/LogicalRelation/Definitions.v @@ -31,84 +31,84 @@ Notation "'EG' A ∈ R ↘ Sb " := (R Sb A : ((Prop : (Type : Type)) : (Type : T Inductive glu_nat : ctx -> exp -> domain -> Prop := | glu_nat_zero : - `{ {{ Γ ⊢ m ≈ zero : ℕ }} -> - glu_nat Γ m d{{{ zero }}} } + `{ {{ Γ ⊢ M ≈ zero : ℕ }} -> + glu_nat Γ M d{{{ zero }}} } | glu_nat_succ : - `{ {{ Γ ⊢ m ≈ succ m' : ℕ }} -> - glu_nat Γ m' a -> - glu_nat Γ m d{{{ succ a }}} } + `{ {{ Γ ⊢ M ≈ succ M' : ℕ }} -> + glu_nat Γ M' m' -> + glu_nat Γ M d{{{ succ m' }}} } | glu_nat_neut : - `{ per_bot c c -> - (forall {Δ σ v}, {{ Δ ⊢w σ : Γ }} -> {{ Rne c in length Δ ↘ v }} -> {{ Δ ⊢ m[σ] ≈ v : ℕ }}) -> - glu_nat Γ m d{{{ ⇑ ℕ c }}} }. + `{ per_bot m m -> + (forall {Δ σ M'}, {{ Δ ⊢w σ : Γ }} -> {{ Rne m in length Δ ↘ M' }} -> {{ Δ ⊢ M[σ] ≈ M' : ℕ }}) -> + glu_nat Γ M d{{{ ⇑ ℕ m }}} }. #[export] Hint Constructors glu_nat : mcltt. -Definition nat_glu_typ_pred i : glu_typ_pred := fun Γ M => {{ Γ ⊢ M ≈ ℕ : Type@i }}. -Arguments nat_glu_typ_pred i Γ M/. +Definition nat_glu_typ_pred i : glu_typ_pred := fun Γ A => {{ Γ ⊢ A ≈ ℕ : Type@i }}. +Arguments nat_glu_typ_pred i Γ A/. -Definition nat_glu_exp_pred i : glu_exp_pred := fun Γ M m a => {{ Γ ⊢ M ® nat_glu_typ_pred i }} /\ glu_nat Γ m a. -Arguments nat_glu_exp_pred i Γ m M a/. +Definition nat_glu_exp_pred i : glu_exp_pred := fun Γ A M m => {{ Γ ⊢ A ® nat_glu_typ_pred i }} /\ glu_nat Γ M m. +Arguments nat_glu_exp_pred i Γ A M m/. -Definition neut_glu_typ_pred i C : glu_typ_pred := - fun Γ M => {{ Γ ⊢ M : Type@i }} /\ - (forall Δ σ V, {{ Δ ⊢w σ : Γ }} -> {{ Rne C in length Δ ↘ V }} -> {{ Δ ⊢ M[σ] ≈ V : Type@i }}). -Arguments neut_glu_typ_pred i C Γ M/. +Definition neut_glu_typ_pred i a : glu_typ_pred := + fun Γ A => {{ Γ ⊢ A : Type@i }} /\ + (forall Δ σ A', {{ Δ ⊢w σ : Γ }} -> {{ Rne a in length Δ ↘ A' }} -> {{ Δ ⊢ A[σ] ≈ A' : Type@i }}). +Arguments neut_glu_typ_pred i a Γ A/. -Inductive neut_glu_exp_pred i C : glu_exp_pred := +Variant neut_glu_exp_pred i a : glu_exp_pred := | mk_neut_glu_exp_pred : - `{ {{ Γ ⊢ M ® neut_glu_typ_pred i C }} -> - {{ Γ ⊢ m : M }} -> - {{ Dom c ≈ c ∈ per_bot }} -> - (forall Δ σ V v, {{ Δ ⊢w σ : Γ }} -> - {{ Rne C in length Δ ↘ V }} -> - {{ Rne c in length Δ ↘ v }} -> - {{ Δ ⊢ m[σ] ≈ v : M[σ] }}) -> - {{ Γ ⊢ m : M ® ⇑ A c ∈ neut_glu_exp_pred i C }} }. - -Inductive pi_glu_typ_pred i + `{ {{ Γ ⊢ A ® neut_glu_typ_pred i a }} -> + {{ Γ ⊢ M : A }} -> + {{ Dom m ≈ m ∈ per_bot }} -> + (forall Δ σ (* A' *) M', {{ Δ ⊢w σ : Γ }} -> + (* {{ Rne a in length Δ ↘ A' }} -> *) + {{ Rne m in length Δ ↘ M' }} -> + {{ Δ ⊢ M[σ] ≈ M' : A[σ] }}) -> + {{ Γ ⊢ M : A ® ⇑ b m ∈ neut_glu_exp_pred i a }} }. + +Variant pi_glu_typ_pred i (IR : relation domain) (IP : glu_typ_pred) (IEl : glu_exp_pred) (OP : forall c (equiv_c : {{ Dom c ≈ c ∈ IR }}), glu_typ_pred) : glu_typ_pred := | mk_pi_glu_typ_pred : - `{ {{ Γ ⊢ M ≈ Π IT OT : Type@i }} -> - {{ Γ ⊢ IT : Type@i }} -> - {{ Γ , IT ⊢ OT : Type@i }} -> - (forall Δ σ, {{ Δ ⊢w σ : Γ }} -> {{ Δ ⊢ IT[σ] ® IP }}) -> - (forall Δ σ m a, - {{ Δ ⊢w σ : Γ }} -> - {{ Δ ⊢ m : IT[σ] ® a ∈ IEl }} -> - forall (equiv_a : {{ Dom a ≈ a ∈ IR }}), - {{ Δ ⊢ OT[σ,,m] ® OP _ equiv_a }}) -> - {{ Γ ⊢ M ® pi_glu_typ_pred i IR IP IEl OP }} }. - -Inductive pi_glu_exp_pred i + `{ {{ Γ ⊢ A ≈ Π IT OT : Type@i }} -> + {{ Γ ⊢ IT : Type@i }} -> + {{ Γ , IT ⊢ OT : Type@i }} -> + (forall Δ σ, {{ Δ ⊢w σ : Γ }} -> {{ Δ ⊢ IT[σ] ® IP }}) -> + (forall Δ σ M m, + {{ Δ ⊢w σ : Γ }} -> + {{ Δ ⊢ M : IT[σ] ® m ∈ IEl }} -> + forall (equiv_m : {{ Dom m ≈ m ∈ IR }}), + {{ Δ ⊢ OT[σ,,M] ® OP _ equiv_m }}) -> + {{ Γ ⊢ A ® pi_glu_typ_pred i IR IP IEl OP }} }. + +Variant pi_glu_exp_pred i (IR : relation domain) (IP : glu_typ_pred) (IEl : glu_exp_pred) (elem_rel : relation domain) (OEl : forall c (equiv_c : {{ Dom c ≈ c ∈ IR }}), glu_exp_pred): glu_exp_pred := | mk_pi_glu_exp_pred : - `{ {{ Γ ⊢ m : M }} -> - {{ Dom a ≈ a ∈ elem_rel }} -> - {{ Γ ⊢ M ≈ Π IT OT : Type@i }} -> - {{ Γ ⊢ IT : Type@i }} -> + `{ {{ Γ ⊢ M : A }} -> + {{ Dom m ≈ m ∈ elem_rel }} -> + {{ Γ ⊢ A ≈ Π IT OT : Type@i }} -> + {{ Γ ⊢ IT : Type@i }} -> {{ Γ , IT ⊢ OT : Type@i }} -> (forall Δ σ, {{ Δ ⊢w σ : Γ }} -> {{ Δ ⊢ IT[σ] ® IP }}) -> - (forall Δ σ m' b, + (forall Δ σ N n, {{ Δ ⊢w σ : Γ }} -> - {{ Δ ⊢ m' : IT[ σ ] ® b ∈ IEl }} -> - forall (equiv_b : {{ Dom b ≈ b ∈ IR }}), - exists ab, {{ $| a & b |↘ ab }} /\ {{ Δ ⊢ m[σ] m' : OT[σ,,m'] ® ab ∈ OEl _ equiv_b }}) -> - {{ Γ ⊢ m : M ® a ∈ pi_glu_exp_pred i IR IP IEl elem_rel OEl }} }. + {{ Δ ⊢ N : IT[σ] ® n ∈ IEl }} -> + forall (equiv_n : {{ Dom n ≈ n ∈ IR }}), + exists mn, {{ $| m & n |↘ mn }} /\ {{ Δ ⊢ M[σ] N : OT[σ,,N] ® mn ∈ OEl _ equiv_n }}) -> + {{ Γ ⊢ M : A ® m ∈ pi_glu_exp_pred i IR IP IEl elem_rel OEl }} }. #[export] Hint Constructors neut_glu_exp_pred pi_glu_typ_pred pi_glu_exp_pred : mcltt. -Definition univ_glu_typ_pred j i : glu_typ_pred := fun Γ T => {{ Γ ⊢ T ≈ Type@j : Type@i }}. -Arguments univ_glu_typ_pred j i Γ T/. +Definition univ_glu_typ_pred j i : glu_typ_pred := fun Γ A => {{ Γ ⊢ A ≈ Type@j : Type@i }}. +Arguments univ_glu_typ_pred j i Γ A/. Transparent univ_glu_typ_pred. Section Gluing. @@ -117,13 +117,13 @@ Section Gluing. (glu_univ_typ_rec : forall {j}, j < i -> domain -> glu_typ_pred). Definition univ_glu_exp_pred' {j} (lt_j_i : j < i) : glu_exp_pred := - fun Γ M m A => - {{ Γ ⊢ m : M }} /\ - {{ Γ ⊢ M ≈ Type@j : Type@i }} /\ - {{ Γ ⊢ m ® glu_univ_typ_rec lt_j_i A }}. + fun Γ A M m => + {{ Γ ⊢ M : A }} /\ + {{ Γ ⊢ A ≈ Type@j : Type@i }} /\ + {{ Γ ⊢ M ® glu_univ_typ_rec lt_j_i m }}. #[global] - Arguments univ_glu_exp_pred' {j} lt_j_i Γ m M A/. + Arguments univ_glu_exp_pred' {j} lt_j_i Γ A M m/. Inductive glu_univ_elem_core : glu_typ_pred -> glu_exp_pred -> domain -> Prop := | glu_univ_elem_core_univ : @@ -150,12 +150,12 @@ Section Gluing. {{ DG a ∈ glu_univ_elem_core ↘ IP ↘ IEl }} -> {{ DF a ≈ a ∈ per_univ_elem i ↘ in_rel }} -> (forall {c} (equiv_c : {{ Dom c ≈ c ∈ in_rel }}) b, - {{ ⟦ B ⟧ p ↦ c ↘ b }} -> + {{ ⟦ B ⟧ ρ ↦ c ↘ b }} -> {{ DG b ∈ glu_univ_elem_core ↘ OP _ equiv_c ↘ OEl _ equiv_c }}) -> - {{ DF Π a p B ≈ Π a p B ∈ per_univ_elem i ↘ elem_rel }} -> + {{ DF Π a ρ B ≈ Π a ρ B ∈ per_univ_elem i ↘ elem_rel }} -> typ_rel <∙> pi_glu_typ_pred i in_rel IP IEl OP -> el_rel <∙> pi_glu_exp_pred i in_rel IP IEl elem_rel OEl -> - {{ DG Π a p B ∈ glu_univ_elem_core ↘ typ_rel ↘ el_rel }} } + {{ DG Π a ρ B ∈ glu_univ_elem_core ↘ typ_rel ↘ el_rel }} } | glu_univ_elem_core_neut : `{ forall typ_rel el_rel, @@ -169,63 +169,63 @@ End Gluing. Hint Constructors glu_univ_elem_core : mcltt. Equations glu_univ_elem (i : nat) : glu_typ_pred -> glu_exp_pred -> domain -> Prop by wf i := -| i => glu_univ_elem_core i (fun j lt_j_i A Γ M => exists P El, {{ DG A ∈ glu_univ_elem j ↘ P ↘ El }} /\ {{ Γ ⊢ M ® P }}). +| i => glu_univ_elem_core i (fun j lt_j_i a Γ A => exists P El, {{ DG a ∈ glu_univ_elem j ↘ P ↘ El }} /\ {{ Γ ⊢ A ® P }}). -Definition glu_univ_typ (i : nat) (A : domain) : glu_typ_pred := - fun Γ M => exists P El, {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} /\ {{ Γ ⊢ M ® P }}. -Arguments glu_univ_typ i A Γ M/. +Definition glu_univ_typ (i : nat) (a : domain) : glu_typ_pred := + fun Γ A => exists P El, {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} /\ {{ Γ ⊢ A ® P }}. +Arguments glu_univ_typ i a Γ A/. Definition univ_glu_exp_pred j i : glu_exp_pred := - fun Γ M m A => - {{ Γ ⊢ m : M }} /\ {{ Γ ⊢ M ≈ Type@j : Type@i }} /\ - {{ Γ ⊢ m ® glu_univ_typ j A }}. -Arguments univ_glu_exp_pred j i Γ t T a/. + fun Γ A M m => + {{ Γ ⊢ M : A }} /\ {{ Γ ⊢ A ≈ Type@j : Type@i }} /\ + {{ Γ ⊢ M ® glu_univ_typ j m }}. +Arguments univ_glu_exp_pred j i Γ A M m/. Section GluingInduction. Hypothesis (motive : nat -> glu_typ_pred -> glu_exp_pred -> domain -> Prop) (case_univ : - forall i (j : nat) - (typ_rel : glu_typ_pred) (el_rel : glu_exp_pred) (lt_j_i : j < i), - (forall P El A, {{ DG A ∈ glu_univ_elem j ↘ P ↘ El }} -> motive j P El A) -> - typ_rel <∙> univ_glu_typ_pred j i -> - el_rel <∙> univ_glu_exp_pred j i -> - motive i typ_rel el_rel d{{{ 𝕌 @ j }}}) + forall i j + (P : glu_typ_pred) (El : glu_exp_pred) (lt_j_i : j < i), + (forall P' El' a, {{ DG a ∈ glu_univ_elem j ↘ P' ↘ El' }} -> motive j P' El' a) -> + P <∙> univ_glu_typ_pred j i -> + El <∙> univ_glu_exp_pred j i -> + motive i P El d{{{ 𝕌 @ j }}}) (case_nat : - forall i (typ_rel : glu_typ_pred) (el_rel : glu_exp_pred), - typ_rel <∙> nat_glu_typ_pred i -> - el_rel <∙> nat_glu_exp_pred i -> - motive i typ_rel el_rel d{{{ ℕ }}}) + forall i (P : glu_typ_pred) (El : glu_exp_pred), + P <∙> nat_glu_typ_pred i -> + El <∙> nat_glu_exp_pred i -> + motive i P El d{{{ ℕ }}}) (case_pi : - forall i (a : domain) (B : typ) (p : env) (in_rel : relation domain) (IP : glu_typ_pred) + forall i a B (ρ : env) (in_rel : relation domain) (IP : glu_typ_pred) (IEl : glu_exp_pred) (OP : forall c : domain, {{ Dom c ≈ c ∈ in_rel }} -> glu_typ_pred) - (OEl : forall c : domain, {{ Dom c ≈ c ∈ in_rel }} -> glu_exp_pred) (typ_rel : glu_typ_pred) (el_rel : glu_exp_pred) + (OEl : forall c : domain, {{ Dom c ≈ c ∈ in_rel }} -> glu_exp_pred) (P : glu_typ_pred) (El : glu_exp_pred) (elem_rel : relation domain), {{ DG a ∈ glu_univ_elem i ↘ IP ↘ IEl }} -> motive i IP IEl a -> {{ DF a ≈ a ∈ per_univ_elem i ↘ in_rel }} -> (forall (c : domain) (equiv_c : {{ Dom c ≈ c ∈ in_rel }}) (b : domain), - {{ ⟦ B ⟧ p ↦ c ↘ b }} -> + {{ ⟦ B ⟧ ρ ↦ c ↘ b }} -> {{ DG b ∈ glu_univ_elem i ↘ OP c equiv_c ↘ OEl c equiv_c }}) -> (forall (c : domain) (equiv_c : {{ Dom c ≈ c ∈ in_rel }}) (b : domain), - {{ ⟦ B ⟧ p ↦ c ↘ b }} -> + {{ ⟦ B ⟧ ρ ↦ c ↘ b }} -> motive i (OP c equiv_c) (OEl c equiv_c) b) -> - {{ DF Π a p B ≈ Π a p B ∈ per_univ_elem i ↘ elem_rel }} -> - typ_rel <∙> pi_glu_typ_pred i in_rel IP IEl OP -> - el_rel <∙> pi_glu_exp_pred i in_rel IP IEl elem_rel OEl -> - motive i typ_rel el_rel d{{{ Π a p B }}}) + {{ DF Π a ρ B ≈ Π a ρ B ∈ per_univ_elem i ↘ elem_rel }} -> + P <∙> pi_glu_typ_pred i in_rel IP IEl OP -> + El <∙> pi_glu_exp_pred i in_rel IP IEl elem_rel OEl -> + motive i P El d{{{ Π a ρ B }}}) (case_neut : - forall i (b : domain_ne) (a : domain) - (typ_rel : glu_typ_pred) - (el_rel : glu_exp_pred), + forall i b a + (P : glu_typ_pred) + (El : glu_exp_pred), {{ Dom b ≈ b ∈ per_bot }} -> - typ_rel <∙> neut_glu_typ_pred i b -> - el_rel <∙> neut_glu_exp_pred i b -> - motive i typ_rel el_rel d{{{ ⇑ a b }}}) + P <∙> neut_glu_typ_pred i b -> + El <∙> neut_glu_exp_pred i b -> + motive i P El d{{{ ⇑ a b }}}) . #[local] @@ -237,13 +237,13 @@ Section GluingInduction. | i, P, El, a, H => glu_univ_elem_core_ind i - (fun j lt_j_i A Γ M => glu_univ_typ j A Γ M) + (fun j lt_j_i a Γ A => glu_univ_typ j a Γ A) (motive i) - (fun j typ_rel el_rel lt_j_i Hty Hel => - case_univ i j typ_rel el_rel lt_j_i - (fun P El A H => glu_univ_elem_ind j P El A H) - Hty - Hel) + (fun j P' El' lt_j_i HP' HEl' => + case_univ i j P' El' lt_j_i + (fun P'' El'' A H => glu_univ_elem_ind j P'' El'' A H) + HP' + HEl') (case_nat i) _ (* (case_pi i) *) (case_neut i) @@ -254,50 +254,36 @@ Section GluingInduction. Qed. End GluingInduction. -Inductive glu_elem_bot i A Γ T t c : Prop := +Variant glu_elem_bot i a Γ A M m : Prop := | glu_elem_bot_make : forall P El, - {{ Γ ⊢ t : T }} -> - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - {{ Γ ⊢ T ® P }} -> - {{ Dom c ≈ c ∈ per_bot }} -> - (forall Δ σ w, {{ Δ ⊢w σ : Γ }} -> {{ Rne c in length Δ ↘ w }} -> {{ Δ ⊢ t [ σ ] ≈ w : T [ σ ] }}) -> - {{ Γ ⊢ t : T ® c ∈ glu_elem_bot i A }}. + {{ Γ ⊢ M : A }} -> + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + {{ Γ ⊢ A ® P }} -> + {{ Dom m ≈ m ∈ per_bot }} -> + (forall Δ σ M', {{ Δ ⊢w σ : Γ }} -> {{ Rne m in length Δ ↘ M' }} -> {{ Δ ⊢ M[σ] ≈ M' : A[σ] }}) -> + {{ Γ ⊢ M : A ® m ∈ glu_elem_bot i a }}. #[export] - Hint Constructors glu_elem_bot : mcltt. - +Hint Constructors glu_elem_bot : mcltt. -Inductive glu_elem_top i A Γ T t a : Prop := +Variant glu_elem_top i a Γ A M m : Prop := | glu_elem_top_make : forall P El, - {{ Γ ⊢ t : T }} -> - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - {{ Γ ⊢ T ® P }} -> - {{ Dom ⇓ A a ≈ ⇓ A a ∈ per_top }} -> - (forall Δ σ w, {{ Δ ⊢w σ : Γ }} -> {{ Rnf ⇓ A a in length Δ ↘ w }} -> {{ Δ ⊢ t [ σ ] ≈ w : T [ σ ] }}) -> - {{ Γ ⊢ t : T ® a ∈ glu_elem_top i A }}. + {{ Γ ⊢ M : A }} -> + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + {{ Γ ⊢ A ® P }} -> + {{ Dom ⇓ a m ≈ ⇓ a m ∈ per_top }} -> + (forall Δ σ w, {{ Δ ⊢w σ : Γ }} -> {{ Rnf ⇓ a m in length Δ ↘ w }} -> {{ Δ ⊢ M[σ] ≈ w : A[σ] }}) -> + {{ Γ ⊢ M : A ® m ∈ glu_elem_top i a }}. #[export] - Hint Constructors glu_elem_top : mcltt. - +Hint Constructors glu_elem_top : mcltt. -Inductive glu_typ_top i A Γ T : Prop := +Variant glu_typ_top i a Γ A : Prop := | glu_typ_top_make : - {{ Γ ⊢ T : Type@i }} -> - {{ Dom A ≈ A ∈ per_top_typ }} -> - (forall Δ σ W, {{ Δ ⊢w σ : Γ }} -> {{ Rtyp A in length Δ ↘ W }} -> {{ Δ ⊢ T [ σ ] ≈ W : Type@i }}) -> - {{ Γ ⊢ T ® glu_typ_top i A }}. + {{ Γ ⊢ A : Type@i }} -> + {{ Dom a ≈ a ∈ per_top_typ }} -> + (forall Δ σ A', {{ Δ ⊢w σ : Γ }} -> {{ Rtyp a in length Δ ↘ A' }} -> {{ Δ ⊢ A[σ] ≈ A' : Type@i }}) -> + {{ Γ ⊢ A ® glu_typ_top i a }}. #[export] - Hint Constructors glu_typ_top : mcltt. - -Ltac invert_glu_rel1 := - match goal with - | H : pi_glu_typ_pred _ _ _ _ _ _ _ |- _ => - progressive_invert H - | H : pi_glu_exp_pred _ _ _ _ _ _ _ _ _ _ |- _ => - progressive_invert H - | H : neut_glu_typ_pred _ _ _ _ |- _ => - progressive_invert H - | H : neut_glu_exp_pred _ _ _ _ _ _ |- _ => - progressive_invert H - end. +Hint Constructors glu_typ_top : mcltt. Variant glu_rel_typ_with_sub i Δ A σ ρ : Prop := | mk_glu_rel_typ_with_sub : @@ -319,7 +305,10 @@ Variant cons_glu_sub_pred i Γ A (TSb : glu_sub_pred) : glu_sub_pred := {{ Δ ⊢s σ : Γ, A }} -> {{ ⟦ A ⟧ ρ ↯ ↘ a }} -> {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> - {{ Δ ⊢ #0[σ] : A[Wk∘σ] ® ~(ρ 0) ∈ El }} -> + (* Here we use [A[Wk][σ]] instead of [A[Wk∘σ]] + as syntactic judgement derived from that is + a more direct consequence of [{{ Γ, A ⊢ #0 : A[Wk] }}] *) + {{ Δ ⊢ #0[σ] : A[Wk][σ] ® ~(ρ 0) ∈ El }} -> {{ Δ ⊢s Wk ∘ σ ® ρ ↯ ∈ TSb }} -> {{ Δ ⊢s σ ® ρ ∈ cons_glu_sub_pred i Γ A TSb }} }. diff --git a/theories/Core/Soundness/LogicalRelation/Lemmas.v b/theories/Core/Soundness/LogicalRelation/Lemmas.v index 576339b7..b23b245d 100644 --- a/theories/Core/Soundness/LogicalRelation/Lemmas.v +++ b/theories/Core/Soundness/LogicalRelation/Lemmas.v @@ -187,15 +187,13 @@ Section glu_univ_elem_cumulativity. + assert {{ Δ ⊢ IT[σ] ® IP }} by mauto. enough (forall Γ A, {{ Γ ⊢ A ® IP }} -> {{ Γ ⊢ A ® IP' }}) by mauto 4. eapply proj1... - + rename a0 into c. - rename equiv_a into equiv_c. - match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). + + match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). apply_relation_equivalence. destruct_rel_mod_eval. handle_per_univ_elem_irrel. - assert (forall Γ A, {{ Γ ⊢ A ® OP c equiv_c }} -> {{ Γ ⊢ A ® OP' c equiv_c }}) by (eapply proj1; mauto). - enough {{ Δ ⊢ OT[σ,,m] ® OP c equiv_c }} by mauto. - enough {{ Δ ⊢ m : IT[σ] ® c ∈ IEl }} by mauto. + assert (forall Γ A, {{ Γ ⊢ A ® OP m equiv_m }} -> {{ Γ ⊢ A ® OP' m equiv_m }}) by (eapply proj1; mauto). + enough {{ Δ ⊢ OT[σ,,M] ® OP m equiv_m }} by mauto. + enough {{ Δ ⊢ M : IT[σ] ® m ∈ IEl }} by mauto. eapply IHHglu... - rename x into IP'. rename x0 into IEl'. @@ -206,18 +204,16 @@ Section glu_univ_elem_cumulativity. econstructor; intros; mauto 4. + enough (forall Γ A, {{ Γ ⊢ A ® IP }} -> {{ Γ ⊢ A ® IP' }}) by mauto 4. eapply proj1... - + rename b into c. - rename equiv_b into equiv_c. - match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). + + match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). handle_per_univ_elem_irrel. destruct_rel_mod_eval. destruct_rel_mod_app. handle_per_univ_elem_irrel. eexists; split; mauto 4. - assert (forall Γ A M m, {{ Γ ⊢ M : A ® m ∈ OEl c equiv_c }} -> {{ Γ ⊢ M : A ® m ∈ OEl' c equiv_c }}) by (eapply proj1, proj2; mauto 4). - enough {{ Δ ⊢ m[σ] m' : OT[σ,,m'] ® fa ∈ OEl c equiv_c }} by mauto 4. - assert {{ Δ ⊢ m' : IT[σ] ® c ∈ IEl }} by (eapply IHHglu; mauto 3). - assert (exists ac, {{ $| a0 & c |↘ ac }} /\ {{ Δ ⊢ m[σ] m' : OT[σ,,m'] ® ac ∈ OEl c equiv_c }}) by mauto 4. + assert (forall Γ A M m, {{ Γ ⊢ M : A ® m ∈ OEl n equiv_n }} -> {{ Γ ⊢ M : A ® m ∈ OEl' n equiv_n }}) by (eapply proj1, proj2; mauto 4). + enough {{ Δ ⊢ M[σ] N : OT[σ,,N] ® fa ∈ OEl n equiv_n }} by mauto 4. + assert {{ Δ ⊢ N : IT[σ] ® n ∈ IEl }} by (eapply IHHglu; mauto 3). + assert (exists mn, {{ $| m & n |↘ mn }} /\ {{ Δ ⊢ M[σ] N : OT[σ,,N] ® mn ∈ OEl n equiv_n }}) by mauto 4. destruct_conjs. functional_eval_rewrite_clear... - rename x into IP'. @@ -228,29 +224,30 @@ Section glu_univ_elem_cumulativity. destruct_by_head pi_glu_exp_pred. handle_per_univ_elem_irrel. econstructor; intros; mauto. - rename b into c. - rename equiv_b into equiv_c. match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). handle_per_univ_elem_irrel. destruct_rel_mod_eval. destruct_rel_mod_app. handle_per_univ_elem_irrel. - rename a1 into b. + match goal with + | _: {{ ⟦ B ⟧ ρ ↦ n ↘ ~?a }} |- _ => + rename a into b + end. eexists; split; mauto 4. - assert (forall Γ A M m, {{ Γ ⊢ A ® OP c equiv_c }} -> {{ Γ ⊢ M : A ® m ∈ OEl' c equiv_c }} -> {{ Γ ⊢ M : A ® m ∈ OEl c equiv_c }}) by (eapply proj2, proj2; eauto 3). - assert {{ Δ ⊢ OT[σ,,m'] ® OP c equiv_c }} by mauto 4. - enough {{ Δ ⊢ m[σ] m' : OT[σ,,m'] ® fa ∈ OEl' c equiv_c }} by mauto 4. - assert {{ Δ ⊢ m' : IT[σ] ® c ∈ IEl' }} by (eapply IHHglu; mauto 3). + assert (forall Γ A M m, {{ Γ ⊢ A ® OP n equiv_n }} -> {{ Γ ⊢ M : A ® m ∈ OEl' n equiv_n }} -> {{ Γ ⊢ M : A ® m ∈ OEl n equiv_n }}) by (eapply proj2, proj2; eauto 3). + assert {{ Δ ⊢ OT[σ,,N] ® OP n equiv_n }} by mauto 4. + enough {{ Δ ⊢ M[σ] N : OT[σ,,N] ® fa ∈ OEl' n equiv_n }} by mauto 4. + assert {{ Δ ⊢ N : IT[σ] ® n ∈ IEl' }} by (eapply IHHglu; mauto 3). assert {{ Δ ⊢ IT[σ] ® IP' }} by (eapply glu_univ_elem_trm_typ; mauto 2). assert {{ Δ ⊢ IT0[σ] ® IP' }} by mauto 2. assert {{ Δ ⊢ IT[σ] ≈ IT0[σ] : Type@j }} as HITeq by mauto 2. - assert {{ Δ ⊢ m' : IT0[σ] ® c ∈ IEl' }} by (rewrite <- HITeq; mauto 3). - assert (exists ac, {{ $| a0 & c |↘ ac }} /\ {{ Δ ⊢ m[σ] m' : OT0[σ,,m'] ® ac ∈ OEl' c equiv_c }}) by mauto 3. + assert {{ Δ ⊢ N : IT0[σ] ® n ∈ IEl' }} by (rewrite <- HITeq; mauto 3). + assert (exists mn, {{ $| m & n |↘ mn }} /\ {{ Δ ⊢ M[σ] N : OT0[σ,,N] ® mn ∈ OEl' n equiv_n }}) by mauto 3. destruct_conjs. functional_eval_rewrite_clear. - assert {{ DG b ∈ glu_univ_elem j ↘ OP' c equiv_c ↘ OEl' c equiv_c }} by mauto 3. - assert {{ Δ ⊢ OT0[σ,,m'] ® OP' c equiv_c }} by (eapply glu_univ_elem_trm_typ; mauto 3). - enough {{ Δ ⊢ OT[σ,,m'] ≈ OT0[σ,,m'] : Type@j }} as ->... + assert {{ DG b ∈ glu_univ_elem j ↘ OP' n equiv_n ↘ OEl' n equiv_n }} by mauto 3. + assert {{ Δ ⊢ OT0[σ,,N] ® OP' n equiv_n }} by (eapply glu_univ_elem_trm_typ; mauto 3). + enough {{ Δ ⊢ OT[σ,,N] ≈ OT0[σ,,N] : Type@j }} as ->... - destruct_by_head neut_glu_exp_pred. econstructor; mauto. destruct_by_head neut_glu_typ_pred. @@ -412,7 +409,7 @@ Proof. mauto 3. - destruct_by_head pi_glu_typ_pred. rename x into IP. rename x0 into IEl. rename x1 into OP. rename x2 into OEl. - rename M0 into M'. rename IT0 into IT'. rename OT0 into OT'. + rename A0 into A'. rename IT0 into IT'. rename OT0 into OT'. rename x3 into OP'. rename x4 into OEl'. assert {{ Γ ⊢ IT ® IP }}. { @@ -437,8 +434,12 @@ Proof. apply_relation_equivalence. destruct_rel_mod_eval. handle_per_univ_elem_irrel. - rename a1 into b. - rename a3 into b'. + match goal with + | _: {{ ⟦ B ⟧ ρ ↦ ⇑! a (length Γ) ↘ ~?a0 }}, + _: {{ ⟦ B' ⟧ ρ' ↦ ⇑! a' (length Γ) ↘ ~?a0' }} |- _ => + rename a0 into b; + rename a0' into b' + end. assert {{ DG b ∈ glu_univ_elem i ↘ OP d{{{ ⇑! a (length Γ) }}} equiv_len_len ↘ OEl d{{{ ⇑! a (length Γ) }}} equiv_len_len }} by mauto 4. assert {{ DG b' ∈ glu_univ_elem i ↘ OP' d{{{ ⇑! a' (length Γ) }}} equiv_len'_len' ↘ OEl' d{{{ ⇑! a' (length Γ) }}} equiv_len'_len' }} by mauto 4. assert {{ Γ, IT' ⊢ OT ® OP d{{{ ⇑! a (length Γ) }}} equiv_len_len }}. @@ -483,18 +484,14 @@ Proof. handle_functional_glu_univ_elem; repeat invert_glu_rel1; try solve [simpl in *; intuition mauto 3]. - - rename A into a''. - rename M into A. - rename m into M. - clear_dups. - match_by_head1 (per_bot b b') ltac:(fun H => destruct (H (length Γ)) as [V []]). + - match_by_head1 (per_bot b b') ltac:(fun H => destruct (H (length Γ)) as [V []]). econstructor; mauto 3. + econstructor; mauto 3. + intros. match_by_head1 (per_bot b b') ltac:(fun H => destruct (H (length Δ)) as [? []]). functional_read_rewrite_clear. assert {{ Δ ⊢ A[σ] ⊆ A'[σ] }} by mauto 3. - assert {{ Δ ⊢ M[σ] ≈ v : A[σ] }} by mauto 3. + assert {{ Δ ⊢ M[σ] ≈ M' : A[σ] }} by mauto 3. mauto 3. - simpl in *. destruct_conjs. @@ -502,20 +499,16 @@ Proof. assert (exists P El, {{ DG m ∈ glu_univ_elem j ↘ P ↘ El }}) as [? [?]] by mauto 3. do 2 eexists; split; mauto 3. eapply glu_univ_elem_typ_cumu_ge; revgoals; mautosolve 3. - - rename M into A. - rename M0 into A'. - rename m into M. + - rename A0 into A'. rename IT0 into IT'. rename OT0 into OT'. rename x into IP. rename x0 into IEl. rename x1 into OP. rename x2 into OEl. rename x3 into OP'. rename x4 into OEl'. handle_per_univ_elem_irrel. econstructor; mauto 3. - + enough {{ Sub Π a p B <: Π a' p' B' at i }} by (eapply per_elem_subtyping; try eassumption). + + enough {{ Sub Π a ρ B <: Π a' ρ' B' at i }} by (eapply per_elem_subtyping; try eassumption). econstructor; mauto 3. + intros. - rename b into c. - rename equiv_b into equiv_c. assert {{ Γ ⊢w Id : Γ }} by mauto 3. assert {{ Γ ⊢ IT ® IP }}. { @@ -530,20 +523,24 @@ Proof. autorewrite with mcltt in *; mauto 3. } assert {{ Δ ⊢ IT[σ] ≈ IT'[σ] : Type@i }} by mauto 4 using glu_univ_elem_per_univ_typ_escape. - assert {{ Δ ⊢ m' : IT[σ] ® c ∈ IEl }} by (simpl; bulky_rewrite1; eassumption). - assert (exists ac : domain, {{ $| a0 & c |↘ ac }} /\ OEl c equiv_c Δ (a_sub OT {{{ σ,, m' }}}) {{{ ~ (a_sub M σ) m' }}} ac) by mauto 3. + assert {{ Δ ⊢ N : IT[σ] ® n ∈ IEl }} by (simpl; bulky_rewrite1; eassumption). + assert (exists mn : domain, {{ $| m & n |↘ mn }} /\ {{Δ ⊢ M[σ] N : OT[σ,,N] ® mn ∈ OEl n equiv_n }}) by mauto 3. destruct_conjs. eexists; split; mauto 3. match_by_head per_univ_elem ltac:(fun H => directed invert_per_univ_elem H). destruct_rel_mod_eval. handle_per_univ_elem_irrel. - rename a1 into b. - rename a2 into b'. - assert {{ DG b ∈ glu_univ_elem i ↘ OP c equiv_c ↘ OEl c equiv_c }} by mauto 3. - assert {{ DG b' ∈ glu_univ_elem i ↘ OP' c equiv_c ↘ OEl' c equiv_c }} by mauto 3. - assert {{ Δ ⊢ OT[σ,,m'] ® OP c equiv_c }} by (eapply glu_univ_elem_trm_typ; mauto 3). + match goal with + | _: {{ ⟦ B ⟧ ρ ↦ n ↘ ~?a }}, + _: {{ ⟦ B' ⟧ ρ' ↦ n ↘ ~?a' }} |- _ => + rename a into b; + rename a' into b' + end. + assert {{ DG b ∈ glu_univ_elem i ↘ OP n equiv_n ↘ OEl n equiv_n }} by mauto 3. + assert {{ DG b' ∈ glu_univ_elem i ↘ OP' n equiv_n ↘ OEl' n equiv_n }} by mauto 3. + assert {{ Δ ⊢ OT[σ,,N] ® OP n equiv_n }} by (eapply glu_univ_elem_trm_typ; mauto 3). assert {{ Sub b <: b' at i }} by mauto 3. - assert {{ Δ ⊢ OT[σ,,m'] ⊆ OT'[σ,,m'] }} by mauto 3. + assert {{ Δ ⊢ OT[σ,,N] ⊆ OT'[σ,,N] }} by mauto 3. intuition. Qed. @@ -678,10 +675,10 @@ Hint Resolve glu_rel_typ_with_sub_implies_glu_rel_exp_with_sub : mcltt. Lemma glu_ctx_env_sub_resp_ctx_eq : forall {Γ Sb}, {{ EG Γ ∈ glu_ctx_env ↘ Sb }} -> - forall {Δ Δ' σ p}, - {{ Δ ⊢s σ ® p ∈ Sb }} -> + forall {Δ Δ' σ ρ}, + {{ Δ ⊢s σ ® ρ ∈ Sb }} -> {{ ⊢ Δ ≈ Δ' }} -> - {{ Δ' ⊢s σ ® p ∈ Sb }}. + {{ Δ' ⊢s σ ® ρ ∈ Sb }}. Proof. induction 1; intros * HSb Hctxeq; apply_predicate_equivalence; @@ -702,10 +699,10 @@ Qed. Lemma glu_ctx_env_sub_resp_sub_eq : forall {Γ Sb}, {{ EG Γ ∈ glu_ctx_env ↘ Sb }} -> - forall {Δ σ σ' p}, - {{ Δ ⊢s σ ® p ∈ Sb }} -> + forall {Δ σ σ' ρ}, + {{ Δ ⊢s σ ® ρ ∈ Sb }} -> {{ Δ ⊢s σ ≈ σ' : Γ }} -> - {{ Δ ⊢s σ' ® p ∈ Sb }}. + {{ Δ ⊢s σ' ® ρ ∈ Sb }}. Proof. induction 1; intros * HSb Hsubeq; apply_predicate_equivalence; @@ -716,10 +713,8 @@ Proof. destruct_by_head cons_glu_sub_pred. econstructor; mauto 4. assert {{ Γ, A ⊢s Wk : Γ }} by mauto 3. - assert {{ Δ ⊢s Wk∘σ ≈ Wk∘σ' : Γ }} by mauto 3. - assert {{ Δ ⊢ A[Wk∘σ] ≈ A[Wk∘σ'] : Type@i }} as <- by mauto 3. - assert {{ Δ ⊢ #0[σ] ≈ #0[σ'] : A[Wk][σ] }} by mauto 3. - assert {{ Δ ⊢ #0[σ] ≈ #0[σ'] : A[Wk∘σ] }} as <- by mauto 3. + assert {{ Δ ⊢ A[Wk][σ] ≈ A[Wk][σ'] : Type@i }} as <- by mauto 3. + assert {{ Δ ⊢ #0[σ] ≈ #0[σ'] : A[Wk][σ] }} as <- by mauto 3. eassumption. Qed. @@ -743,11 +738,8 @@ Proof. assert {{ Δ ⊢s Wk∘σ' : Γ }} by mauto 3. assert {{ Δ ⊢s Wk∘σ ≈ Wk∘σ' : Γ }} by mauto 3. econstructor; mauto 3. - - assert {{ Δ ⊢ #0[σ] ≈ #0[σ'] : A[Wk][σ] }} by mauto 3. - assert {{ Δ ⊢ A[Wk][σ] ≈ A[Wk∘σ] : Type@i }} by mauto 3. - assert {{ Δ ⊢ A[Wk][σ] ≈ A[Wk∘σ'] : Type@i }} by mauto 3. - assert {{ Δ ⊢ #0[σ] ≈ #0[σ'] : A[Wk∘σ'] }} as <- by mauto 3. - assert {{ Δ ⊢ A[Wk∘σ] ≈ A[Wk∘σ'] : Type@i }} as <-; mauto 2. + - assert {{ Δ ⊢ A[Wk][σ] ≈ A[Wk][σ'] : Type@i }} as <- by mauto 4. + assert {{ Δ ⊢ #0[σ] ≈ #0[σ'] : A[Wk][σ] }} as <-; mauto 3. - assert {{ Δ ⊢s Wk∘σ ≈ Wk∘σ' : Γ }} as <-; eassumption. Qed. @@ -757,14 +749,14 @@ Proof. split; mauto using cons_glu_sub_pred_resp_wf_sub_eq. Qed. -Lemma glu_ctx_env_per_env : forall {Γ Sb env_rel Δ σ p}, +Lemma glu_ctx_env_per_env : forall {Γ Sb env_rel Δ σ ρ}, {{ EG Γ ∈ glu_ctx_env ↘ Sb }} -> {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }} -> - {{ Δ ⊢s σ ® p ∈ Sb }} -> - {{ Dom p ≈ p ∈ env_rel }}. + {{ Δ ⊢s σ ® ρ ∈ Sb }} -> + {{ Dom ρ ≈ ρ ∈ env_rel }}. Proof. intros * Hglu Hper. - gen p σ Δ env_rel. + gen ρ σ Δ env_rel. induction Hglu; intros; invert_per_ctx_env Hper; apply_predicate_equivalence; @@ -773,7 +765,7 @@ Proof. rename i0 into j. inversion_clear_by_head cons_glu_sub_pred. - assert {{ Dom p ↯ ≈ p ↯ ∈ tail_rel }} by intuition. + assert {{ Dom ρ ↯ ≈ ρ ↯ ∈ tail_rel }} by intuition. destruct_rel_typ. handle_per_univ_elem_irrel. assert (exists P' El', {{ DG a ∈ glu_univ_elem (max i j) ↘ P' ↘ El' }}) as [? []] by mauto 3 using glu_univ_elem_cumu_max_left. @@ -790,8 +782,8 @@ Qed. Lemma glu_ctx_env_sub_escape : forall {Γ Sb}, {{ EG Γ ∈ glu_ctx_env ↘ Sb }} -> - forall Δ σ p, - {{ Δ ⊢s σ ® p ∈ Sb }} -> + forall Δ σ ρ, + {{ Δ ⊢s σ ® ρ ∈ Sb }} -> {{ Δ ⊢s σ : Γ }}. Proof. induction 1; intros; @@ -840,12 +832,12 @@ Proof. rename i1 into k. inversion HΓΓ' as [|? ? l]; subst. assert (TSb -∙> TSb') by intuition. - intros Δ σ p []. + intros Δ σ ρ []. saturate_refl_for per_ctx_env. - assert {{ Dom ρ ↯ ≈ ρ ↯ ∈ tail_rel }} + assert {{ Dom ρ0 ↯ ≈ ρ0 ↯ ∈ tail_rel }} by (eapply glu_ctx_env_per_env; [| | eassumption]; eassumption). - assert {{ Δ0 ⊢s Wk∘σ0 ® ρ ↯ ∈ TSb' }} by intuition. - assert (glu_rel_typ_with_sub j Δ0 A' {{{ Wk∘σ0 }}} d{{{ ρ ↯ }}}) as [] by mauto 3. + assert {{ Δ0 ⊢s Wk∘σ0 ® ρ0 ↯ ∈ TSb' }} by intuition. + assert (glu_rel_typ_with_sub j Δ0 A' {{{ Wk∘σ0 }}} d{{{ ρ0 ↯ }}}) as [] by mauto 3. destruct_rel_typ. handle_functional_glu_univ_elem. rename a0 into a'. @@ -855,13 +847,16 @@ Proof. assert (exists Pmax Elmax, {{ DG a ∈ glu_univ_elem (max i l) ↘ Pmax ↘ Elmax }}) as [Pmax [Elmax]] by mauto using glu_univ_elem_cumu_max_left. assert (i <= max i l) by lia. - assert {{ Δ0 ⊢ #0[σ0] : A'[Wk∘σ0] ® ~(ρ 0) ∈ Elmax }}. + assert {{ Δ0 ⊢ #0[σ0] : A'[Wk][σ0] ® ~(ρ0 0) ∈ Elmax }}. { - assert {{ Δ0 ⊢s Wk∘σ0 : Γ }} by mauto 4. - assert {{ Δ0 ⊢ A'[Wk∘σ0] ≈ A[Wk∘σ0] : Type@(max i l) }} as -> by mauto 4 using lift_exp_eq_max_right. + assert {{ Γ, A ⊢s Wk : Γ }} by mauto 3. + assert {{ Δ0 ⊢ A[Wk][σ0] ≈ A'[Wk][σ0] : Type@l }} by mauto 3. + assert {{ Δ0 ⊢ A[Wk][σ0] ≈ A'[Wk][σ0] : Type@(max i l) }} as <- by mauto 2 using lift_exp_eq_max_right. eapply glu_univ_elem_exp_cumu_ge; mauto 3. } eapply glu_univ_elem_exp_conv; [eexists | | | |]; intuition. + autorewrite with mcltt. + eassumption. Qed. Corollary functional_glu_ctx_env : forall {Γ Sb Sb'}, @@ -923,15 +918,15 @@ Ltac invert_glu_ctx_env H := destruct H as [? [? []]]) + (inversion H; subst). -Lemma glu_ctx_env_subtyp_sub_if : forall Γ Γ' Sb Sb' Δ σ p, +Lemma glu_ctx_env_subtyp_sub_if : forall Γ Γ' Sb Sb' Δ σ ρ, {{ ⊢ Γ ⊆ Γ' }} -> {{ EG Γ ∈ glu_ctx_env ↘ Sb }} -> {{ EG Γ' ∈ glu_ctx_env ↘ Sb' }} -> - {{ Δ ⊢s σ ® p ∈ Sb }} -> - {{ Δ ⊢s σ ® p ∈ Sb' }}. + {{ Δ ⊢s σ ® ρ ∈ Sb }} -> + {{ Δ ⊢s σ ® ρ ∈ Sb' }}. Proof. intros * Hsubtyp Hglu Hglu'. - gen p σ Δ. gen Sb' Sb. + gen ρ σ Δ. gen Sb' Sb. induction Hsubtyp; intros; invert_glu_ctx_env Hglu; invert_glu_ctx_env Hglu'; @@ -958,14 +953,16 @@ Proof. destruct_by_head rel_exp. handle_functional_glu_ctx_env. eapply glu_univ_elem_per_subtyp_trm_conv; mauto 3. + autorewrite with mcltt. + eassumption. Qed. Lemma glu_ctx_env_sub_monotone : forall Γ Sb, {{ EG Γ ∈ glu_ctx_env ↘ Sb }} -> - forall Δ' σ Δ τ p, - {{ Δ ⊢s τ ® p ∈ Sb }} -> + forall Δ' σ Δ τ ρ, + {{ Δ ⊢s τ ® ρ ∈ Sb }} -> {{ Δ' ⊢w σ : Δ }} -> - {{ Δ' ⊢s τ ∘ σ ® p ∈ Sb }}. + {{ Δ' ⊢s τ ∘ σ ® ρ ∈ Sb }}. Proof. induction 1; intros * HSb Hσ; apply_predicate_equivalence; @@ -974,17 +971,12 @@ Proof. destruct_by_head cons_glu_sub_pred. econstructor; mauto 3. - - assert {{ Δ' ⊢ #0[σ0][σ] : A[Wk∘σ0][σ] ® ~(ρ 0) ∈ El }} by (eapply glu_univ_elem_exp_monotone; mauto 3). + - assert {{ Δ' ⊢ #0[σ0][σ] : A[Wk][σ0][σ] ® ~(ρ 0) ∈ El }} by (eapply glu_univ_elem_exp_monotone; mauto 3). assert {{ Γ, A ⊢ #0 : A[Wk] }} by mauto 3. - assert {{ Δ ⊢ A[Wk∘σ0] : Type@i }} by (eapply glu_univ_elem_trm_univ_lvl; mauto 3). - assert {{ ⊢ Γ, A }} by mauto 3. - assert {{ Γ,A ⊢s Wk : Γ }} by mauto 3. + assert {{ Γ, A ⊢s Wk : Γ }} by mauto 3. + assert {{ Δ' ⊢ #0[σ0∘σ] ≈ #0[σ0][σ] : A[Wk][σ0∘σ] }} as -> by mauto 3. assert {{ Δ' ⊢s σ : Δ }} by mauto 3. - assert {{ Δ' ⊢ A[Wk][σ0∘σ] ≈ A[Wk∘(σ0∘σ)] : Type@i }} as <- by mauto 3. - assert {{ Δ' ⊢ #0[σ0][σ] ≈ #0[σ0∘σ] : A[Wk][σ0∘σ] }} as <- by mauto 3. assert {{ Δ' ⊢ A[Wk][σ0][σ] ≈ A[Wk][σ0∘σ] : Type@i }} as <- by mauto 3. - assert {{ Δ ⊢ A[Wk∘σ0] ≈ A[Wk][σ0] : Type@i }} by mauto 3. - assert {{ Δ' ⊢ A[Wk∘σ0][σ] ≈ A[Wk][σ0][σ] : Type@i }} as <- by mauto 3. eassumption. - assert {{ Δ' ⊢s σ : Δ }} by mauto 3. assert {{ Γ, A ⊢s Wk : Γ }} by mauto 3. @@ -1004,8 +996,9 @@ Proof. intros. assert {{ Δ ⊢s σ : Γ }} by mauto 2. assert {{ Δ ⊢ M : A[σ] }} by mauto 2 using glu_univ_elem_trm_escape. + assert {{ Δ ⊢s σ,,M : Γ, A }} by mauto 2. econstructor; mauto 3; - assert {{ Δ ⊢s Wk∘(σ,,M) ≈ σ : Γ }} as ->; mauto 3. + autorewrite with mcltt; mauto 3. assert {{ Δ ⊢ #0[σ,,M] ≈ M : A[σ] }} as ->; mauto 2. Qed. @@ -1013,29 +1006,31 @@ Qed. #[export] Hint Resolve cons_glu_sub_pred_helper : mcltt. -Lemma initial_env_glu_rel_exp : forall {Γ p Sb}, - initial_env Γ p -> +Lemma initial_env_glu_rel_exp : forall {Γ ρ Sb}, + initial_env Γ ρ -> {{ EG Γ ∈ glu_ctx_env ↘ Sb }} -> - {{ Γ ⊢s Id ® p ∈ Sb }}. + {{ Γ ⊢s Id ® ρ ∈ Sb }}. Proof. intros * Hinit HΓ. - gen p. + gen ρ. induction HΓ; intros * Hinit; dependent destruction Hinit; apply_predicate_equivalence; try solve [econstructor; mauto]. - assert (glu_rel_typ_with_sub i Γ A {{{ Id }}} p) as [] by mauto. + assert (glu_rel_typ_with_sub i Γ A {{{ Id }}} ρ) as [] by mauto. functional_eval_rewrite_clear. econstructor; mauto. - - eapply realize_glu_elem_bot; mauto. + - match goal with + | H: P Γ {{{ A[Id] }}} |- _ => + bulky_rewrite_in H + end. + eapply realize_glu_elem_bot; mauto. assert {{ ⊢ Γ }} by mauto 3. - assert {{ Γ ⊢ A[Id] : Type@i }} by mauto 3. - assert {{ ⊢ Γ, A[Id] ≈ Γ, A }} as <- by mauto 4. - assert {{ Γ, A[Id] ⊢ A[Wk] : Type@i }} by mauto 4. - assert {{ Γ, A[Id] ⊢s Wk : Γ }} by mauto 4. - assert {{ Γ, A[Id] ⊢ A[Id][Wk] ≈ A[Wk∘Id] : Type@i }} as <- by (transitivity {{{ A[Wk] }}}; mauto 4). - assert {{ Γ, A[Id] ⊢ #0 ≈ #0[Id] : A[Id][Wk] }} as <- by mauto. + assert {{ Γ, A ⊢s Wk : Γ }} by mauto 4. + assert {{ Γ, A ⊢ A[Wk] : Type@i }} by mauto 4. + assert {{ Γ, A ⊢ A[Wk] ≈ A[Wk][Id] : Type@i }} as <- by mauto 3. + assert {{ Γ, A ⊢ #0 ≈ #0[Id] : A[Wk] }} as <- by mauto. eapply var_glu_elem_bot; mauto. - assert {{ Γ, A ⊢s Wk : Γ }} by mauto 4. assert {{ Γ, A ⊢s Wk∘Id : Γ }} by mauto 4. @@ -1141,10 +1136,10 @@ Proof. intros * [Sb]. destruct_conjs. assert (exists env_rel, {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_rel }}) as [env_rel] by mauto 3. - assert (exists p p', initial_env Γ p /\ initial_env Γ p' /\ {{ Dom p ≈ p' ∈ env_rel }}) as [p] by mauto using per_ctx_then_per_env_initial_env. + assert (exists ρ ρ', initial_env Γ ρ /\ initial_env Γ ρ' /\ {{ Dom ρ ≈ ρ' ∈ env_rel }}) as [ρ] by mauto using per_ctx_then_per_env_initial_env. destruct_conjs. functional_initial_env_rewrite_clear. - assert {{ Γ ⊢s Id ® p ∈ Sb }} by (eapply initial_env_glu_rel_exp; mauto 3). + assert {{ Γ ⊢s Id ® ρ ∈ Sb }} by (eapply initial_env_glu_rel_exp; mauto 3). destruct_glu_rel_exp_with_sub. enough {{ Γ ⊢ M[Id] : A[Id] }} as HId; mauto 3 using glu_univ_elem_trm_escape. Qed. @@ -1216,10 +1211,10 @@ Proof. intros * [SbΓ [SbΔ]]. destruct_conjs. assert (exists env_relΓ, {{ EF Γ ≈ Γ ∈ per_ctx_env ↘ env_relΓ }}) as [env_relΓ] by mauto 3. - assert (exists p p', initial_env Γ p /\ initial_env Γ p' /\ {{ Dom p ≈ p' ∈ env_relΓ }}) as [p] by mauto 3 using per_ctx_then_per_env_initial_env. + assert (exists ρ ρ', initial_env Γ ρ /\ initial_env Γ ρ' /\ {{ Dom ρ ≈ ρ' ∈ env_relΓ }}) as [ρ] by mauto 3 using per_ctx_then_per_env_initial_env. destruct_conjs. functional_initial_env_rewrite_clear. - assert {{ Γ ⊢s Id ® p ∈ SbΓ }} by (eapply initial_env_glu_rel_exp; mauto 3). + assert {{ Γ ⊢s Id ® ρ ∈ SbΓ }} by (eapply initial_env_glu_rel_exp; mauto 3). destruct_glu_rel_sub_with_sub. mauto 3. Qed. diff --git a/theories/Core/Soundness/NatCases.v b/theories/Core/Soundness/NatCases.v index c8e55c7a..49919751 100644 --- a/theories/Core/Soundness/NatCases.v +++ b/theories/Core/Soundness/NatCases.v @@ -131,18 +131,18 @@ Qed. #[export] Hint Resolve glu_rel_sub_extend_nat : mcltt. -Lemma glu_rel_exp_natrec_zero_helper : forall {i Γ SbΓ A MZ MS Δ M σ p am P El}, +Lemma glu_rel_exp_natrec_zero_helper : forall {i Γ SbΓ A MZ MS Δ M σ ρ am P El}, {{ EG Γ ∈ glu_ctx_env ↘ SbΓ }} -> {{ Γ, ℕ ⊢ A : Type@i }} -> {{ Γ ⊩ A[Id,,zero] : Type@i }} -> {{ Γ ⊩ MZ : A[Id,,zero] }} -> {{ Γ, ℕ, A ⊢ MS : A[Wk∘Wk,,succ #1] }} -> {{ Δ ⊢ M ≈ zero : ℕ }} -> - {{ Δ ⊢s σ ® p ∈ SbΓ }} -> - {{ ⟦ A ⟧ p ↦ zero ↘ am }} -> + {{ Δ ⊢s σ ® ρ ∈ SbΓ }} -> + {{ ⟦ A ⟧ ρ ↦ zero ↘ am }} -> {{ DG am ∈ glu_univ_elem i ↘ P ↘ El }} -> exists r, - {{ rec zero ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} /\ + {{ rec zero ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} /\ {{ Δ ⊢ rec M return A[q σ] | zero -> MZ[σ] | succ -> MS[q (q σ)] end : A[σ,,M] ® r ∈ El }}. Proof. intros * ? ? ? HMZ **. @@ -153,7 +153,6 @@ Proof. assert {{ EG Γ, ℕ ∈ glu_ctx_env ↘ SbΓℕ }} by (invert_glu_rel_exp Hℕ; econstructor; mauto 3; reflexivity). destruct_glu_rel_exp_with_sub. simplify_evals. - rename p'0 into p. rename m into mz. eexists mz; split; mauto 3. handle_functional_glu_univ_elem. @@ -201,7 +200,7 @@ Qed. #[local] Hint Resolve cons_glu_sub_pred_nat_helper : mcltt. -Lemma glu_rel_exp_natrec_succ_helper : forall {i Γ SbΓ A MZ MS Δ M M' m' σ p am P El}, +Lemma glu_rel_exp_natrec_succ_helper : forall {i Γ SbΓ A MZ MS Δ M M' m' σ ρ am P El}, {{ EG Γ ∈ glu_ctx_env ↘ SbΓ }} -> {{ Γ, ℕ ⊩ A : Type@i }} -> {{ Γ ⊢ MZ : A[Id,,zero] }} -> @@ -209,18 +208,18 @@ Lemma glu_rel_exp_natrec_succ_helper : forall {i Γ SbΓ A MZ MS Δ M M' m' σ p {{ Γ, ℕ, A ⊩ MS : A[Wk∘Wk,,succ #1] }} -> {{ Δ ⊢ M ≈ succ M' : ℕ }} -> glu_nat Δ M' m' -> - (forall σ p am P El, - {{ Δ ⊢s σ ® p ∈ SbΓ }} -> - {{ ⟦ A ⟧ p ↦ m' ↘ am }} -> + (forall σ ρ am P El, + {{ Δ ⊢s σ ® ρ ∈ SbΓ }} -> + {{ ⟦ A ⟧ ρ ↦ m' ↘ am }} -> {{ DG am ∈ glu_univ_elem i ↘ P ↘ El }} -> exists r, - {{ rec m' ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} /\ + {{ rec m' ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} /\ {{ Δ ⊢ rec M' return A[q σ] | zero -> MZ[σ] | succ -> MS[q (q σ)] end : A[σ,,M'] ® r ∈ El }}) -> - {{ Δ ⊢s σ ® p ∈ SbΓ }} -> - {{ ⟦ A ⟧ p ↦ succ m' ↘ am }} -> + {{ Δ ⊢s σ ® ρ ∈ SbΓ }} -> + {{ ⟦ A ⟧ ρ ↦ succ m' ↘ am }} -> {{ DG am ∈ glu_univ_elem i ↘ P ↘ El }} -> exists r, - {{ rec succ m' ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} /\ + {{ rec succ m' ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} /\ {{ Δ ⊢ rec M return A[q σ] | zero -> MZ[σ] | succ -> MS[q (q σ)] end : A[σ,,M] ® r ∈ El }}. Proof. intros * ? HA ? ? HMS **. @@ -234,7 +233,7 @@ Proof. assert {{ EG Γ, ℕ, A ∈ glu_ctx_env ↘ SbΓℕA }} by (econstructor; mauto 3; reflexivity). assert {{ Γ, ℕ, A ⊢ MS : A[Wk∘Wk,,succ #1] }} by mauto 2. invert_glu_rel_exp HMS. - assert {{ Δ ⊢s σ,,M' ® p ↦ m' ∈ SbΓℕ }} by (unfold SbΓℕ; mauto 3). + assert {{ Δ ⊢s σ,,M' ® ρ ↦ m' ∈ SbΓℕ }} by (unfold SbΓℕ; mauto 3). destruct_glu_rel_exp_with_sub. simplify_evals. match_by_head glu_univ_elem ltac:(fun H => directed invert_glu_univ_elem H). @@ -242,7 +241,7 @@ Proof. unfold univ_glu_exp_pred' in *. destruct_conjs. match goal with - | _: {{ ⟦ A ⟧ p ↦ m' ↘ ~?m }}, _: {{ DG ~?m ∈ glu_univ_elem i ↘ ?P ↘ ?El }} |- _ => + | _: {{ ⟦ A ⟧ ρ ↦ m' ↘ ~?m }}, _: {{ DG ~?m ∈ glu_univ_elem i ↘ ?P ↘ ?El }} |- _ => rename m into am'; rename P into P'; rename El into El' @@ -268,9 +267,9 @@ Proof. assert {{ Δ, ℕ, A[q σ] ⊢s Wk∘Wk,,succ #1 : Δ, ℕ }} by mauto 3. assert {{ Δ, ℕ, A[q σ] ⊢ MS[q (q σ)] : A[q σ][Wk∘Wk,,succ #1] }} by (rewrite @exp_eq_typ_q_sigma_then_weak_weak_extend_succ_var_1; mauto 3). pose (R := {{{ rec M' return A[q σ] | zero -> MZ[σ] | succ -> MS[q (q σ)] end }}}). - assert (exists r, {{ rec m' ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} /\ {{ Δ ⊢ R : A[σ,,M'] ® r ∈ El' }}) as [r' []] by mauto 3. + assert (exists r, {{ rec m' ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} /\ {{ Δ ⊢ R : A[σ,,M'] ® r ∈ El' }}) as [r' []] by mauto 3. assert {{ Δ ⊢ R : A[σ,,M'] }} by (erewrite <- @exp_eq_elim_sub_rhs_typ; mauto 3). - assert {{ Δ ⊢s σ,,M',,R ® p ↦ m' ↦ r' ∈ SbΓℕA }} by (unfold SbΓℕA; mauto 3). + assert {{ Δ ⊢s σ,,M',,R ® ρ ↦ m' ↦ r' ∈ SbΓℕA }} by (unfold SbΓℕA; mauto 3). destruct_glu_rel_exp_with_sub. simplify_evals. match_by_head glu_univ_elem ltac:(fun H => directed invert_glu_univ_elem H). @@ -280,7 +279,7 @@ Proof. destruct_conjs. handle_functional_glu_univ_elem. match goal with - | _: {{ ⟦ MS ⟧ p ↦ m' ↦ r' ↘ ~?m }} |- _ => + | _: {{ ⟦ MS ⟧ ρ ↦ m' ↦ r' ↘ ~?m }} |- _ => rename m into ms end. exists ms; split; mauto 3. @@ -349,7 +348,7 @@ Qed. #[local] Hint Resolve cons_glu_sub_pred_q_nat_helper : mcltt. -Lemma glu_rel_exp_natrec_neut_helper : forall {i Γ SbΓ A MZ MS Δ M m σ p am P El}, +Lemma glu_rel_exp_natrec_neut_helper : forall {i Γ SbΓ A MZ MS Δ M m σ ρ am P El}, {{ EG Γ ∈ glu_ctx_env ↘ SbΓ }} -> {{ Γ, ℕ ⊩ A : Type@i }} -> {{ Γ ⊩ A[Id,,zero] : Type@i }} -> @@ -358,11 +357,11 @@ Lemma glu_rel_exp_natrec_neut_helper : forall {i Γ SbΓ A MZ MS Δ M m σ p am {{ Γ, ℕ, A ⊩ MS : A[Wk∘Wk,,succ #1] }} -> {{ Dom m ≈ m ∈ per_bot }} -> (forall Δ' τ V, {{ Δ' ⊢w τ : Δ }} -> {{ Rne m in length Δ' ↘ V }} -> {{ Δ' ⊢ M[τ] ≈ V : ℕ }}) -> - {{ Δ ⊢s σ ® p ∈ SbΓ }} -> - {{ ⟦ A ⟧ p ↦ ⇑ ℕ m ↘ am }} -> + {{ Δ ⊢s σ ® ρ ∈ SbΓ }} -> + {{ ⟦ A ⟧ ρ ↦ ⇑ ℕ m ↘ am }} -> {{ DG am ∈ glu_univ_elem i ↘ P ↘ El }} -> exists r, - {{ rec ⇑ ℕ m ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} /\ + {{ rec ⇑ ℕ m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} /\ {{ Δ ⊢ rec M return A[q σ] | zero -> MZ[σ] | succ -> MS[q (q σ)] end : A[σ,,M] ® r ∈ El }}. Proof. intros * ? HA ? HMZ ? HMS **. @@ -377,7 +376,7 @@ Proof. invert_glu_rel_exp HA. pose (SbΓℕA := cons_glu_sub_pred i {{{ Γ, ℕ }}} A SbΓℕ). assert {{ EG Γ, ℕ, A ∈ glu_ctx_env ↘ SbΓℕA }} by (econstructor; mauto 3; reflexivity). - assert {{ Δ ⊢s σ,,M ® p ↦ ⇑ ℕ m ∈ SbΓℕ }} by (unfold SbΓℕ; mauto 3). + assert {{ Δ ⊢s σ,,M ® ρ ↦ ⇑ ℕ m ∈ SbΓℕ }} by (unfold SbΓℕ; mauto 3). assert {{ Γ, ℕ, A ⊢ MS : A[Wk∘Wk,,succ #1] }} by mauto 2. invert_glu_rel_exp HMS. destruct_glu_rel_exp_with_sub. @@ -388,8 +387,8 @@ Proof. destruct_conjs. handle_functional_glu_univ_elem. match goal with - | _: {{ ⟦ MZ ⟧ ~?p0 ↘ ~?m }}, _: {{ ⟦ A ⟧ ~?p0 ↦ zero ↘ ~?a }} |- _ => - rename p0 into p; + | _: {{ ⟦ MZ ⟧ ~?ρ0 ↘ ~?m }}, _: {{ ⟦ A ⟧ ~?ρ0 ↦ zero ↘ ~?a }} |- _ => + rename ρ0 into ρ; rename m into mz; rename a into az end. @@ -414,7 +413,7 @@ Proof. assert {{ Δ, ℕ, A[q σ] ⊢s Wk∘Wk,,succ #1 : Δ, ℕ }} by mauto 3. assert {{ Δ, ℕ, A[q σ] ⊢ MS[q (q σ)] : A[q σ][Wk∘Wk,,succ #1] }} by (rewrite @exp_eq_typ_q_sigma_then_weak_weak_extend_succ_var_1; mauto 3). pose (R := {{{ rec M return A[q σ] | zero -> MZ[σ] | succ -> MS[q (q σ)] end }}}). - enough {{ Δ ⊢ R : A[σ,,M] ® rec m under p return A | zero -> mz | succ -> MS end ∈ glu_elem_bot i am }} by (eapply realize_glu_elem_bot; mauto 3). + enough {{ Δ ⊢ R : A[σ,,M] ® rec m under ρ return A | zero -> mz | succ -> MS end ∈ glu_elem_bot i am }} by (eapply realize_glu_elem_bot; mauto 3). econstructor; mauto 3. - erewrite <- @exp_eq_elim_sub_rhs_typ; mauto 3. - assert {{ Δ ⊢ MZ[σ] : A[Id,,zero][σ] ® mz ∈ glu_elem_top i az }} as [] by (eapply realize_glu_elem_top; eassumption). @@ -427,21 +426,21 @@ Proof. match_by_head (per_ctx_env env_relΓℕA) ltac:(fun H => invert_per_ctx_env H). match_by_head (per_ctx_env env_relΓℕ) ltac:(fun H => invert_per_ctx_env H). intros s. - enough (exists r, {{ Rne rec m under p return A | zero -> mz | succ -> MS end in s ↘ r }}) as [] by (eexists; split; eassumption). - assert {{ Dom p ≈ p ∈ env_relΓ }} by (eapply glu_ctx_env_per_env; revgoals; eassumption). + enough (exists r, {{ Rne rec m under ρ return A | zero -> mz | succ -> MS end in s ↘ r }}) as [] by (eexists; split; eassumption). + assert {{ Dom ρ ≈ ρ ∈ env_relΓ }} by (eapply glu_ctx_env_per_env; revgoals; eassumption). destruct_rel_typ. invert_rel_typ_body. assert {{ Dom ! s ≈ ! s ∈ per_bot }} by mauto 3. - assert {{ Dom p ↦ ⇑! ℕ s ≈ p ↦ ⇑! ℕ s ∈ env_relΓℕ }} by (apply_relation_equivalence; unshelve eexists; simpl; intuition). - assert {{ Dom p ↦ succ ⇑! ℕ s ≈ p ↦ succ ⇑! ℕ s ∈ env_relΓℕ }} by (apply_relation_equivalence; unshelve eexists; simpl; intuition). + assert {{ Dom ρ ↦ ⇑! ℕ s ≈ ρ ↦ ⇑! ℕ s ∈ env_relΓℕ }} by (apply_relation_equivalence; unshelve eexists; simpl; intuition). + assert {{ Dom ρ ↦ succ ⇑! ℕ s ≈ ρ ↦ succ ⇑! ℕ s ∈ env_relΓℕ }} by (apply_relation_equivalence; unshelve eexists; simpl; intuition). destruct_rel_typ. invert_rel_typ_body. match goal with - | _: {{ ⟦ A ⟧ p ↦ ⇑! ℕ s ↘ ~?a }}, _: {{ ⟦ A ⟧ p ↦ (succ ⇑! ℕ s) ↘ ~?a' }} |- _ => + | _: {{ ⟦ A ⟧ ρ ↦ ⇑! ℕ s ↘ ~?a }}, _: {{ ⟦ A ⟧ ρ ↦ (succ ⇑! ℕ s) ↘ ~?a' }} |- _ => rename a into as'; (* We cannot use [as] as a name *) rename a' into asucc end. - assert {{ Dom p ↦ ⇑! ℕ s ↦ ⇑! as' (S s) ≈ p ↦ ⇑! ℕ s ↦ ⇑! as' (S s) ∈ env_relΓℕA }} as HΓℕA + assert {{ Dom ρ ↦ ⇑! ℕ s ↦ ⇑! as' (S s) ≈ ρ ↦ ⇑! ℕ s ↦ ⇑! as' (S s) ∈ env_relΓℕA }} as HΓℕA by (apply_relation_equivalence; unshelve eexists; simpl; intuition; eapply per_bot_then_per_elem; mauto 3). apply_relation_equivalence. (on_all_hyp_rev: fun H => destruct (H _ _ HΓℕA)). @@ -451,7 +450,7 @@ Proof. destruct_by_head rel_exp. functional_eval_rewrite_clear. match goal with - | _: {{ ⟦ MS ⟧ p ↦ ⇑! ℕ s ↦ ⇑! as' (S s) ↘ ~?m }} |- _ => + | _: {{ ⟦ MS ⟧ ρ ↦ ⇑! ℕ s ↦ ⇑! as' (S s) ↘ ~?m }} |- _ => rename m into ms end. assert {{ Dom as' ≈ as' ∈ per_top_typ }} as [? []]%(fun {a} (f : per_top_typ a a) => f (S s)) by mauto 3. @@ -465,8 +464,8 @@ Proof. assert {{ ⊢ Δ', ℕ }} by mauto 3. assert {{ Δ' ⊢s τ : Δ }} by mauto 3. assert {{ Δ' ⊢s σ∘τ : Γ }} by mauto 3. - assert {{ Δ' ⊢s σ∘τ ® p ∈ SbΓ }} by (eapply glu_ctx_env_sub_monotone; eassumption). - assert {{ Δ', ℕ ⊢s q (σ∘τ) ® p ↦ ⇑! ℕ (length Δ') ∈ SbΓℕ }} by (unfold SbΓℕ; mauto 3). + assert {{ Δ' ⊢s σ∘τ ® ρ ∈ SbΓ }} by (eapply glu_ctx_env_sub_monotone; eassumption). + assert {{ Δ', ℕ ⊢s q (σ∘τ) ® ρ ↦ ⇑! ℕ (length Δ') ∈ SbΓℕ }} by (unfold SbΓℕ; mauto 3). destruct_glu_rel_exp_with_sub. simplify_evals. match_by_head glu_univ_elem ltac:(fun H => directed invert_glu_univ_elem H). @@ -477,11 +476,11 @@ Proof. match_by_head read_ne ltac:(fun H => directed inversion_clear H). handle_functional_glu_univ_elem. match goal with - | _: {{ ⟦ A ⟧ ~?p' ↦ ⇑! ℕ (length Δ') ↘ ~?a }} |- _ => - rename p' into p; + | _: {{ ⟦ A ⟧ ~?ρ' ↦ ⇑! ℕ (length Δ') ↘ ~?a }} |- _ => + rename ρ' into ρ; rename a into aΔ' end. - assert {{ Δ', ℕ, A[q (σ∘τ)] ⊢s q (q (σ∘τ)) ® p ↦ ⇑! ℕ (length Δ') ↦ ⇑! aΔ' (length {{{ Δ', ℕ }}}) ∈ SbΓℕA }} + assert {{ Δ', ℕ, A[q (σ∘τ)] ⊢s q (q (σ∘τ)) ® ρ ↦ ⇑! ℕ (length Δ') ↦ ⇑! aΔ' (length {{{ Δ', ℕ }}}) ∈ SbΓℕA }} by (unfold SbΓℕA; mauto 3). destruct_glu_rel_exp_with_sub. simplify_evals. @@ -492,12 +491,12 @@ Proof. clear_dups. handle_functional_glu_univ_elem. match goal with - | _: {{ ⟦ A ⟧ ~?p' ↦ succ (⇑! ℕ (length Δ')) ↘ ~?a }}, + | _: {{ ⟦ A ⟧ ~?ρ' ↦ succ (⇑! ℕ (length Δ')) ↘ ~?a }}, _: {{ Rtyp aΔ' in S (length Δ') ↘ ~?A }}, _: {{ Rnf ⇓ az mz in length Δ' ↘ ~?MZ }}, _: {{ Rne m in length Δ' ↘ ~?M }} |- _ => rename A into A'; - rename p' into p; + rename ρ' into ρ; rename a into asucc; rename MZ into MZ'; rename M into M' @@ -574,12 +573,12 @@ Lemma glu_rel_exp_natrec_helper : forall {i Γ SbΓ A MZ MS}, {{ Γ, ℕ, A ⊩ MS : A[Wk∘Wk,,succ #1] }} -> forall {Δ M m}, glu_nat Δ M m -> - forall {σ p am P El}, - {{ Δ ⊢s σ ® p ∈ SbΓ }} -> - {{ ⟦ A ⟧ p ↦ m ↘ am }} -> + forall {σ ρ am P El}, + {{ Δ ⊢s σ ® ρ ∈ SbΓ }} -> + {{ ⟦ A ⟧ ρ ↦ m ↘ am }} -> {{ DG am ∈ glu_univ_elem i ↘ P ↘ El }} -> exists r, - {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ p ↘ r }} /\ + {{ rec m ⟦return A | zero -> MZ | succ -> MS end⟧ ρ ↘ r }} /\ {{ Δ ⊢ rec M return A[q σ] | zero -> MZ[σ] | succ -> MS[q (q σ)] end : A[σ,,M] ® r ∈ El }}. Proof. intros * ? HA ? ?. @@ -607,7 +606,7 @@ Proof. assert {{ Γ, ℕ, A ⊩ #1 : ℕ }} by mauto 3. mauto. } - induction 1; intros; rename m into M; rename Γ0 into Δ. + induction 1; intros; rename Γ0 into Δ. - (* glu_nat_zero *) mauto 4 using glu_rel_exp_natrec_zero_helper. - (* glu_nat_succ *) diff --git a/theories/Core/Soundness/Realizability.v b/theories/Core/Soundness/Realizability.v index a8564a35..ae0be94c 100644 --- a/theories/Core/Soundness/Realizability.v +++ b/theories/Core/Soundness/Realizability.v @@ -8,16 +8,16 @@ Import Domain_Notations. Open Scope list_scope. -Lemma wf_ctx_sub_ctx_lookup : forall n T Γ, - {{ #n : T ∈ Γ }} -> +Lemma wf_ctx_sub_ctx_lookup : forall n A Γ, + {{ #n : A ∈ Γ }} -> forall Δ, {{ ⊢ Δ ⊆ Γ}} -> - exists Δ1 T0 Δ2 T', - Δ = Δ1 ++ T0 :: Δ2 /\ + exists Δ1 A0 Δ2 A', + Δ = Δ1 ++ A0 :: Δ2 /\ n = length Δ1 /\ - T' = iter (S n) (fun T => {{{T [ Wk ]}}}) T0 /\ - {{ #n : T' ∈ Δ }} /\ - {{ Δ ⊢ T' ⊆ T }}. + A' = iter (S n) (fun A => {{{ A[Wk] }}}) A0 /\ + {{ #n : A' ∈ Δ }} /\ + {{ Δ ⊢ A' ⊆ A }}. Proof. induction 1; intros; progressive_inversion. - exists nil. @@ -27,9 +27,8 @@ Proof. repeat eexists; mauto 4. Qed. - -Lemma var_arith : forall Γ1 Γ2 (T : typ), - length (Γ1 ++ T :: Γ2) - length Γ2 - 1 = length Γ1. +Lemma var_arith : forall Γ1 Γ2 (A : typ), + length (Γ1 ++ A :: Γ2) - length Γ2 - 1 = length Γ1. Proof. intros. rewrite List.app_length. simpl. @@ -38,9 +37,9 @@ Qed. Lemma var_weaken_gen : forall Δ σ Γ, {{ Δ ⊢w σ : Γ }} -> - forall Γ1 Γ2 T, - Γ = Γ1 ++ T :: Γ2 -> - {{Δ ⊢ (# (length Γ1)) [σ] ≈ # (length Δ - length Γ2 - 1) : ~(iter (S (length Γ1)) (fun T => {{{T [ Wk ]}}}) T) [ σ ] }}. + forall Γ1 Γ2 A0, + Γ = Γ1 ++ A0 :: Γ2 -> + {{ Δ ⊢ #(length Γ1)[σ] ≈ #(length Δ - length Γ2 - 1) : ~(iter (S (length Γ1)) (fun A => {{{ A[Wk] }}}) A0)[σ] }}. Proof. induction 1; intros; subst; gen_presups. - pose proof (app_ctx_vlookup _ _ _ _ HΔ eq_refl) as Hvar. @@ -48,16 +47,16 @@ Proof. clear_dups. apply wf_sub_id_inversion in Hτ. pose proof (wf_ctx_sub_length _ _ Hτ). - transitivity {{{#(length Γ1) [Id]}}}; [mauto 3 |]. - replace (length Γ) with (length (Γ1 ++ {{{ Γ2, T }}})) by eassumption. + transitivity {{{ #(length Γ1)[Id] }}}; [mauto 3 |]. + replace (length Γ) with (length (Γ1 ++ {{{ Γ2, A0 }}})) by eassumption. rewrite var_arith, H. bulky_rewrite. - pose proof (app_ctx_vlookup _ _ _ _ HΔ0 eq_refl) as Hvar. - pose proof (app_ctx_lookup Γ1 T0 Γ2 _ eq_refl). + pose proof (app_ctx_lookup Γ1 A0 Γ2 _ eq_refl). gen_presup Hvar. clear_dups. - assert {{ ⊢ Δ', T }} by mauto 3. - assert {{ Δ', T ⊢s Wk : ~ (Γ1 ++ {{{ Γ2, T0 }}}) }} by mauto 3. + assert {{ ⊢ Δ', A }} by mauto 3. + assert {{ Δ', A ⊢s Wk : ~ (Γ1 ++ {{{ Γ2, A0 }}}) }} by mauto 3. transitivity {{{ #(length Γ1)[Wk∘τ] }}}; [mauto 3 |]. rewrite H1. etransitivity; [eapply wf_exp_eq_sub_compose; mauto 3 |]. @@ -76,16 +75,16 @@ Proof. * eapply wf_exp_eq_var_weaken; [mauto 3|]; eauto. * mauto 4. + eapply wf_exp_eq_subtyp'. - * eapply IHweakening with (Γ1 := T :: _). + * eapply IHweakening with (Γ1 := A :: _). reflexivity. * eapply wf_subtyp_subst; [ |mauto 3]. simpl. eapply wf_subtyp_subst; mauto 3. Qed. -Lemma var_glu_elem_bot : forall A i P El Γ T, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - {{ Γ ⊢ T ® P }} -> - {{ Γ, T ⊢ #0 : T[Wk] ® ! (length Γ) ∈ glu_elem_bot i A }}. +Lemma var_glu_elem_bot : forall a i P El Γ A, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + {{ Γ ⊢ A ® P }} -> + {{ Γ, A ⊢ #0 : A[Wk] ® !(length Γ) ∈ glu_elem_bot i a }}. Proof. intros. saturate_glu_info. econstructor; mauto 4. @@ -98,19 +97,26 @@ Qed. #[local] Hint Rewrite -> wf_sub_eq_extend_compose using mauto 4 : mcltt. -Theorem realize_glu_univ_elem_gen : forall A i P El, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - (forall Γ T R, {{ DF A ≈ A ∈ per_univ_elem i ↘ R }} -> - {{ Γ ⊢ T ® P }} -> - {{ Γ ⊢ T ® glu_typ_top i A }}) /\ - (forall Γ t T c, {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - {{ Γ ⊢ t : T ® c ∈ glu_elem_bot i A }} -> - {{ Γ ⊢ t : T ® ⇑ A c ∈ El }}) /\ - (forall Γ t T a R, {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - {{ Γ ⊢ t : T ® a ∈ El }} -> - {{ DF A ≈ A ∈ per_univ_elem i ↘ R }} -> - {{ Dom a ≈ a ∈ R }} -> - {{ Γ ⊢ t : T ® a ∈ glu_elem_top i A }}). +Theorem realize_glu_univ_elem_gen : forall a i P El, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + (forall Γ A R, + {{ DF a ≈ a ∈ per_univ_elem i ↘ R }} -> + {{ Γ ⊢ A ® P }} -> + {{ Γ ⊢ A ® glu_typ_top i a }}) /\ + (forall Γ M A m, + (* We repeat this to get the relation between [a] and [P] + more easily after applying [induction 1.] *) + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + {{ Γ ⊢ M : A ® m ∈ glu_elem_bot i a }} -> + {{ Γ ⊢ M : A ® ⇑ a m ∈ El }}) /\ + (forall Γ M A m R, + (* We repeat this to get the relation between [a] and [P] + more easily after applying [induction 1.] *) + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + {{ Γ ⊢ M : A ® m ∈ El }} -> + {{ DF a ≈ a ∈ per_univ_elem i ↘ R }} -> + {{ Dom m ≈ m ∈ R }} -> + {{ Γ ⊢ M : A ® m ∈ glu_elem_top i a }}). Proof. simpl. induction 1 using glu_univ_elem_ind. all:split; [| split]; intros; @@ -121,7 +127,7 @@ Proof. destruct_all. - econstructor; eauto; intros. progressive_inversion. - transitivity {{{Type@j [ σ ]}}}; mauto 4. + transitivity {{{ Type@j[σ] }}}; mauto 4. - handle_functional_glu_univ_elem. match_by_head glu_univ_elem invert_glu_univ_elem. clear_dups. @@ -139,9 +145,8 @@ Proof. - deepexec glu_univ_elem_per_univ ltac:(fun H => pose proof H). firstorder. specialize (H _ _ _ H10) as [? []]. - econstructor; eauto. + econstructor; mauto 3. + apply_equiv_left. trivial. - + mauto 2. + intros. saturate_weakening_escape. deepexec H ltac:(fun H => destruct H). @@ -152,7 +157,7 @@ Proof. - econstructor; eauto; intros. progressive_inversion. - transitivity {{{ℕ [ σ ]}}}; mauto 3. + transitivity {{{ ℕ[σ] }}}; mauto 3. - handle_functional_glu_univ_elem. match_by_head glu_univ_elem invert_glu_univ_elem. apply_equiv_left. @@ -161,13 +166,12 @@ Proof. intros. saturate_weakening_escape. - assert {{ Δ ⊢ T [σ] ≈ ℕ[σ] : Type @ i }} by mauto 3. + assert {{ Δ ⊢ A[σ] ≈ ℕ[σ] : Type @ i }} by mauto 3. rewrite <- wf_exp_eq_nat_sub; try eassumption. mauto 3. - - econstructor; eauto. - + rewrite H2. mauto 3. + - econstructor; mauto 3. + + bulky_rewrite. mauto 3. + apply_equiv_left. trivial. - + mauto 2. + intros. saturate_weakening_escape. bulky_rewrite. @@ -180,23 +184,23 @@ Proof. + gen_presups. trivial. + saturate_weakening_escape. assert {{ Γ ⊢w Id : Γ }} by mauto 4. - pose proof (H13 _ _ H16). - specialize (H13 _ _ H19). - assert {{ Γ ⊢ IT[Id] ≈ IT : Type@i}} by mauto 3. - bulky_rewrite_in H13. - progressive_invert H17. - destruct (H9 _ _ _ H0 H13) as []. + assert {{ Δ ⊢ IT[σ] ® IP }} by mauto 3. + assert (IP Γ {{{ IT[Id] }}}) as HITId by mauto 3. + bulky_rewrite_in HITId. + assert {{ Γ ⊢ IT[Id] ≈ IT : Type@i }} by mauto 3. + dir_inversion_clear_by_head read_typ. + assert {{ Γ ⊢ IT ® glu_typ_top i a }} as [] by mauto 3. bulky_rewrite. simpl. apply wf_exp_eq_pi_cong'; [firstorder |]. pose proof (var_per_elem (length Δ) H0). destruct_rel_mod_eval. simplify_evals. destruct (H2 _ ltac:(eassumption) _ ltac:(eassumption)) as [? []]. - specialize (H10 _ _ _ _ ltac:(trivial) (var_glu_elem_bot _ _ _ _ _ _ ltac:(eassumption) ltac:(eassumption))). - autorewrite with mcltt in H10. - specialize (H14 {{{Δ, IT[σ]}}} {{{σ ∘ Wk}}} _ _ ltac:(mauto) ltac:(eassumption) ltac:(eassumption)). + assert (IEl {{{ Δ, IT[σ] }}} {{{ IT[σ][Wk] }}} {{{ #0 }}} d{{{ ⇑! a (length Δ) }}}) by mauto 3 using var_glu_elem_bot. + autorewrite with mcltt in H31. + specialize (H14 {{{ Δ, IT[σ] }}} {{{ σ∘Wk }}} _ _ ltac:(mauto) ltac:(eassumption) ltac:(eassumption)). specialize (H8 _ _ _ ltac:(eassumption) ltac:(eassumption)) as []. - etransitivity; [| eapply H30]; mauto 3. + etransitivity; [| eapply H33]; mauto 3. - handle_functional_glu_univ_elem. apply_equiv_left. invert_glu_rel1. @@ -210,18 +214,18 @@ Proof. simplify_evals. eexists; repeat split; mauto 3. eapply H2; eauto. - assert {{ Δ ⊢ t [σ] : M[σ] }} by mauto 3. + assert {{ Δ ⊢ M[σ] : A[σ] }} by mauto 3. bulky_rewrite_in H23. unshelve (econstructor; eauto). + trivial. - + eassert {{ Δ ⊢ ~ (a_sub t σ) m' : ~_ }} by (eapply wf_app'; eassumption). + + eassert {{ Δ ⊢ M[σ] N : ~_ }} by (eapply wf_app'; eassumption). autorewrite with mcltt in H25. trivial. + mauto using domain_app_per. + intros. saturate_weakening_escape. progressive_invert H26. - destruct (H15 _ _ _ _ _ ltac:(eassumption) ltac:(eassumption) ltac:(eassumption) equiv_b). + destruct (H15 _ _ _ _ _ ltac:(eassumption) ltac:(eassumption) ltac:(eassumption) equiv_n). handle_functional_glu_univ_elem. autorewrite with mcltt. @@ -232,7 +236,7 @@ Proof. rewrite <- @exp_eq_sub_compose_typ; mauto 3; [eapply wf_exp_eq_app_cong' |]. -- specialize (H12 _ {{{σ ∘ σ0}}} _ ltac:(mauto 3) ltac:(eassumption)). - rewrite wf_exp_eq_sub_compose with (M := t) in H12; mauto 3. + rewrite wf_exp_eq_sub_compose with (M := M) in H12; mauto 3. bulky_rewrite_in H12. -- rewrite <- @exp_eq_sub_compose_typ; mauto 3. -- econstructor; mauto 3. @@ -254,15 +258,15 @@ Proof. assert {{ Γ ⊢w Id : Γ }} by mauto 4. pose proof (H10 _ _ H24). specialize (H10 _ _ H19). - assert {{ Γ ⊢ IT[Id] ≈ IT : Type@i}} by mauto 3. + assert {{ Γ ⊢ IT[Id] ≈ IT : Type@i }} by mauto 3. bulky_rewrite_in H25. destruct (H11 _ _ _ ltac:(eassumption) ltac:(eassumption)) as []. specialize (H29 _ _ _ H19 H9). rewrite H5 in *. autorewrite with mcltt. - eassert {{ Δ ⊢ m[σ] : ~_ }} by (mauto 2). + eassert {{ Δ ⊢ M[σ] : ~_ }} by (mauto 2). autorewrite with mcltt in H30. - rewrite @wf_exp_eq_pi_eta' with (M := {{{m[σ]}}}); [| trivial]. + rewrite @wf_exp_eq_pi_eta' with (M := {{{ M[σ] }}}); [| trivial]. cbn [nf_to_exp]. eapply wf_exp_eq_fn_cong'; eauto. @@ -282,7 +286,7 @@ Proof. do 2 (rewrite wf_exp_eq_sub_id in H40; mauto 3). etransitivity; [|eassumption]. simpl. - assert {{ Δ, IT[σ] ⊢ # 0 : IT[σ ∘ Wk] }} by (rewrite <- @exp_eq_sub_compose_typ; mauto 3). + assert {{ Δ, IT[σ] ⊢ #0 : IT[σ∘Wk] }} by (rewrite <- @exp_eq_sub_compose_typ; mauto 3). rewrite <- @sub_eq_q_sigma_id_extend; mauto 4. rewrite <- @exp_eq_sub_compose_typ; mauto 2. 2:eapply sub_q; mauto 4. @@ -309,12 +313,11 @@ Proof. firstorder. Qed. - -Corollary realize_glu_typ_top : forall A i P El, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ T, - {{ Γ ⊢ T ® P }} -> - {{ Γ ⊢ T ® glu_typ_top i A }}. +Corollary realize_glu_typ_top : forall a i P El, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ A, + {{ Γ ⊢ A ® P }} -> + {{ Γ ⊢ A ® glu_typ_top i a }}. Proof. intros. pose proof H. @@ -323,21 +326,21 @@ Proof. eapply realize_glu_univ_elem_gen; eauto. Qed. -Theorem realize_glu_elem_bot : forall A i P El, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ t T c, - {{ Γ ⊢ t : T ® c ∈ glu_elem_bot i A }} -> - {{ Γ ⊢ t : T ® ⇑ A c ∈ El }}. +Theorem realize_glu_elem_bot : forall a i P El, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ A M m, + {{ Γ ⊢ M : A ® m ∈ glu_elem_bot i a }} -> + {{ Γ ⊢ M : A ® ⇑ a m ∈ El }}. Proof. intros. eapply realize_glu_univ_elem_gen; eauto. Qed. -Theorem realize_glu_elem_top : forall A i P El, - {{ DG A ∈ glu_univ_elem i ↘ P ↘ El }} -> - forall Γ t T a, - {{ Γ ⊢ t : T ® a ∈ El }} -> - {{ Γ ⊢ t : T ® a ∈ glu_elem_top i A }}. +Theorem realize_glu_elem_top : forall a i P El, + {{ DG a ∈ glu_univ_elem i ↘ P ↘ El }} -> + forall Γ A M m, + {{ Γ ⊢ M : A ® m ∈ El }} -> + {{ Γ ⊢ M : A ® m ∈ glu_elem_top i a }}. Proof. intros. pose proof H. diff --git a/theories/Core/Soundness/SubstitutionCases.v b/theories/Core/Soundness/SubstitutionCases.v index d52b7788..5fd963f7 100644 --- a/theories/Core/Soundness/SubstitutionCases.v +++ b/theories/Core/Soundness/SubstitutionCases.v @@ -134,10 +134,9 @@ Proof. assert {{ Δ0 ⊢s Wk∘((σ,,M)∘σ0) ≈ σ∘σ0 : Δ }} by mauto 3. econstructor; mauto 4. + assert {{ Δ, A ⊢s Wk : Δ }} by mauto 4. - assert {{ Δ0 ⊢ A[Wk][(σ,,M)∘σ0] ≈ A[Wk∘((σ,,M)∘σ0)] : Type@i }} by mauto 4. - assert {{ Δ0 ⊢ A[Wk∘((σ,,M)∘σ0)] ≈ A[σ∘σ0] : Type@i }} by mauto 3. - assert {{ Δ0 ⊢ A[σ∘σ0] ≈ A[σ][σ0] : Type@i }} by mauto 4. - assert {{ Δ0 ⊢ A[Wk∘((σ,,M)∘σ0)] ≈ A[σ][σ0] : Type@i }} as -> by mauto 3. + assert {{ Δ0 ⊢ A[Wk][(σ,,M)∘σ0] ≈ A[Wk][(σ∘σ0),,M[σ0]] : Type@i }} as -> by mauto 3. + assert {{ Δ0 ⊢ A[Wk][(σ∘σ0),,M[σ0]] ≈ A[σ∘σ0] : Type@i }} as -> by mauto 3. + assert {{ Δ0 ⊢ A[σ∘σ0] ≈ A[σ][σ0] : Type@i }} as -> by mauto 3. assert {{ Δ, A ⊢ #0 : A[Wk] }} by mauto 3. assert {{ Δ0 ⊢ #0[(σ,,M)∘σ0] ≈ #0[(σ∘σ0),,M[σ0]] : A[Wk][(σ,,M)∘σ0] }} by mauto 3. assert {{ Δ0 ⊢ #0[(σ,,M)∘σ0] ≈ #0[(σ∘σ0),,M[σ0]] : A[σ][σ0] }} as -> by mauto 3. diff --git a/theories/Core/Soundness/TermStructureCases.v b/theories/Core/Soundness/TermStructureCases.v index deb89c73..b0467bb1 100644 --- a/theories/Core/Soundness/TermStructureCases.v +++ b/theories/Core/Soundness/TermStructureCases.v @@ -53,9 +53,6 @@ Proof. intros. destruct_by_head cons_glu_sub_pred. econstructor; mauto. - enough {{ Δ ⊢ A[Wk][σ] ≈ A[Wk∘σ] : Type@i }} as -> by eassumption. - assert {{ Γ, A ⊢s Wk : Γ }} by mauto 3. - mauto 3. - assert {{ Γ ⊩ #n : A }} as Hn by mauto. assert (exists i, {{ Γ ⊢ A : Type@i }}) as [j] by (gen_presups; mauto 3). invert_glu_rel_exp Hn. diff --git a/theories/Core/Soundness/Weakening/Definition.v b/theories/Core/Soundness/Weakening/Definition.v index 2acabf81..46a678ea 100644 --- a/theories/Core/Soundness/Weakening/Definition.v +++ b/theories/Core/Soundness/Weakening/Definition.v @@ -10,7 +10,7 @@ Inductive weakening : ctx -> sub -> ctx -> Prop := `( {{ Γ ⊢s σ ≈ Id : Δ }} -> {{ Γ ⊢w σ : Δ }} ) | wk_p : - `( {{ Γ ⊢w τ : Δ', T }} -> + `( {{ Γ ⊢w τ : Δ', A }} -> {{ ⊢ Δ' ⊆ Δ }} -> {{ Γ ⊢s σ ≈ Wk ∘ τ : Δ }} -> {{ Γ ⊢w σ : Δ }} ) diff --git a/theories/Core/Soundness/Weakening/Lemmas.v b/theories/Core/Soundness/Weakening/Lemmas.v index bcdaa1e1..88cce9de 100644 --- a/theories/Core/Soundness/Weakening/Lemmas.v +++ b/theories/Core/Soundness/Weakening/Lemmas.v @@ -93,9 +93,9 @@ Proof. mauto. Qed. -Lemma weakening_wk : forall Γ T, - {{ ⊢ Γ , T }} -> - {{ Γ , T ⊢w Wk : Γ }}. +Lemma weakening_wk : forall Γ A, + {{ ⊢ Γ, A }} -> + {{ Γ, A ⊢w Wk : Γ }}. Proof. intros. econstructor; mautosolve. diff --git a/theories/Extraction/TypeCheck.v b/theories/Extraction/TypeCheck.v index 317e840c..316e978e 100644 --- a/theories/Extraction/TypeCheck.v +++ b/theories/Extraction/TypeCheck.v @@ -50,7 +50,13 @@ Section type_check. | _ => inright _ . - Extraction Inline get_level_of_type_nf. + (* Don't forget to use 9th bit of [Extraction Flag] (for example, [Set Extraction Flag 1007.]) *) + Equations get_subterms_of_type_pi (A : nf) : { B & { C | A = n{{{ Π B C }}} } } + { forall B C, A <> n{{{ Π B C }}} } := + | n{{{ Π B C }}} => pureo (existT _ B (exist _ C _)) + | _ => inright _ + . + + Extraction Inline get_level_of_type_nf get_subterms_of_type_pi. Inductive type_check_order : exp -> Prop := | tc_ti : forall {A}, type_infer_order A -> type_check_order A @@ -153,8 +159,6 @@ Section type_check. | _ => idtac end. - Definition inspect {A : Set} (a : A) : { b : A | a = b } := (exist _ a eq_refl). - #[tactic="impl_obl_tac",derive(equations=no,eliminator=no)] Equations type_check G A (HA : (exists i, {{ G ⊢ A : Type@i }})) M (H : type_check_order M) : { {{ G ⊢a M ⟸ A }} } + { ~ {{ G ⊢a M ⟸ A }} } by struct H := | G, A, HA, M, H => @@ -192,13 +196,12 @@ Section type_check. let*o (exist _ B' _) := type_infer {{{ G, A' }}} _ M' _ while _ in let (A'', _) := nbe_ty_impl G A' _ in pureo (exist _ n{{{ Π A'' B' }}} _) - | {{{ M' N' }}} with type_infer G _ M' _ => { - | inleft (exist _ (nf_pi A B) _) => - let*b->o _ := type_check G A _ N' _ while _ in - let (B', _) := nbe_ty_impl G {{{ B[Id,,N'] }}} _ in - pureo (exist _ B' _) - | _ => inright _ - } + | {{{ M' N' }}} => + let*o (exist _ C _) := type_infer G _ M' _ while _ in + let*o (existT _ A (exist _ B _)) := get_subterms_of_type_pi C while _ in + let*b->o _ := type_check G A _ N' _ while _ in + let (B', _) := nbe_ty_impl G {{{ B[Id,,N'] }}} _ in + pureo (exist _ B' _) | {{{ #x }}} => let*o (exist _ A _) := lookup G _ x while _ in let (A', _) := nbe_ty_impl G A _ in @@ -302,9 +305,9 @@ Section type_check. functional_alg_type_infer_rewrite_clear. progressive_inversion. assert {{ G ⊢ A : ~n{{{ Type@i }}} }} by mauto 4 using alg_type_infer_sound. - assert {{ G, ~(A : exp) ⊢ B : ~n{{{ Type@j }}} }} by mauto 4 using alg_type_infer_sound. + assert {{ G, ~(A : exp) ⊢ s : ~n{{{ Type@j }}} }} by mauto 4 using alg_type_infer_sound. assert {{ G ⊢ N' : A }} by mauto 3 using alg_type_check_sound. - assert {{ G ⊢ B[Id,,N'] : ~n{{{ Type@j }}} }} as [? []]%soundness_ty by mauto 3. + assert {{ G ⊢ s[Id,,N'] : ~n{{{ Type@j }}} }} as [? []]%soundness_ty by mauto 3. mauto 3 using nbe_ty_order_sound. Qed. @@ -314,20 +317,13 @@ Section type_check. progressive_inversion. split; [mauto 3 |]. assert {{ G ⊢ A : ~n{{{ Type@i }}} }} by mauto 4 using alg_type_infer_sound. - assert {{ G, ~(A : exp) ⊢ B : ~n{{{ Type@j }}} }} by mauto 4 using alg_type_infer_sound. - assert {{ G ⊢ B[Id,,N'] ≈ B' : Type@j }} by (eapply soundness_ty'; mauto 4 using alg_type_check_sound). + assert {{ G, ~(A : exp) ⊢ s : ~n{{{ Type@j }}} }} by mauto 4 using alg_type_infer_sound. + assert {{ G ⊢ s[Id,,N'] ≈ B' : Type@j }} by (eapply soundness_ty'; mauto 4 using alg_type_check_sound). assert (user_exp B') by trivial using user_exp_nf. assert (exists k, {{ G ⊢a B' ⟹ Type@k }} /\ k <= j) as [? []] by (gen_presups; mauto 3). firstorder. Qed. - Next Obligation. (* False *) - clear_defs. - functional_alg_type_infer_rewrite_clear. - progressive_inversion. - intuition. - Qed. - Next Obligation. (* {{ ⊢ G, B }} *) clear_defs. assert {{ G ⊢ B : Type@i }} by mauto 4 using alg_type_infer_sound.