diff --git a/theories/Core/Semantic/PERDefinitions.v b/theories/Core/Semantic/PERDefinitions.v index e0ac8419..de2e898d 100644 --- a/theories/Core/Semantic/PERDefinitions.v +++ b/theories/Core/Semantic/PERDefinitions.v @@ -77,7 +77,9 @@ Inductive per_ne : relation domain := (** Universe/Element PER Definition *) Section Per_univ_elem_core_def. - Variable (i : nat) (per_univ_rec : forall {j}, j < i -> relation domain). + Variable + (i : nat) + (per_univ_rec : forall {j}, j < i -> relation domain). Inductive per_univ_elem_core : domain -> domain -> relation domain -> Prop := | per_univ_elem_core_univ : @@ -100,28 +102,20 @@ Section Per_univ_elem_core_def. . Hypothesis - (motive : domain -> domain -> relation domain -> Prop). - - Hypothesis - (case_U : forall (j j' : nat) (lt_j_i : j < i), j = j' -> motive d{{{ 𝕌@j }}} d{{{ 𝕌@j' }}} (per_univ_rec lt_j_i)). - - Hypothesis - (case_nat : motive d{{{ ℕ }}} d{{{ ℕ }}} per_nat). - - Hypothesis - (case_Pi : - forall {A p B A' p' B' in_rel elem_rel} - (out_rel : forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), relation domain), - {{ DF A ≈ A' ∈ per_univ_elem_core ↘ in_rel }} -> - motive A A' in_rel -> - PER in_rel -> - (forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), - rel_mod_eval (fun x y R => {{ DF x ≈ y ∈ per_univ_elem_core ↘ R }} /\ motive x y R) B d{{{ p ↦ c }}} B' d{{{ p' ↦ c' }}} (out_rel equiv_c_c')) -> - (forall f f', elem_rel f f' = forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_mod_app (out_rel equiv_c_c') f c f' c') -> - motive d{{{ Π A p B }}} d{{{ Π A' p' B' }}} elem_rel). - - Hypothesis - (case_ne : (forall {a b a' b'}, {{ Dom b ≈ b' ∈ per_bot }} -> motive d{{{ ⇑ a b }}} d{{{ ⇑ a' b' }}} per_ne)). + (motive : domain -> domain -> relation domain -> Prop) + (case_U : forall (j j' : nat) (lt_j_i : j < i), j = j' -> motive d{{{ 𝕌@j }}} d{{{ 𝕌@j' }}} (per_univ_rec lt_j_i)) + (case_nat : motive d{{{ ℕ }}} d{{{ ℕ }}} per_nat) + (case_Pi : + forall {A p B A' p' B' in_rel elem_rel} + (out_rel : forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), relation domain), + {{ DF A ≈ A' ∈ per_univ_elem_core ↘ in_rel }} -> + motive A A' in_rel -> + PER in_rel -> + (forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), + rel_mod_eval (fun x y R => {{ DF x ≈ y ∈ per_univ_elem_core ↘ R }} /\ motive x y R) B d{{{ p ↦ c }}} B' d{{{ p' ↦ c' }}} (out_rel equiv_c_c')) -> + (forall f f', elem_rel f f' = forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_mod_app (out_rel equiv_c_c') f c f' c') -> + motive d{{{ Π A p B }}} d{{{ Π A' p' B' }}} elem_rel) + (case_ne : (forall {a b a' b'}, {{ Dom b ≈ b' ∈ per_bot }} -> motive d{{{ ⇑ a b }}} d{{{ ⇑ a' b' }}} per_ne)). #[derive(equations=no, eliminator=no)] Equations per_univ_elem_core_strong_ind a b R (H : {{ DF a ≈ b ∈ per_univ_elem_core ↘ R }}) : {{ DF a ≈ b ∈ motive ↘ R }} := @@ -163,30 +157,22 @@ Global Hint Resolve per_univ_elem_core_univ' : mcltt. Section Per_univ_elem_ind_def. Hypothesis - (motive : nat -> domain -> domain -> relation domain -> Prop). - - Hypothesis - (case_U : forall j j' i, j < i -> j = j' -> - (forall A B R, {{ DF A ≈ B ∈ per_univ_elem j ↘ R }} -> motive j A B R) -> - motive i d{{{ 𝕌@j }}} d{{{ 𝕌@j' }}} (per_univ j)). - - Hypothesis - (case_N : forall i, motive i d{{{ ℕ }}} d{{{ ℕ }}} per_nat). - - Hypothesis - (case_Pi : - forall i {A p B A' p' B' in_rel elem_rel} - (out_rel : forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), relation domain), - {{ DF A ≈ A' ∈ per_univ_elem i ↘ in_rel }} -> - motive i A A' in_rel -> - PER in_rel -> - (forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), - rel_mod_eval (fun x y R => {{ DF x ≈ y ∈ per_univ_elem i ↘ R }} /\ motive i x y R) B d{{{ p ↦ c }}} B' d{{{ p' ↦ c' }}} (out_rel equiv_c_c')) -> - (forall f f', elem_rel f f' = forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_mod_app (out_rel equiv_c_c') f c f' c') -> - motive i d{{{ Π A p B }}} d{{{ Π A' p' B' }}} elem_rel). - - Hypothesis - (case_ne : (forall i {a b a' b'}, {{ Dom b ≈ b' ∈ per_bot }} -> motive i d{{{ ⇑ a b }}} d{{{ ⇑ a' b' }}} per_ne)). + (motive : nat -> domain -> domain -> relation domain -> Prop) + (case_U : forall j j' i, j < i -> j = j' -> + (forall A B R, {{ DF A ≈ B ∈ per_univ_elem j ↘ R }} -> motive j A B R) -> + motive i d{{{ 𝕌@j }}} d{{{ 𝕌@j' }}} (per_univ j)) + (case_N : forall i, motive i d{{{ ℕ }}} d{{{ ℕ }}} per_nat) + (case_Pi : + forall i {A p B A' p' B' in_rel elem_rel} + (out_rel : forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), relation domain), + {{ DF A ≈ A' ∈ per_univ_elem i ↘ in_rel }} -> + motive i A A' in_rel -> + PER in_rel -> + (forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), + rel_mod_eval (fun x y R => {{ DF x ≈ y ∈ per_univ_elem i ↘ R }} /\ motive i x y R) B d{{{ p ↦ c }}} B' d{{{ p' ↦ c' }}} (out_rel equiv_c_c')) -> + (forall f f', elem_rel f f' = forall {c c'} (equiv_c_c' : {{ Dom c ≈ c' ∈ in_rel }}), rel_mod_app (out_rel equiv_c_c') f c f' c') -> + motive i d{{{ Π A p B }}} d{{{ Π A' p' B' }}} elem_rel) + (case_ne : (forall i {a b a' b'}, {{ Dom b ≈ b' ∈ per_bot }} -> motive i d{{{ ⇑ a b }}} d{{{ ⇑ a' b' }}} per_ne)). #[local] Ltac def_simp := simp per_univ_elem in *. @@ -225,18 +211,21 @@ Ltac per_univ_elem_econstructor := 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 per_total : relation env := fun p p' => True. + Inductive per_ctx_env : ctx -> ctx -> relation env -> Prop := | per_ctx_env_nil : - `{ (Env = fun p p' => True) -> - {{ EF ⋅ ≈ ⋅ ∈ per_ctx_env ↘ Env }} } + `{ {{ EF ⋅ ≈ ⋅ ∈ per_ctx_env ↘ per_total }} } | per_ctx_env_cons : `{ forall (tail_rel : relation env) - (head_rel : forall {p p'}, {{ Dom p ≈ p' ∈ tail_rel }} -> relation domain) + (head_rel : forall {p p'} (equiv_p_p' : {{ Dom p ≈ p' ∈ tail_rel }}), relation domain) (equiv_Γ_Γ' : {{ 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 = 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 }}) -> - {{ EF Γ, A ≈ Γ', A' ∈ per_ctx_env ↘ Env }} } + 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 }}) -> + {{ EF Γ, A ≈ Γ', A' ∈ per_ctx_env ↘ env_rel }} } . Definition per_ctx : relation ctx := fun Γ Γ' => exists R', per_ctx_env Γ Γ' R'. diff --git a/theories/Core/Semantic/PERLemmas.v b/theories/Core/Semantic/PERLemmas.v index 0448818f..57778502 100644 --- a/theories/Core/Semantic/PERLemmas.v +++ b/theories/Core/Semantic/PERLemmas.v @@ -2,47 +2,10 @@ From Coq Require Import Lia PeanoNat Relation_Definitions RelationClasses. From Mcltt Require Import Axioms Base Evaluation LibTactics PERDefinitions Readback. Import Domain_Notations. -(* Lemma rel_mod_eval_ex_pull : *) -(* forall (A : Type) (P : domain -> domain -> relation domain -> A -> Prop) {T p T' p'} R, *) -(* rel_mod_eval (fun a b R => exists x : A, P a b R x) T p T' p' R <-> *) -(* exists x : A, rel_mod_eval (fun a b R => P a b R x) T p T' p' R. *) -(* Proof. *) -(* split; intros. *) -(* - destruct H. *) -(* destruct H1 as [? ?]. *) -(* eexists; econstructor; eauto. *) -(* - do 2 destruct H; econstructor; eauto. *) -(* Qed. *) - -(* Lemma rel_mod_eval_simp_ex : *) -(* forall (A : Type) (P : domain -> domain -> relation domain -> Prop) (Q : domain -> domain -> relation domain -> A -> Prop) {T p T' p'} R, *) -(* rel_mod_eval (fun a b R => P a b R /\ exists x : A, Q a b R x) T p T' p' R <-> *) -(* exists x : A, rel_mod_eval (fun a b R => P a b R /\ Q a b R x) T p T' p' R. *) -(* Proof. *) -(* split; intros. *) -(* - destruct H. *) -(* destruct H1 as [? [? ?]]. *) -(* eexists; econstructor; eauto. *) -(* - do 2 destruct H; econstructor; eauto. *) -(* firstorder. *) -(* Qed. *) - -(* Lemma rel_mod_eval_simp_and : *) -(* forall (P : domain -> domain -> relation domain -> Prop) (Q : relation domain -> Prop) {T p T' p'} R, *) -(* rel_mod_eval (fun a b R => P a b R /\ Q R) T p T' p' R <-> *) -(* rel_mod_eval P T p T' p' R /\ Q R. *) -(* Proof. *) -(* split; intros. *) -(* - destruct H. *) -(* destruct H1 as [? ?]. *) -(* split; try econstructor; eauto. *) -(* - do 2 destruct H; econstructor; eauto. *) -(* Qed. *) - Lemma per_bot_sym : forall m n, {{ Dom m ≈ n ∈ per_bot }} -> {{ Dom n ≈ m ∈ per_bot }}. -Proof with solve [mauto]. +Proof with solve [eauto]. intros * H s. destruct (H s) as [? []]... Qed. @@ -54,7 +17,7 @@ Lemma per_bot_trans : forall m n l, {{ Dom m ≈ n ∈ per_bot }} -> {{ Dom n ≈ l ∈ per_bot }} -> {{ Dom m ≈ l ∈ per_bot }}. -Proof with solve [mauto]. +Proof with solve [eauto]. intros * Hmn Hnl s. destruct (Hmn s) as [? []]. destruct (Hnl s) as [? []]. @@ -76,7 +39,7 @@ Hint Resolve var_per_bot : mcltt. Lemma per_top_sym : forall m n, {{ Dom m ≈ n ∈ per_top }} -> {{ Dom n ≈ m ∈ per_top }}. -Proof with solve [mauto]. +Proof with solve [eauto]. intros * H s. destruct (H s) as [? []]... Qed. @@ -88,7 +51,7 @@ Lemma per_top_trans : forall m n l, {{ Dom m ≈ n ∈ per_top }} -> {{ Dom n ≈ l ∈ per_top }} -> {{ Dom m ≈ l ∈ per_top }}. -Proof with solve [mauto]. +Proof with solve [eauto]. intros * Hmn Hnl s. destruct (Hmn s) as [? []]. destruct (Hnl s) as [? []]. @@ -113,7 +76,7 @@ Hint Resolve per_bot_then_per_top : mcltt. Lemma per_top_typ_sym : forall m n, {{ Dom m ≈ n ∈ per_top_typ }} -> {{ Dom n ≈ m ∈ per_top_typ }}. -Proof with solve [mauto]. +Proof with solve [eauto]. intros * H s. destruct (H s) as [? []]... Qed. @@ -125,7 +88,7 @@ Lemma per_top_typ_trans : forall m n l, {{ Dom m ≈ n ∈ per_top_typ }} -> {{ Dom n ≈ l ∈ per_top_typ }} -> {{ Dom m ≈ l ∈ per_top_typ }}. -Proof with solve [mauto]. +Proof with solve [eauto]. intros * Hmn Hnl s. destruct (Hmn s) as [? []]. destruct (Hnl s) as [? []]. @@ -186,18 +149,19 @@ Qed. #[export] Hint Resolve per_ne_trans : mcltt. -Lemma per_univ_elem_right_irrel_gen : forall i A B R, +Lemma per_univ_elem_right_irrel : forall i i' A B R B' R', per_univ_elem i A B R -> - forall A' B' R', - A = A' -> - per_univ_elem i A' B' R' -> - R = R'. + per_univ_elem i' A B' R' -> + R = R'. Proof. - induction 1 using per_univ_elem_ind; intros * Heq Hright; + intros * Horig. + remember A as A' in |- *. + gen A' B' R'. + induction Horig using per_univ_elem_ind; intros * Heq Hright; try solve [induction Hright using per_univ_elem_ind; congruence]. subst. invert_per_univ_elem Hright; try congruence. - specialize (IHper_univ_elem _ _ _ eq_refl equiv_a_a'). + specialize (IHHorig _ _ _ eq_refl equiv_a_a'). subst. extensionality f. extensionality f'. @@ -211,80 +175,74 @@ Proof. congruence. Qed. -Lemma per_univ_elem_right_irrel : forall i A B R B' R', - per_univ_elem i A B R -> - per_univ_elem i A B' R' -> - R = R'. -Proof. - intros. eauto using per_univ_elem_right_irrel_gen. -Qed. - Ltac per_univ_elem_right_irrel_rewrite1 := match goal with | H1 : {{ DF ~?A ≈ ~?B ∈ per_univ_elem ?i ↘ ?R1 }}, H2 : {{ DF ~?A ≈ ~?B' ∈ per_univ_elem ?i ↘ ?R2 }} |- _ => clean replace R2 with R1 by (eauto using per_univ_elem_right_irrel) end. -Ltac per_univ_elem_right_irrel_rewrite := repeat per_univ_elem_right_irrel_rewrite1. +Ltac per_univ_elem_right_irrel_rewrite := repeat per_univ_elem_right_irrel_rewrite1. Lemma per_univ_elem_sym : forall i A B R, per_univ_elem i A B R -> - per_univ_elem i B A R /\ (forall a b, {{ Dom a ≈ b ∈ R }} <-> {{ Dom b ≈ a ∈ R }}). -Proof. + per_univ_elem i B A R /\ (forall a b, {{ Dom a ≈ b ∈ R }} -> {{ Dom b ≈ a ∈ R }}). +Proof with solve [try econstructor; mauto]. induction 1 using per_univ_elem_ind; subst. - split. - + apply per_univ_elem_core_univ'; trivial. - + intros. split; intros []. - * specialize (H1 _ _ _ H0). - firstorder. - * specialize (H1 _ _ _ H0). - firstorder. - - split. - + econstructor. - + intros; split; mauto. + + apply per_univ_elem_core_univ'... + + intros * []. + specialize (H1 _ _ _ H0) as []... + - split... - destruct IHper_univ_elem as [? ?]. setoid_rewrite H2. split. + per_univ_elem_econstructor; eauto. intros. - assert (equiv_c'_c : in_rel c' c) by firstorder. - assert (equiv_c_c : in_rel c c) by (etransitivity; eassumption). + assert (in_rel c' c) by firstorder. + assert (in_rel c c) by (etransitivity; eassumption). destruct_rel_mod_eval. - econstructor; eauto. functional_eval_rewrite_clear. - per_univ_elem_right_irrel_rewrite. - assumption. - + assert (forall a b : domain, + per_univ_elem_right_irrel_rewrite... + + enough (forall a b : domain, (forall (c c' : domain) (equiv_c_c' : in_rel c c'), rel_mod_app (out_rel c c' equiv_c_c') a c b c') -> - (forall (c c' : domain) (equiv_c_c' : in_rel c c'), rel_mod_app (out_rel c c' equiv_c_c') b c a c')). - { - intros. - assert (equiv_c'_c : in_rel c' c) by firstorder. - assert (equiv_c_c : in_rel c c) by (etransitivity; eassumption). - destruct_rel_mod_eval. - destruct_rel_mod_app. - functional_eval_rewrite_clear. - per_univ_elem_right_irrel_rewrite. - econstructor; eauto. - firstorder. - } - firstorder. - - split. - + econstructor; mauto. - + intros; split; mauto. + (forall (c c' : domain) (equiv_c_c' : in_rel c c'), rel_mod_app (out_rel c c' equiv_c_c') b c a c')) by firstorder. + intros. + assert (in_rel c' c) by firstorder. + assert (in_rel c c) by (etransitivity; eassumption). + destruct_rel_mod_eval. + destruct_rel_mod_app. + functional_eval_rewrite_clear. + per_univ_elem_right_irrel_rewrite... + - split... +Qed. + +Corollary per_univ_sym : forall i A B R, + per_univ_elem i A B R -> + per_univ_elem i B A R. +Proof. + intros * HS. + apply per_univ_elem_sym; assumption. +Qed. + +Corollary per_elem_sym : forall i A B a b R, + per_univ_elem i A B R -> + R a b -> R b a. +Proof. + intros * HS. + eapply per_univ_elem_sym; eassumption. Qed. -Lemma per_univ_elem_left_irrel : forall i A B R A' R', +Lemma per_univ_elem_left_irrel : forall i i' A B R A' R', per_univ_elem i A B R -> - per_univ_elem i A' B R' -> + per_univ_elem i' A' B R' -> R = R'. Proof. intros * []%per_univ_elem_sym []%per_univ_elem_sym. eauto using per_univ_elem_right_irrel. Qed. -Lemma per_univ_elem_cross_irrel : forall i A B R B' R', +Lemma per_univ_elem_cross_irrel : forall i i' A B R B' R', per_univ_elem i A B R -> - per_univ_elem i B' A R' -> + per_univ_elem i' B' A R' -> R = R'. Proof. intros * ? []%per_univ_elem_sym. @@ -294,13 +252,13 @@ Qed. Ltac do_per_univ_elem_irrel_rewrite1 := match goal with | H1 : {{ DF ~?A ≈ ~_ ∈ per_univ_elem ?i ↘ ?R1 }}, - H2 : {{ DF ~?A ≈ ~_ ∈ per_univ_elem ?i ↘ ?R2 }} |- _ => + H2 : {{ DF ~?A ≈ ~_ ∈ per_univ_elem ?i' ↘ ?R2 }} |- _ => clean replace R2 with R1 by (eauto using per_univ_elem_right_irrel) | H1 : {{ DF ~_ ≈ ~?B ∈ per_univ_elem ?i ↘ ?R1 }}, - H2 : {{ DF ~_ ≈ ~?B ∈ per_univ_elem ?i ↘ ?R2 }} |- _ => + H2 : {{ DF ~_ ≈ ~?B ∈ per_univ_elem ?i' ↘ ?R2 }} |- _ => clean replace R2 with R1 by (eauto using per_univ_elem_left_irrel) | H1 : {{ DF ~?A ≈ ~_ ∈ per_univ_elem ?i ↘ ?R1 }}, - H2 : {{ DF ~_ ≈ ~?A ∈ per_univ_elem ?i ↘ ?R2 }} |- _ => + H2 : {{ DF ~_ ≈ ~?A ∈ per_univ_elem ?i' ↘ ?R2 }} |- _ => (* Order matters less here as H1 and H2 cannot be exchanged *) clean replace R2 with R1 by (symmetry; eauto using per_univ_elem_cross_irrel) end. @@ -315,42 +273,59 @@ Ltac per_univ_elem_irrel_rewrite := Lemma per_univ_elem_trans : forall i A1 A2 R, per_univ_elem i A1 A2 R -> - forall A3, - per_univ_elem i A2 A3 R -> - per_univ_elem i A1 A3 R /\ (forall a1 a2 a3, R a1 a2 -> R a2 a3 -> R a1 a3). -Proof with solve [mauto]. - induction 1 using per_univ_elem_ind; intros * HT2; - invert_per_univ_elem HT2; clear HT2. - - split; mauto. - intros * [] []. + (forall j A3, + per_univ_elem j A2 A3 R -> + per_univ_elem i A1 A3 R) /\ + (forall a1 a2 a3, R a1 a2 -> R a2 a3 -> R a1 a3). +Proof with solve [mauto| firstorder (econstructor; eauto)]. + induction 1 using per_univ_elem_ind; + [> split; + [ intros * HT2; invert_per_univ_elem HT2; clear HT2 + | intros * HTR1 HTR2 ] ..]; mauto. + - destruct HTR1, HTR2. per_univ_elem_irrel_rewrite. - destruct (H1 _ _ _ H0 _ H2) as [? ?]. - econstructor... - - split; try econstructor... + specialize (H1 _ _ _ H2) as []. + specialize (H0 _ _ H3)... + - idtac... - per_univ_elem_irrel_rewrite. rename in_rel0 into in_rel. - specialize (IHper_univ_elem _ equiv_a_a') as [? _]. - split. - + per_univ_elem_econstructor; mauto. - intros. - assert (equiv_c_c : in_rel c c) by (etransitivity; [ | symmetry]; eassumption). - destruct_rel_mod_eval. - per_univ_elem_irrel_rewrite. - firstorder (econstructor; eauto). - + setoid_rewrite H2. - intros. - assert (equiv_c'_c' : in_rel c' c') by (etransitivity; [symmetry | ]; eassumption). - destruct_rel_mod_eval. - destruct_rel_mod_app. - per_univ_elem_irrel_rewrite. - firstorder (econstructor; eauto). - - split; try per_univ_elem_econstructor... + destruct IHper_univ_elem as [? _]. + per_univ_elem_econstructor; mauto. + intros. + assert (in_rel c c) by (etransitivity; [ | symmetry]; eassumption). + destruct_rel_mod_eval. + per_univ_elem_irrel_rewrite... + - destruct IHper_univ_elem as [? ?]. + rewrite H2 in *. + intros. + assert (in_rel c' c') by (etransitivity; [symmetry | ]; eassumption). + destruct_rel_mod_eval. + destruct_rel_mod_app. + per_univ_elem_irrel_rewrite... + - try per_univ_elem_econstructor... +Qed. + +Corollary per_univ_trans : forall i j A1 A2 A3 R, + per_univ_elem i A1 A2 R -> + per_univ_elem j A2 A3 R -> + per_univ_elem i A1 A3 R. +Proof. + intros * HT1. + apply per_univ_elem_trans; assumption. +Qed. + +Corollary per_elem_trans : forall i A1 A2 a1 a2 a3 R, + per_univ_elem i A1 A2 R -> + R a1 a2 -> R a2 a3 -> R a1 a3. +Proof. + intros * HT1. + eapply per_univ_elem_trans; eassumption. Qed. Lemma per_univ_elem_cumu : forall {i a0 a1 R}, {{ DF a0 ≈ a1 ∈ per_univ_elem i ↘ R }} -> {{ DF a0 ≈ a1 ∈ per_univ_elem (S i) ↘ R }}. -Proof. +Proof with solve [mauto]. simpl. induction 1 using per_univ_elem_ind; subst. - eapply per_univ_elem_core_univ'. @@ -359,6 +334,121 @@ Proof. - per_univ_elem_econstructor; mauto. intros. destruct_rel_mod_eval. - econstructor; mauto. - - per_univ_elem_econstructor; mauto. + econstructor... + - per_univ_elem_econstructor... +Qed. + +Lemma per_ctx_env_right_irrel : forall Γ Δ Δ' R R', + {{ DF Γ ≈ Δ ∈ per_ctx_env ↘ R }} -> + {{ DF Γ ≈ Δ' ∈ per_ctx_env ↘ R' }} -> + R = R'. +Proof. + intros * Horig; gen Δ' R'. + induction Horig; intros * Hright; + try solve [inversion Hright; congruence]. + subst. + inversion_clear Hright. + specialize (IHHorig _ _ equiv_Γ_Γ'0). + subst. + enough (head_rel = head_rel0) by (subst; easy). + extensionality p. + extensionality p'. + extensionality equiv_p_p'. + unfold rel_typ in *. + destruct_rel_mod_eval. + per_univ_elem_irrel_rewrite. + congruence. +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 }}). +Proof with solve [eauto]. + simpl. + induction 1; firstorder; try solve [econstructor; eauto]; unfold rel_typ in *. + - econstructor; eauto; firstorder. + assert (tail_rel p' p) by eauto. + assert (tail_rel p p) by (etransitivity; eauto). + destruct_rel_mod_eval. + per_univ_elem_irrel_rewrite. + econstructor; eauto. + apply per_univ_sym... + - subst. + firstorder. + assert (tail_rel d{{{ p ↯ }}} d{{{ o ↯ }}}) by (unfold Symmetric in *; eauto). + assert (tail_rel d{{{ p ↯ }}} d{{{ p ↯ }}}) by (etransitivity; eauto). + destruct_rel_mod_eval. + eexists. + per_univ_elem_irrel_rewrite. + eapply per_elem_sym... +Qed. + +Lemma per_ctx_env_left_irrel : forall Γ Γ' Δ R R', + {{ DF Γ ≈ Δ ∈ per_ctx_env ↘ R }} -> + {{ DF Γ' ≈ Δ ∈ per_ctx_env ↘ R' }} -> + R = R'. +Proof. + intros * []%per_ctx_env_sym []%per_ctx_env_sym. + eauto using per_ctx_env_right_irrel. +Qed. + +Lemma per_ctx_env_cross_irrel : forall Γ Δ Δ' R R', + {{ DF Γ ≈ Δ ∈ per_ctx_env ↘ R }} -> + {{ DF Δ' ≈ Γ ∈ per_ctx_env ↘ R' }} -> + R = R'. +Proof. + intros * ? []%per_ctx_env_sym. + eauto using per_ctx_env_right_irrel. +Qed. + +Ltac do_per_ctx_env_irrel_rewrite1 := + match goal with + | H1 : {{ DF ~?Γ ≈ ~_ ∈ per_ctx_env ↘ ?R1 }}, + H2 : {{ DF ~?Γ ≈ ~_ ∈ per_ctx_env ↘ ?R2 }} |- _ => + clean replace R2 with R1 by (eauto using per_ctx_env_right_irrel) + | H1 : {{ DF ~_ ≈ ~?Δ ∈ per_ctx_env ↘ ?R1 }}, + H2 : {{ DF ~_ ≈ ~?Δ ∈ per_ctx_env ↘ ?R2 }} |- _ => + clean replace R2 with R1 by (eauto using per_ctx_env_left_irrel) + | H1 : {{ DF ~?Γ ≈ ~_ ∈ per_ctx_env ↘ ?R1 }}, + H2 : {{ DF ~_ ≈ ~?Γ ∈ per_ctx_env ↘ ?R2 }} |- _ => + (* Order matters less here as H1 and H2 cannot be exchanged *) + clean replace R2 with R1 by (symmetry; eauto using per_ctx_env_cross_irrel) + end. + +Ltac do_per_ctx_env_irrel_rewrite := + repeat do_per_ctx_env_irrel_rewrite1. + +Ltac per_ctx_env_irrel_rewrite := + functional_eval_rewrite_clear; + do_per_ctx_env_irrel_rewrite; + clear_dups. + +Lemma per_ctx_env_trans : forall Γ1 Γ2 R, + {{ DF Γ1 ≈ Γ2 ∈ per_ctx_env ↘ 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 }}). +Proof with solve [eauto]. + simpl. + induction 1; subst; + [> split; [intros * HT2; inversion HT2; subst; eauto; clear HT2 | intros * HRT1 HRT2; eauto] ..]; + per_ctx_env_irrel_rewrite. + - rename tail_rel0 into tail_rel. + econstructor; eauto. + + eapply IHper_ctx_env... + + unfold rel_typ in *. + intros. + assert (tail_rel p p) by (etransitivity; [| symmetry]; eassumption). + destruct_rel_mod_eval. + per_univ_elem_irrel_rewrite. + econstructor; eauto. + eapply per_univ_trans; [| eassumption]... + - destruct HRT1, HRT2, IHper_ctx_env. + assert (tail_rel d{{{ p1 ↯ }}} d{{{ p3 ↯ }}}) by mauto. + eexists. + unfold rel_typ in *. + destruct_rel_mod_eval. + per_univ_elem_irrel_rewrite. + eapply per_elem_trans... Qed. diff --git a/theories/Core/Semantic/Realizability.v b/theories/Core/Semantic/Realizability.v index 02f840e6..528907f1 100644 --- a/theories/Core/Semantic/Realizability.v +++ b/theories/Core/Semantic/Realizability.v @@ -1,6 +1,6 @@ From Coq Require Import Lia PeanoNat Relation_Definitions. From Equations Require Import Equations. -From Mcltt Require Import Base Evaluation LibTactics. +From Mcltt Require Import Base Evaluation LibTactics Readback. From Mcltt Require Export PER. Import Domain_Notations.