diff --git a/Makefile b/Makefile index c56e771..ef5b76f 100644 --- a/Makefile +++ b/Makefile @@ -6,7 +6,7 @@ theories: $(COQMAKEFILE) FORCE validate: $(COQMAKEFILE) FORCE $(MAKE) -f $(COQMAKEFILE) validate -$(COQMAKEFILE) : theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/unscoped.v +$(COQMAKEFILE) : theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/fintype.v $(COQBIN)coq_makefile -f _CoqProject -o $(COQMAKEFILE) install: $(COQMAKEFILE) @@ -15,13 +15,13 @@ install: $(COQMAKEFILE) uninstall: $(COQMAKEFILE) $(MAKE) -f $(COQMAKEFILE) uninstall -theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/unscoped.v : syntax.sig - autosubst -f -v ge813 -s ucoq -o theories/Autosubst2/syntax.v syntax.sig +theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/fintype.v : syntax.sig + autosubst -f -v ge813 -s coq -o theories/Autosubst2/syntax.v syntax.sig .PHONY: clean FORCE clean: test ! -f $(COQMAKEFILE) || ( $(MAKE) -f $(COQMAKEFILE) clean && rm $(COQMAKEFILE) ) - rm -f theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/fintype.v theories/Autosubst2/unscoped.v + rm -f theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/fintype.v FORCE: diff --git a/README.md b/README.md deleted file mode 100644 index fc9458e..0000000 --- a/README.md +++ /dev/null @@ -1 +0,0 @@ -This repository contains a proof of pi injectivity for an untyped equational theory with surjective pairing and extensionality. diff --git a/syntax.sig b/syntax.sig index e63c3b4..9c83c4f 100644 --- a/syntax.sig +++ b/syntax.sig @@ -1,29 +1,8 @@ -nat : Type -PTm(VarPTm) : Type Tm(VarTm) : Type -PTag : Type -TTag : Type -bool : Type - -PL : PTag -PR : PTag -TPi : TTag -TSig : TTag - -PAbs : (bind PTm in PTm) -> PTm -PApp : PTm -> PTm -> PTm -PPair : PTm -> PTm -> PTm -PProj : PTag -> PTm -> PTm -PConst : TTag -> PTm -PUniv : nat -> PTm -PBot : PTm +-- nat : Type Abs : (bind Tm in Tm) -> Tm App : Tm -> Tm -> Tm Pair : Tm -> Tm -> Tm -Proj : PTag -> Tm -> Tm -TBind : TTag -> Tm -> (bind Tm in Tm) -> Tm -Univ : nat -> Tm -BVal : bool -> Tm -Bool : Tm -If : Tm -> Tm -> Tm -> Tm +Proj1 : Tm -> Tm +Proj2 : Tm -> Tm diff --git a/theories/Autosubst2/fintype.v b/theories/Autosubst2/fintype.v new file mode 100644 index 0000000..99508b6 --- /dev/null +++ b/theories/Autosubst2/fintype.v @@ -0,0 +1,419 @@ +(** * Autosubst Header for Scoped Syntax + Our development utilises well-scoped de Bruijn syntax. This means that the de Bruijn indices are taken from finite types. As a consequence, any kind of substitution or environment used in conjunction with well-scoped syntax takes the form of a mapping from some finite type _I^n_. In particular, _renamings_ are mappings _I^n -> I^m_. Here we develop the theory of how these parts interact. + +Version: December 11, 2019. +*) +Require Import core. +Require Import Setoid Morphisms Relation_Definitions. + +Set Implicit Arguments. +Unset Strict Implicit. + +Definition ap {X Y} (f : X -> Y) {x y : X} (p : x = y) : f x = f y := + match p with eq_refl => eq_refl end. + +Definition apc {X Y} {f g : X -> Y} {x y : X} (p : f = g) (q : x = y) : f x = g y := + match q with eq_refl => match p with eq_refl => eq_refl end end. + +(** ** Primitives of the Sigma Calculus + We implement the finite type with _n_ elements, _I^n_, as the _n_-fold iteration of the Option Type. _I^0_ is implemented as the empty type. +*) + +Fixpoint fin (n : nat) : Type := + match n with + | 0 => False + | S m => option (fin m) + end. + +(** Renamings and Injective Renamings + _Renamings_ are mappings between finite types. +*) +Definition ren (m n : nat) : Type := fin m -> fin n. + +Definition id {X} := @Datatypes.id X. + +Definition idren {k: nat} : ren k k := @Datatypes.id (fin k). + +(** We give a special name, to the newest element in a non-empty finite type, as it usually corresponds to a freshly bound variable. *) +Definition var_zero {n : nat} : fin (S n) := None. + +Definition null {T} (i : fin 0) : T := match i with end. + +Definition shift {n : nat} : ren n (S n) := + Some. + +(** Extension of Finite Mappings + Assume we are given a mapping _f_ from _I^n_ to some type _X_, then we can _extend_ this mapping with a new value from _x : X_ to a mapping from _I^n+1_ to _X_. We denote this operation by _x . f_ and define it as follows: +*) +Definition scons {X : Type} {n : nat} (x : X) (f : fin n -> X) (m : fin (S n)) : X := + match m with + | None => x + | Some i => f i + end. + +#[ export ] +Hint Opaque scons : rewrite. + +(** ** Type Class Instances for Notation *) + +(** *** Type classes for renamings. *) + +Class Ren1 (X1 : Type) (Y Z : Type) := + ren1 : X1 -> Y -> Z. + +Class Ren2 (X1 X2 : Type) (Y Z : Type) := + ren2 : X1 -> X2 -> Y -> Z. + +Class Ren3 (X1 X2 X3 : Type) (Y Z : Type) := + ren3 : X1 -> X2 -> X3 -> Y -> Z. + +Class Ren4 (X1 X2 X3 X4 : Type) (Y Z : Type) := + ren4 : X1 -> X2 -> X3 -> X4 -> Y -> Z. + +Class Ren5 (X1 X2 X3 X4 X5 : Type) (Y Z : Type) := + ren5 : X1 -> X2 -> X3 -> X4 -> X5 -> Y -> Z. + +Module RenNotations. + Notation "s ⟨ xi1 ⟩" := (ren1 xi1 s) (at level 7, left associativity, format "s ⟨ xi1 ⟩") : subst_scope. + + Notation "s ⟨ xi1 ; xi2 ⟩" := (ren2 xi1 xi2 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ⟩") : subst_scope. + + Notation "s ⟨ xi1 ; xi2 ; xi3 ⟩" := (ren3 xi1 xi2 xi3 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ; xi3 ⟩") : subst_scope. + + Notation "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ⟩" := (ren4 xi1 xi2 xi3 xi4 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ⟩") : subst_scope. + + Notation "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ; xi5 ⟩" := (ren5 xi1 xi2 xi3 xi4 xi5 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ; xi5 ⟩") : subst_scope. + + Notation "⟨ xi ⟩" := (ren1 xi) (at level 1, left associativity, format "⟨ xi ⟩") : fscope. + + Notation "⟨ xi1 ; xi2 ⟩" := (ren2 xi1 xi2) (at level 1, left associativity, format "⟨ xi1 ; xi2 ⟩") : fscope. +End RenNotations. + +(** *** Type Classes for Substiution *) + +Class Subst1 (X1 : Type) (Y Z: Type) := + subst1 : X1 -> Y -> Z. + +Class Subst2 (X1 X2 : Type) (Y Z: Type) := + subst2 : X1 -> X2 -> Y -> Z. + +Class Subst3 (X1 X2 X3 : Type) (Y Z: Type) := + subst3 : X1 -> X2 -> X3 -> Y -> Z. + +Class Subst4 (X1 X2 X3 X4: Type) (Y Z: Type) := + subst4 : X1 -> X2 -> X3 -> X4 -> Y -> Z. + +Class Subst5 (X1 X2 X3 X4 X5 : Type) (Y Z: Type) := + subst5 : X1 -> X2 -> X3 -> X4 -> X5 -> Y -> Z. + +Module SubstNotations. + Notation "s [ sigma ]" := (subst1 sigma s) (at level 7, left associativity, format "s '/' [ sigma ]") : subst_scope. + + Notation "s [ sigma ; tau ]" := (subst2 sigma tau s) (at level 7, left associativity, format "s '/' [ sigma ; '/' tau ]") : subst_scope. +End SubstNotations. + +(** ** Type Class for Variables *) +Class Var X Y := + ids : X -> Y. + + +(** ** Proofs for substitution primitives *) + +(** Forward Function Composition + Substitutions represented as functions are ubiquitious in this development and we often have to compose them, without talking about their pointwise behaviour. + That is, we are interested in the forward compostion of functions, _f o g_, for which we introduce a convenient notation, "f >> g". The direction of the arrow serves as a reminder of the _forward_ nature of this composition, that is first apply _f_, then _g_. *) + +Arguments funcomp {X Y Z} (g)%fscope (f)%fscope. + +Module CombineNotations. + Notation "f >> g" := (funcomp g f) (at level 50) : fscope. + + Notation "s .: sigma" := (scons s sigma) (at level 55, sigma at next level, right associativity) : subst_scope. + + #[ global ] + Open Scope fscope. + #[ global ] + Open Scope subst_scope. +End CombineNotations. + +Import CombineNotations. + + +(** Generic lifting operation for renamings *) +Definition up_ren {m n} (xi : ren m n) : ren (S m) (S n) := + var_zero .: xi >> shift. + +(** Generic proof that lifting of renamings composes. *) +Lemma up_ren_ren k l m (xi: ren k l) (zeta : ren l m) (rho: ren k m) (E: forall x, (xi >> zeta) x = rho x) : + forall x, (up_ren xi >> up_ren zeta) x = up_ren rho x. +Proof. + intros [x|]. + - cbn. unfold funcomp. now rewrite <- E. + - reflexivity. +Qed. + +Arguments up_ren_ren {k l m} xi zeta rho E. + +Lemma fin_eta {X} (f g : fin 0 -> X) : + pointwise_relation _ eq f g. +Proof. intros []. Qed. + +(** Eta laws *) +Lemma scons_eta' {T} {n : nat} (f : fin (S n) -> T) : + pointwise_relation _ eq (f var_zero .: (funcomp f shift)) f. +Proof. intros x. destruct x; reflexivity. Qed. + +Lemma scons_eta_id' {n : nat} : + pointwise_relation (fin (S n)) eq (var_zero .: shift) id. +Proof. intros x. destruct x; reflexivity. Qed. + +Lemma scons_comp' {T:Type} {U} {m} (s: T) (sigma: fin m -> T) (tau: T -> U) : + pointwise_relation _ eq (funcomp tau (s .: sigma)) ((tau s) .: (funcomp tau sigma)). +Proof. intros x. destruct x; reflexivity. Qed. + +(* Lemma scons_tail'_pointwise {X} {n} (s : X) (f : fin n -> X) : *) +(* pointwise_relation _ eq (funcomp (scons s f) shift) f. *) +(* Proof. *) +(* reflexivity. *) +(* Qed. *) + +(* Lemma scons_tail' {X} {n} (s : X) (f : fin n -> X) x : *) +(* (scons s f (shift x)) = f x. *) +(* Proof. *) +(* reflexivity. *) +(* Qed. *) + +(* Morphism for Setoid Rewriting. The only morphism that can be defined statically. *) +#[export] Instance scons_morphism {X: Type} {n:nat} : + Proper (eq ==> pointwise_relation _ eq ==> pointwise_relation _ eq) (@scons X n). +Proof. + intros t t' -> sigma tau H. + intros [x|]. + cbn. apply H. + reflexivity. +Qed. + +#[export] Instance scons_morphism2 {X: Type} {n: nat} : + Proper (eq ==> pointwise_relation _ eq ==> eq ==> eq) (@scons X n). +Proof. + intros ? t -> sigma tau H ? x ->. + destruct x as [x|]. + cbn. apply H. + reflexivity. +Qed. + +(** ** Variadic Substitution Primitives *) + +Fixpoint shift_p (p : nat) {n} : ren n (p + n) := + fun n => match p with + | 0 => n + | S p => Some (shift_p p n) + end. + +Fixpoint scons_p {X: Type} {m : nat} : forall {n} (f : fin m -> X) (g : fin n -> X), fin (m + n) -> X. +Proof. + destruct m. + - intros n f g. exact g. + - intros n f g. cbn. apply scons. + + exact (f var_zero). + + apply scons_p. + * intros z. exact (f (Some z)). + * exact g. +Defined. + +#[export] Hint Opaque scons_p : rewrite. + +#[export] Instance scons_p_morphism {X: Type} {m n:nat} : + Proper (pointwise_relation _ eq ==> pointwise_relation _ eq ==> pointwise_relation _ eq) (@scons_p X m n). +Proof. + intros sigma sigma' Hsigma tau tau' Htau. + intros x. + induction m. + - cbn. apply Htau. + - cbn. change (fin (S m + n)) with (fin (S (m + n))) in x. + destruct x as [x|]. + + cbn. apply IHm. + intros ?. apply Hsigma. + + cbn. apply Hsigma. +Qed. + +#[export] Instance scons_p_morphism2 {X: Type} {m n:nat} : + Proper (pointwise_relation _ eq ==> pointwise_relation _ eq ==> eq ==> eq) (@scons_p X m n). +Proof. + intros sigma sigma' Hsigma tau tau' Htau ? x ->. + induction m. + - cbn. apply Htau. + - cbn. change (fin (S m + n)) with (fin (S (m + n))) in x. + destruct x as [x|]. + + cbn. apply IHm. + intros ?. apply Hsigma. + + cbn. apply Hsigma. +Qed. + +Definition zero_p {m : nat} {n} : fin m -> fin (m + n). +Proof. + induction m. + - intros []. + - intros [x|]. + + exact (shift_p 1 (IHm x)). + + exact var_zero. +Defined. + +Lemma scons_p_head' {X} {m n} (f : fin m -> X) (g : fin n -> X) z: + (scons_p f g) (zero_p z) = f z. +Proof. + induction m. + - inversion z. + - destruct z. + + simpl. simpl. now rewrite IHm. + + reflexivity. +Qed. + +Lemma scons_p_head_pointwise {X} {m n} (f : fin m -> X) (g : fin n -> X) : + pointwise_relation _ eq (funcomp (scons_p f g) zero_p) f. +Proof. + intros z. + unfold funcomp. + induction m. + - inversion z. + - destruct z. + + simpl. now rewrite IHm. + + reflexivity. +Qed. + +Lemma scons_p_tail' X m n (f : fin m -> X) (g : fin n -> X) z : + scons_p f g (shift_p m z) = g z. +Proof. induction m; cbn; eauto. Qed. + +Lemma scons_p_tail_pointwise X m n (f : fin m -> X) (g : fin n -> X) : + pointwise_relation _ eq (funcomp (scons_p f g) (shift_p m)) g. +Proof. intros z. induction m; cbn; eauto. Qed. + +Lemma destruct_fin {m n} (x : fin (m + n)): + (exists x', x = zero_p x') \/ exists x', x = shift_p m x'. +Proof. + induction m; simpl in *. + - right. eauto. + - destruct x as [x|]. + + destruct (IHm x) as [[x' ->] |[x' ->]]. + * left. now exists (Some x'). + * right. eauto. + + left. exists None. eauto. +Qed. + +Lemma scons_p_comp' X Y m n (f : fin m -> X) (g : fin n -> X) (h : X -> Y) : + pointwise_relation _ eq (funcomp h (scons_p f g)) (scons_p (f >> h) (g >> h)). +Proof. + intros x. + destruct (destruct_fin x) as [[x' ->]|[x' ->]]. + (* TODO better way to solve this? *) + - revert x'. + apply pointwise_forall. + change (fun x => (scons_p f g >> h) (zero_p x)) with (zero_p >> scons_p f g >> h). + now setoid_rewrite scons_p_head_pointwise. + - revert x'. + apply pointwise_forall. + change (fun x => (scons_p f g >> h) (shift_p m x)) with (shift_p m >> scons_p f g >> h). + change (fun x => scons_p (f >> h) (g >> h) (shift_p m x)) with (shift_p m >> scons_p (f >> h) (g >> h)). + now rewrite !scons_p_tail_pointwise. +Qed. + + +Lemma scons_p_congr {X} {m n} (f f' : fin m -> X) (g g': fin n -> X) z: + (forall x, f x = f' x) -> (forall x, g x = g' x) -> scons_p f g z = scons_p f' g' z. +Proof. intros H1 H2. induction m; eauto. cbn. destruct z; eauto. Qed. + +(** Generic n-ary lifting operation. *) +Definition upRen_p p { m : nat } { n : nat } (xi : (fin) (m) -> (fin) (n)) : fin (p + m) -> fin (p + n) := + scons_p (zero_p ) (xi >> shift_p _). + +Arguments upRen_p p {m n} xi. + +(** Generic proof for composition of n-ary lifting. *) +Lemma up_ren_ren_p p k l m (xi: ren k l) (zeta : ren l m) (rho: ren k m) (E: forall x, (xi >> zeta) x = rho x) : + forall x, (upRen_p p xi >> upRen_p p zeta) x = upRen_p p rho x. +Proof. + intros x. destruct (destruct_fin x) as [[? ->]|[? ->]]. + - unfold upRen_p. unfold funcomp. now repeat rewrite scons_p_head'. + - unfold upRen_p. unfold funcomp. repeat rewrite scons_p_tail'. + now rewrite <- E. +Qed. + + +Arguments zero_p m {n}. +Arguments scons_p {X} m {n} f g. + +Lemma scons_p_eta {X} {m n} {f : fin m -> X} + {g : fin n -> X} (h: fin (m + n) -> X) {z: fin (m + n)}: + (forall x, g x = h (shift_p m x)) -> (forall x, f x = h (zero_p m x)) -> scons_p m f g z = h z. +Proof. + intros H1 H2. destruct (destruct_fin z) as [[? ->] |[? ->]]. + - rewrite scons_p_head'. eauto. + - rewrite scons_p_tail'. eauto. +Qed. + +Arguments scons_p_eta {X} {m n} {f g} h {z}. +Arguments scons_p_congr {X} {m n} {f f'} {g g'} {z}. + +(** ** Notations for Scoped Syntax *) + +Module ScopedNotations. + Include RenNotations. + Include SubstNotations. + Include CombineNotations. + +(* Notation "s , sigma" := (scons s sigma) (at level 60, format "s , sigma", right associativity) : subst_scope. *) + + Notation "s '..'" := (scons s ids) (at level 1, format "s ..") : subst_scope. + + Notation "↑" := (shift) : subst_scope. + + #[global] + Open Scope fscope. + #[global] + Open Scope subst_scope. +End ScopedNotations. + + +(** ** Tactics for Scoped Syntax *) + +Tactic Notation "auto_case" tactic(t) := (match goal with + | [|- forall (i : fin 0), _] => intros []; t + | [|- forall (i : fin (S (S (S (S _))))), _] => intros [[[[|]|]|]|]; t + | [|- forall (i : fin (S (S (S _)))), _] => intros [[[|]|]|]; t + | [|- forall (i : fin (S (S _))), _] => intros [[?|]|]; t + | [|- forall (i : fin (S _)), _] => intros [?|]; t + end). + +#[export] Hint Rewrite @scons_p_head' @scons_p_tail' : FunctorInstances. + +(** Generic fsimpl tactic: simplifies the above primitives in a goal. *) +Ltac fsimpl := + repeat match goal with + | [|- context[id >> ?f]] => change (id >> f) with f (* AsimplCompIdL *) + | [|- context[?f >> id]] => change (f >> id) with f (* AsimplCompIdR *) + | [|- context [id ?s]] => change (id s) with s + | [|- context[(?f >> ?g) >> ?h]] => change ((f >> g) >> h) with (f >> (g >> h)) (* AsimplComp *) + (* | [|- zero_p >> scons_p ?f ?g] => rewrite scons_p_head *) + | [|- context[(?s .: ?sigma) var_zero]] => change ((s.:sigma) var_zero) with s + | [|- context[(?s .: ?sigma) (shift ?m)]] => change ((s.:sigma) (shift m)) with (sigma m) + (* first [progress setoid_rewrite scons_tail' | progress setoid_rewrite scons_tail'_pointwise ] *) + | [|- context[idren >> ?f]] => change (idren >> f) with f + | [|- context[?f >> idren]] => change (f >> idren) with f + | [|- context[?f >> (?x .: ?g)]] => change (f >> (x .: g)) with g (* f should evaluate to shift *) + | [|- context[?x2 .: (funcomp ?f shift)]] => change (scons x2 (funcomp f shift)) with (scons (f var_zero) (funcomp f shift)); setoid_rewrite (@scons_eta' _ _ f); eta_reduce + | [|- context[?f var_zero .: ?g]] => change (scons (f var_zero) g) with (scons (f var_zero) (funcomp f shift)); setoid_rewrite scons_eta'; eta_reduce + | [|- _ = ?h (?f ?s)] => change (h (f s)) with ((f >> h) s) + | [|- ?h (?f ?s) = _] => change (h (f s)) with ((f >> h) s) + | [|- context[funcomp _ (scons _ _)]] => setoid_rewrite scons_comp'; eta_reduce + | [|- context[funcomp _ (scons_p _ _ _)]] => setoid_rewrite scons_p_comp'; eta_reduce + | [|- context[scons (@var_zero _) shift]] => setoid_rewrite scons_eta_id'; eta_reduce + (* | _ => progress autorewrite with FunctorInstances *) + | [|- context[funcomp (scons_p _ _ _) (zero_p _)]] => + first [progress setoid_rewrite scons_p_head_pointwise | progress setoid_rewrite scons_p_head' ] + | [|- context[scons_p _ _ _ (zero_p _ _)]] => setoid_rewrite scons_p_head' + | [|- context[funcomp (scons_p _ _ _) (shift_p _)]] => + first [progress setoid_rewrite scons_p_tail_pointwise | progress setoid_rewrite scons_p_tail' ] + | [|- context[scons_p _ _ _ (shift_p _ _)]] => setoid_rewrite scons_p_tail' + | _ => first [progress minimize | progress cbn [shift scons scons_p] ] + end. diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v index 659d8b0..79e4f87 100644 --- a/theories/Autosubst2/syntax.v +++ b/theories/Autosubst2/syntax.v @@ -1,838 +1,305 @@ -Require Import core unscoped. +Require Import core fintype. Require Import Setoid Morphisms Relation_Definitions. Module Core. -Inductive PTag : Type := - | PL : PTag - | PR : PTag. +Inductive Tm (n_Tm : nat) : Type := + | VarTm : fin n_Tm -> Tm n_Tm + | Abs : Tm (S n_Tm) -> Tm n_Tm + | App : Tm n_Tm -> Tm n_Tm -> Tm n_Tm + | Pair : Tm n_Tm -> Tm n_Tm -> Tm n_Tm + | Proj1 : Tm n_Tm -> Tm n_Tm + | Proj2 : Tm n_Tm -> Tm n_Tm. -Lemma congr_PL : PL = PL. +Lemma congr_Abs {m_Tm : nat} {s0 : Tm (S m_Tm)} {t0 : Tm (S m_Tm)} + (H0 : s0 = t0) : Abs m_Tm s0 = Abs m_Tm t0. Proof. -exact (eq_refl). +exact (eq_trans eq_refl (ap (fun x => Abs m_Tm x) H0)). Qed. -Lemma congr_PR : PR = PR. +Lemma congr_App {m_Tm : nat} {s0 : Tm m_Tm} {s1 : Tm m_Tm} {t0 : Tm m_Tm} + {t1 : Tm m_Tm} (H0 : s0 = t0) (H1 : s1 = t1) : + App m_Tm s0 s1 = App m_Tm t0 t1. Proof. -exact (eq_refl). +exact (eq_trans (eq_trans eq_refl (ap (fun x => App m_Tm x s1) H0)) + (ap (fun x => App m_Tm t0 x) H1)). Qed. -Inductive TTag : Type := - | TPi : TTag - | TSig : TTag. - -Lemma congr_TPi : TPi = TPi. +Lemma congr_Pair {m_Tm : nat} {s0 : Tm m_Tm} {s1 : Tm m_Tm} {t0 : Tm m_Tm} + {t1 : Tm m_Tm} (H0 : s0 = t0) (H1 : s1 = t1) : + Pair m_Tm s0 s1 = Pair m_Tm t0 t1. Proof. -exact (eq_refl). +exact (eq_trans (eq_trans eq_refl (ap (fun x => Pair m_Tm x s1) H0)) + (ap (fun x => Pair m_Tm t0 x) H1)). Qed. -Lemma congr_TSig : TSig = TSig. +Lemma congr_Proj1 {m_Tm : nat} {s0 : Tm m_Tm} {t0 : Tm m_Tm} (H0 : s0 = t0) : + Proj1 m_Tm s0 = Proj1 m_Tm t0. Proof. -exact (eq_refl). +exact (eq_trans eq_refl (ap (fun x => Proj1 m_Tm x) H0)). Qed. -Inductive PTm : Type := - | VarPTm : nat -> PTm - | PAbs : PTm -> PTm - | PApp : PTm -> PTm -> PTm - | PPair : PTm -> PTm -> PTm - | PProj : PTag -> PTm -> PTm - | PConst : TTag -> PTm - | PUniv : nat -> PTm - | PBot : PTm. - -Lemma congr_PAbs {s0 : PTm} {t0 : PTm} (H0 : s0 = t0) : PAbs s0 = PAbs t0. +Lemma congr_Proj2 {m_Tm : nat} {s0 : Tm m_Tm} {t0 : Tm m_Tm} (H0 : s0 = t0) : + Proj2 m_Tm s0 = Proj2 m_Tm t0. Proof. -exact (eq_trans eq_refl (ap (fun x => PAbs x) H0)). +exact (eq_trans eq_refl (ap (fun x => Proj2 m_Tm x) H0)). Qed. -Lemma congr_PApp {s0 : PTm} {s1 : PTm} {t0 : PTm} {t1 : PTm} (H0 : s0 = t0) - (H1 : s1 = t1) : PApp s0 s1 = PApp t0 t1. -Proof. -exact (eq_trans (eq_trans eq_refl (ap (fun x => PApp x s1) H0)) - (ap (fun x => PApp t0 x) H1)). -Qed. - -Lemma congr_PPair {s0 : PTm} {s1 : PTm} {t0 : PTm} {t1 : PTm} (H0 : s0 = t0) - (H1 : s1 = t1) : PPair s0 s1 = PPair t0 t1. -Proof. -exact (eq_trans (eq_trans eq_refl (ap (fun x => PPair x s1) H0)) - (ap (fun x => PPair t0 x) H1)). -Qed. - -Lemma congr_PProj {s0 : PTag} {s1 : PTm} {t0 : PTag} {t1 : PTm} - (H0 : s0 = t0) (H1 : s1 = t1) : PProj s0 s1 = PProj t0 t1. -Proof. -exact (eq_trans (eq_trans eq_refl (ap (fun x => PProj x s1) H0)) - (ap (fun x => PProj t0 x) H1)). -Qed. - -Lemma congr_PConst {s0 : TTag} {t0 : TTag} (H0 : s0 = t0) : - PConst s0 = PConst t0. -Proof. -exact (eq_trans eq_refl (ap (fun x => PConst x) H0)). -Qed. - -Lemma congr_PUniv {s0 : nat} {t0 : nat} (H0 : s0 = t0) : PUniv s0 = PUniv t0. -Proof. -exact (eq_trans eq_refl (ap (fun x => PUniv x) H0)). -Qed. - -Lemma congr_PBot : PBot = PBot. -Proof. -exact (eq_refl). -Qed. - -Lemma upRen_PTm_PTm (xi : nat -> nat) : nat -> nat. +Lemma upRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n) : + fin (S m) -> fin (S n). Proof. exact (up_ren xi). Defined. -Fixpoint ren_PTm (xi_PTm : nat -> nat) (s : PTm) {struct s} : PTm := - match s with - | VarPTm s0 => VarPTm (xi_PTm s0) - | PAbs s0 => PAbs (ren_PTm (upRen_PTm_PTm xi_PTm) s0) - | PApp s0 s1 => PApp (ren_PTm xi_PTm s0) (ren_PTm xi_PTm s1) - | PPair s0 s1 => PPair (ren_PTm xi_PTm s0) (ren_PTm xi_PTm s1) - | PProj s0 s1 => PProj s0 (ren_PTm xi_PTm s1) - | PConst s0 => PConst s0 - | PUniv s0 => PUniv s0 - | PBot => PBot - end. - -Lemma up_PTm_PTm (sigma : nat -> PTm) : nat -> PTm. +Lemma upRen_list_Tm_Tm (p : nat) {m : nat} {n : nat} (xi : fin m -> fin n) : + fin (plus p m) -> fin (plus p n). Proof. -exact (scons (VarPTm var_zero) (funcomp (ren_PTm shift) sigma)). +exact (upRen_p p xi). Defined. -Fixpoint subst_PTm (sigma_PTm : nat -> PTm) (s : PTm) {struct s} : PTm := +Fixpoint ren_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm) +(s : Tm m_Tm) {struct s} : Tm n_Tm := match s with - | VarPTm s0 => sigma_PTm s0 - | PAbs s0 => PAbs (subst_PTm (up_PTm_PTm sigma_PTm) s0) - | PApp s0 s1 => PApp (subst_PTm sigma_PTm s0) (subst_PTm sigma_PTm s1) - | PPair s0 s1 => PPair (subst_PTm sigma_PTm s0) (subst_PTm sigma_PTm s1) - | PProj s0 s1 => PProj s0 (subst_PTm sigma_PTm s1) - | PConst s0 => PConst s0 - | PUniv s0 => PUniv s0 - | PBot => PBot + | VarTm _ s0 => VarTm n_Tm (xi_Tm s0) + | Abs _ s0 => Abs n_Tm (ren_Tm (upRen_Tm_Tm xi_Tm) s0) + | App _ s0 s1 => App n_Tm (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1) + | Pair _ s0 s1 => Pair n_Tm (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1) + | Proj1 _ s0 => Proj1 n_Tm (ren_Tm xi_Tm s0) + | Proj2 _ s0 => Proj2 n_Tm (ren_Tm xi_Tm s0) end. -Lemma upId_PTm_PTm (sigma : nat -> PTm) (Eq : forall x, sigma x = VarPTm x) : - forall x, up_PTm_PTm sigma x = VarPTm x. +Lemma up_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm) : + fin (S m) -> Tm (S n_Tm). Proof. -exact (fun n => - match n with - | S n' => ap (ren_PTm shift) (Eq n') - | O => eq_refl - end). -Qed. - -Fixpoint idSubst_PTm (sigma_PTm : nat -> PTm) -(Eq_PTm : forall x, sigma_PTm x = VarPTm x) (s : PTm) {struct s} : -subst_PTm sigma_PTm s = s := - match s with - | VarPTm s0 => Eq_PTm s0 - | PAbs s0 => - congr_PAbs - (idSubst_PTm (up_PTm_PTm sigma_PTm) (upId_PTm_PTm _ Eq_PTm) s0) - | PApp s0 s1 => - congr_PApp (idSubst_PTm sigma_PTm Eq_PTm s0) - (idSubst_PTm sigma_PTm Eq_PTm s1) - | PPair s0 s1 => - congr_PPair (idSubst_PTm sigma_PTm Eq_PTm s0) - (idSubst_PTm sigma_PTm Eq_PTm s1) - | PProj s0 s1 => congr_PProj (eq_refl s0) (idSubst_PTm sigma_PTm Eq_PTm s1) - | PConst s0 => congr_PConst (eq_refl s0) - | PUniv s0 => congr_PUniv (eq_refl s0) - | PBot => congr_PBot - end. - -Lemma upExtRen_PTm_PTm (xi : nat -> nat) (zeta : nat -> nat) - (Eq : forall x, xi x = zeta x) : - forall x, upRen_PTm_PTm xi x = upRen_PTm_PTm zeta x. -Proof. -exact (fun n => match n with - | S n' => ap shift (Eq n') - | O => eq_refl - end). -Qed. - -Fixpoint extRen_PTm (xi_PTm : nat -> nat) (zeta_PTm : nat -> nat) -(Eq_PTm : forall x, xi_PTm x = zeta_PTm x) (s : PTm) {struct s} : -ren_PTm xi_PTm s = ren_PTm zeta_PTm s := - match s with - | VarPTm s0 => ap (VarPTm) (Eq_PTm s0) - | PAbs s0 => - congr_PAbs - (extRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm) - (upExtRen_PTm_PTm _ _ Eq_PTm) s0) - | PApp s0 s1 => - congr_PApp (extRen_PTm xi_PTm zeta_PTm Eq_PTm s0) - (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1) - | PPair s0 s1 => - congr_PPair (extRen_PTm xi_PTm zeta_PTm Eq_PTm s0) - (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1) - | PProj s0 s1 => - congr_PProj (eq_refl s0) (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1) - | PConst s0 => congr_PConst (eq_refl s0) - | PUniv s0 => congr_PUniv (eq_refl s0) - | PBot => congr_PBot - end. - -Lemma upExt_PTm_PTm (sigma : nat -> PTm) (tau : nat -> PTm) - (Eq : forall x, sigma x = tau x) : - forall x, up_PTm_PTm sigma x = up_PTm_PTm tau x. -Proof. -exact (fun n => - match n with - | S n' => ap (ren_PTm shift) (Eq n') - | O => eq_refl - end). -Qed. - -Fixpoint ext_PTm (sigma_PTm : nat -> PTm) (tau_PTm : nat -> PTm) -(Eq_PTm : forall x, sigma_PTm x = tau_PTm x) (s : PTm) {struct s} : -subst_PTm sigma_PTm s = subst_PTm tau_PTm s := - match s with - | VarPTm s0 => Eq_PTm s0 - | PAbs s0 => - congr_PAbs - (ext_PTm (up_PTm_PTm sigma_PTm) (up_PTm_PTm tau_PTm) - (upExt_PTm_PTm _ _ Eq_PTm) s0) - | PApp s0 s1 => - congr_PApp (ext_PTm sigma_PTm tau_PTm Eq_PTm s0) - (ext_PTm sigma_PTm tau_PTm Eq_PTm s1) - | PPair s0 s1 => - congr_PPair (ext_PTm sigma_PTm tau_PTm Eq_PTm s0) - (ext_PTm sigma_PTm tau_PTm Eq_PTm s1) - | PProj s0 s1 => - congr_PProj (eq_refl s0) (ext_PTm sigma_PTm tau_PTm Eq_PTm s1) - | PConst s0 => congr_PConst (eq_refl s0) - | PUniv s0 => congr_PUniv (eq_refl s0) - | PBot => congr_PBot - end. - -Lemma up_ren_ren_PTm_PTm (xi : nat -> nat) (zeta : nat -> nat) - (rho : nat -> nat) (Eq : forall x, funcomp zeta xi x = rho x) : - forall x, - funcomp (upRen_PTm_PTm zeta) (upRen_PTm_PTm xi) x = upRen_PTm_PTm rho x. -Proof. -exact (up_ren_ren xi zeta rho Eq). -Qed. - -Fixpoint compRenRen_PTm (xi_PTm : nat -> nat) (zeta_PTm : nat -> nat) -(rho_PTm : nat -> nat) -(Eq_PTm : forall x, funcomp zeta_PTm xi_PTm x = rho_PTm x) (s : PTm) {struct - s} : ren_PTm zeta_PTm (ren_PTm xi_PTm s) = ren_PTm rho_PTm s := - match s with - | VarPTm s0 => ap (VarPTm) (Eq_PTm s0) - | PAbs s0 => - congr_PAbs - (compRenRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm) - (upRen_PTm_PTm rho_PTm) (up_ren_ren _ _ _ Eq_PTm) s0) - | PApp s0 s1 => - congr_PApp (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s0) - (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1) - | PPair s0 s1 => - congr_PPair (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s0) - (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1) - | PProj s0 s1 => - congr_PProj (eq_refl s0) - (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1) - | PConst s0 => congr_PConst (eq_refl s0) - | PUniv s0 => congr_PUniv (eq_refl s0) - | PBot => congr_PBot - end. - -Lemma up_ren_subst_PTm_PTm (xi : nat -> nat) (tau : nat -> PTm) - (theta : nat -> PTm) (Eq : forall x, funcomp tau xi x = theta x) : - forall x, - funcomp (up_PTm_PTm tau) (upRen_PTm_PTm xi) x = up_PTm_PTm theta x. -Proof. -exact (fun n => - match n with - | S n' => ap (ren_PTm shift) (Eq n') - | O => eq_refl - end). -Qed. - -Fixpoint compRenSubst_PTm (xi_PTm : nat -> nat) (tau_PTm : nat -> PTm) -(theta_PTm : nat -> PTm) -(Eq_PTm : forall x, funcomp tau_PTm xi_PTm x = theta_PTm x) (s : PTm) {struct - s} : subst_PTm tau_PTm (ren_PTm xi_PTm s) = subst_PTm theta_PTm s := - match s with - | VarPTm s0 => Eq_PTm s0 - | PAbs s0 => - congr_PAbs - (compRenSubst_PTm (upRen_PTm_PTm xi_PTm) (up_PTm_PTm tau_PTm) - (up_PTm_PTm theta_PTm) (up_ren_subst_PTm_PTm _ _ _ Eq_PTm) s0) - | PApp s0 s1 => - congr_PApp (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s0) - (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1) - | PPair s0 s1 => - congr_PPair (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s0) - (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1) - | PProj s0 s1 => - congr_PProj (eq_refl s0) - (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1) - | PConst s0 => congr_PConst (eq_refl s0) - | PUniv s0 => congr_PUniv (eq_refl s0) - | PBot => congr_PBot - end. - -Lemma up_subst_ren_PTm_PTm (sigma : nat -> PTm) (zeta_PTm : nat -> nat) - (theta : nat -> PTm) - (Eq : forall x, funcomp (ren_PTm zeta_PTm) sigma x = theta x) : - forall x, - funcomp (ren_PTm (upRen_PTm_PTm zeta_PTm)) (up_PTm_PTm sigma) x = - up_PTm_PTm theta x. -Proof. -exact (fun n => - match n with - | S n' => - eq_trans - (compRenRen_PTm shift (upRen_PTm_PTm zeta_PTm) - (funcomp shift zeta_PTm) (fun x => eq_refl) (sigma n')) - (eq_trans - (eq_sym - (compRenRen_PTm zeta_PTm shift (funcomp shift zeta_PTm) - (fun x => eq_refl) (sigma n'))) - (ap (ren_PTm shift) (Eq n'))) - | O => eq_refl - end). -Qed. - -Fixpoint compSubstRen_PTm (sigma_PTm : nat -> PTm) (zeta_PTm : nat -> nat) -(theta_PTm : nat -> PTm) -(Eq_PTm : forall x, funcomp (ren_PTm zeta_PTm) sigma_PTm x = theta_PTm x) -(s : PTm) {struct s} : -ren_PTm zeta_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s := - match s with - | VarPTm s0 => Eq_PTm s0 - | PAbs s0 => - congr_PAbs - (compSubstRen_PTm (up_PTm_PTm sigma_PTm) (upRen_PTm_PTm zeta_PTm) - (up_PTm_PTm theta_PTm) (up_subst_ren_PTm_PTm _ _ _ Eq_PTm) s0) - | PApp s0 s1 => - congr_PApp (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s0) - (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1) - | PPair s0 s1 => - congr_PPair (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s0) - (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1) - | PProj s0 s1 => - congr_PProj (eq_refl s0) - (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1) - | PConst s0 => congr_PConst (eq_refl s0) - | PUniv s0 => congr_PUniv (eq_refl s0) - | PBot => congr_PBot - end. - -Lemma up_subst_subst_PTm_PTm (sigma : nat -> PTm) (tau_PTm : nat -> PTm) - (theta : nat -> PTm) - (Eq : forall x, funcomp (subst_PTm tau_PTm) sigma x = theta x) : - forall x, - funcomp (subst_PTm (up_PTm_PTm tau_PTm)) (up_PTm_PTm sigma) x = - up_PTm_PTm theta x. -Proof. -exact (fun n => - match n with - | S n' => - eq_trans - (compRenSubst_PTm shift (up_PTm_PTm tau_PTm) - (funcomp (up_PTm_PTm tau_PTm) shift) (fun x => eq_refl) - (sigma n')) - (eq_trans - (eq_sym - (compSubstRen_PTm tau_PTm shift - (funcomp (ren_PTm shift) tau_PTm) (fun x => eq_refl) - (sigma n'))) (ap (ren_PTm shift) (Eq n'))) - | O => eq_refl - end). -Qed. - -Fixpoint compSubstSubst_PTm (sigma_PTm : nat -> PTm) (tau_PTm : nat -> PTm) -(theta_PTm : nat -> PTm) -(Eq_PTm : forall x, funcomp (subst_PTm tau_PTm) sigma_PTm x = theta_PTm x) -(s : PTm) {struct s} : -subst_PTm tau_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s := - match s with - | VarPTm s0 => Eq_PTm s0 - | PAbs s0 => - congr_PAbs - (compSubstSubst_PTm (up_PTm_PTm sigma_PTm) (up_PTm_PTm tau_PTm) - (up_PTm_PTm theta_PTm) (up_subst_subst_PTm_PTm _ _ _ Eq_PTm) s0) - | PApp s0 s1 => - congr_PApp (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s0) - (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1) - | PPair s0 s1 => - congr_PPair (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s0) - (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1) - | PProj s0 s1 => - congr_PProj (eq_refl s0) - (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1) - | PConst s0 => congr_PConst (eq_refl s0) - | PUniv s0 => congr_PUniv (eq_refl s0) - | PBot => congr_PBot - end. - -Lemma renRen_PTm (xi_PTm : nat -> nat) (zeta_PTm : nat -> nat) (s : PTm) : - ren_PTm zeta_PTm (ren_PTm xi_PTm s) = ren_PTm (funcomp zeta_PTm xi_PTm) s. -Proof. -exact (compRenRen_PTm xi_PTm zeta_PTm _ (fun n => eq_refl) s). -Qed. - -Lemma renRen'_PTm_pointwise (xi_PTm : nat -> nat) (zeta_PTm : nat -> nat) : - pointwise_relation _ eq (funcomp (ren_PTm zeta_PTm) (ren_PTm xi_PTm)) - (ren_PTm (funcomp zeta_PTm xi_PTm)). -Proof. -exact (fun s => compRenRen_PTm xi_PTm zeta_PTm _ (fun n => eq_refl) s). -Qed. - -Lemma renSubst_PTm (xi_PTm : nat -> nat) (tau_PTm : nat -> PTm) (s : PTm) : - subst_PTm tau_PTm (ren_PTm xi_PTm s) = subst_PTm (funcomp tau_PTm xi_PTm) s. -Proof. -exact (compRenSubst_PTm xi_PTm tau_PTm _ (fun n => eq_refl) s). -Qed. - -Lemma renSubst_PTm_pointwise (xi_PTm : nat -> nat) (tau_PTm : nat -> PTm) : - pointwise_relation _ eq (funcomp (subst_PTm tau_PTm) (ren_PTm xi_PTm)) - (subst_PTm (funcomp tau_PTm xi_PTm)). -Proof. -exact (fun s => compRenSubst_PTm xi_PTm tau_PTm _ (fun n => eq_refl) s). -Qed. - -Lemma substRen_PTm (sigma_PTm : nat -> PTm) (zeta_PTm : nat -> nat) (s : PTm) - : - ren_PTm zeta_PTm (subst_PTm sigma_PTm s) = - subst_PTm (funcomp (ren_PTm zeta_PTm) sigma_PTm) s. -Proof. -exact (compSubstRen_PTm sigma_PTm zeta_PTm _ (fun n => eq_refl) s). -Qed. - -Lemma substRen_PTm_pointwise (sigma_PTm : nat -> PTm) (zeta_PTm : nat -> nat) - : - pointwise_relation _ eq (funcomp (ren_PTm zeta_PTm) (subst_PTm sigma_PTm)) - (subst_PTm (funcomp (ren_PTm zeta_PTm) sigma_PTm)). -Proof. -exact (fun s => compSubstRen_PTm sigma_PTm zeta_PTm _ (fun n => eq_refl) s). -Qed. - -Lemma substSubst_PTm (sigma_PTm : nat -> PTm) (tau_PTm : nat -> PTm) - (s : PTm) : - subst_PTm tau_PTm (subst_PTm sigma_PTm s) = - subst_PTm (funcomp (subst_PTm tau_PTm) sigma_PTm) s. -Proof. -exact (compSubstSubst_PTm sigma_PTm tau_PTm _ (fun n => eq_refl) s). -Qed. - -Lemma substSubst_PTm_pointwise (sigma_PTm : nat -> PTm) - (tau_PTm : nat -> PTm) : - pointwise_relation _ eq (funcomp (subst_PTm tau_PTm) (subst_PTm sigma_PTm)) - (subst_PTm (funcomp (subst_PTm tau_PTm) sigma_PTm)). -Proof. -exact (fun s => compSubstSubst_PTm sigma_PTm tau_PTm _ (fun n => eq_refl) s). -Qed. - -Lemma rinstInst_up_PTm_PTm (xi : nat -> nat) (sigma : nat -> PTm) - (Eq : forall x, funcomp (VarPTm) xi x = sigma x) : - forall x, funcomp (VarPTm) (upRen_PTm_PTm xi) x = up_PTm_PTm sigma x. -Proof. -exact (fun n => - match n with - | S n' => ap (ren_PTm shift) (Eq n') - | O => eq_refl - end). -Qed. - -Fixpoint rinst_inst_PTm (xi_PTm : nat -> nat) (sigma_PTm : nat -> PTm) -(Eq_PTm : forall x, funcomp (VarPTm) xi_PTm x = sigma_PTm x) (s : PTm) -{struct s} : ren_PTm xi_PTm s = subst_PTm sigma_PTm s := - match s with - | VarPTm s0 => Eq_PTm s0 - | PAbs s0 => - congr_PAbs - (rinst_inst_PTm (upRen_PTm_PTm xi_PTm) (up_PTm_PTm sigma_PTm) - (rinstInst_up_PTm_PTm _ _ Eq_PTm) s0) - | PApp s0 s1 => - congr_PApp (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s0) - (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1) - | PPair s0 s1 => - congr_PPair (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s0) - (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1) - | PProj s0 s1 => - congr_PProj (eq_refl s0) (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1) - | PConst s0 => congr_PConst (eq_refl s0) - | PUniv s0 => congr_PUniv (eq_refl s0) - | PBot => congr_PBot - end. - -Lemma rinstInst'_PTm (xi_PTm : nat -> nat) (s : PTm) : - ren_PTm xi_PTm s = subst_PTm (funcomp (VarPTm) xi_PTm) s. -Proof. -exact (rinst_inst_PTm xi_PTm _ (fun n => eq_refl) s). -Qed. - -Lemma rinstInst'_PTm_pointwise (xi_PTm : nat -> nat) : - pointwise_relation _ eq (ren_PTm xi_PTm) - (subst_PTm (funcomp (VarPTm) xi_PTm)). -Proof. -exact (fun s => rinst_inst_PTm xi_PTm _ (fun n => eq_refl) s). -Qed. - -Lemma instId'_PTm (s : PTm) : subst_PTm (VarPTm) s = s. -Proof. -exact (idSubst_PTm (VarPTm) (fun n => eq_refl) s). -Qed. - -Lemma instId'_PTm_pointwise : pointwise_relation _ eq (subst_PTm (VarPTm)) id. -Proof. -exact (fun s => idSubst_PTm (VarPTm) (fun n => eq_refl) s). -Qed. - -Lemma rinstId'_PTm (s : PTm) : ren_PTm id s = s. -Proof. -exact (eq_ind_r (fun t => t = s) (instId'_PTm s) (rinstInst'_PTm id s)). -Qed. - -Lemma rinstId'_PTm_pointwise : pointwise_relation _ eq (@ren_PTm id) id. -Proof. -exact (fun s => - eq_ind_r (fun t => t = s) (instId'_PTm s) (rinstInst'_PTm id s)). -Qed. - -Lemma varL'_PTm (sigma_PTm : nat -> PTm) (x : nat) : - subst_PTm sigma_PTm (VarPTm x) = sigma_PTm x. -Proof. -exact (eq_refl). -Qed. - -Lemma varL'_PTm_pointwise (sigma_PTm : nat -> PTm) : - pointwise_relation _ eq (funcomp (subst_PTm sigma_PTm) (VarPTm)) sigma_PTm. -Proof. -exact (fun x => eq_refl). -Qed. - -Lemma varLRen'_PTm (xi_PTm : nat -> nat) (x : nat) : - ren_PTm xi_PTm (VarPTm x) = VarPTm (xi_PTm x). -Proof. -exact (eq_refl). -Qed. - -Lemma varLRen'_PTm_pointwise (xi_PTm : nat -> nat) : - pointwise_relation _ eq (funcomp (ren_PTm xi_PTm) (VarPTm)) - (funcomp (VarPTm) xi_PTm). -Proof. -exact (fun x => eq_refl). -Qed. - -Inductive Tm : Type := - | VarTm : nat -> Tm - | Abs : Tm -> Tm - | App : Tm -> Tm -> Tm - | Pair : Tm -> Tm -> Tm - | Proj : PTag -> Tm -> Tm - | TBind : TTag -> Tm -> Tm -> Tm - | Univ : nat -> Tm - | BVal : bool -> Tm - | Bool : Tm - | If : Tm -> Tm -> Tm -> Tm. - -Lemma congr_Abs {s0 : Tm} {t0 : Tm} (H0 : s0 = t0) : Abs s0 = Abs t0. -Proof. -exact (eq_trans eq_refl (ap (fun x => Abs x) H0)). -Qed. - -Lemma congr_App {s0 : Tm} {s1 : Tm} {t0 : Tm} {t1 : Tm} (H0 : s0 = t0) - (H1 : s1 = t1) : App s0 s1 = App t0 t1. -Proof. -exact (eq_trans (eq_trans eq_refl (ap (fun x => App x s1) H0)) - (ap (fun x => App t0 x) H1)). -Qed. - -Lemma congr_Pair {s0 : Tm} {s1 : Tm} {t0 : Tm} {t1 : Tm} (H0 : s0 = t0) - (H1 : s1 = t1) : Pair s0 s1 = Pair t0 t1. -Proof. -exact (eq_trans (eq_trans eq_refl (ap (fun x => Pair x s1) H0)) - (ap (fun x => Pair t0 x) H1)). -Qed. - -Lemma congr_Proj {s0 : PTag} {s1 : Tm} {t0 : PTag} {t1 : Tm} (H0 : s0 = t0) - (H1 : s1 = t1) : Proj s0 s1 = Proj t0 t1. -Proof. -exact (eq_trans (eq_trans eq_refl (ap (fun x => Proj x s1) H0)) - (ap (fun x => Proj t0 x) H1)). -Qed. - -Lemma congr_TBind {s0 : TTag} {s1 : Tm} {s2 : Tm} {t0 : TTag} {t1 : Tm} - {t2 : Tm} (H0 : s0 = t0) (H1 : s1 = t1) (H2 : s2 = t2) : - TBind s0 s1 s2 = TBind t0 t1 t2. -Proof. -exact (eq_trans - (eq_trans (eq_trans eq_refl (ap (fun x => TBind x s1 s2) H0)) - (ap (fun x => TBind t0 x s2) H1)) - (ap (fun x => TBind t0 t1 x) H2)). -Qed. - -Lemma congr_Univ {s0 : nat} {t0 : nat} (H0 : s0 = t0) : Univ s0 = Univ t0. -Proof. -exact (eq_trans eq_refl (ap (fun x => Univ x) H0)). -Qed. - -Lemma congr_BVal {s0 : bool} {t0 : bool} (H0 : s0 = t0) : BVal s0 = BVal t0. -Proof. -exact (eq_trans eq_refl (ap (fun x => BVal x) H0)). -Qed. - -Lemma congr_Bool : Bool = Bool. -Proof. -exact (eq_refl). -Qed. - -Lemma congr_If {s0 : Tm} {s1 : Tm} {s2 : Tm} {t0 : Tm} {t1 : Tm} {t2 : Tm} - (H0 : s0 = t0) (H1 : s1 = t1) (H2 : s2 = t2) : If s0 s1 s2 = If t0 t1 t2. -Proof. -exact (eq_trans - (eq_trans (eq_trans eq_refl (ap (fun x => If x s1 s2) H0)) - (ap (fun x => If t0 x s2) H1)) (ap (fun x => If t0 t1 x) H2)). -Qed. - -Lemma upRen_Tm_Tm (xi : nat -> nat) : nat -> nat. -Proof. -exact (up_ren xi). +exact (scons (VarTm (S n_Tm) var_zero) (funcomp (ren_Tm shift) sigma)). Defined. -Fixpoint ren_Tm (xi_Tm : nat -> nat) (s : Tm) {struct s} : Tm := - match s with - | VarTm s0 => VarTm (xi_Tm s0) - | Abs s0 => Abs (ren_Tm (upRen_Tm_Tm xi_Tm) s0) - | App s0 s1 => App (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1) - | Pair s0 s1 => Pair (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1) - | Proj s0 s1 => Proj s0 (ren_Tm xi_Tm s1) - | TBind s0 s1 s2 => - TBind s0 (ren_Tm xi_Tm s1) (ren_Tm (upRen_Tm_Tm xi_Tm) s2) - | Univ s0 => Univ s0 - | BVal s0 => BVal s0 - | Bool => Bool - | If s0 s1 s2 => If (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1) (ren_Tm xi_Tm s2) - end. - -Lemma up_Tm_Tm (sigma : nat -> Tm) : nat -> Tm. +Lemma up_list_Tm_Tm (p : nat) {m : nat} {n_Tm : nat} + (sigma : fin m -> Tm n_Tm) : fin (plus p m) -> Tm (plus p n_Tm). Proof. -exact (scons (VarTm var_zero) (funcomp (ren_Tm shift) sigma)). +exact (scons_p p (funcomp (VarTm (plus p n_Tm)) (zero_p p)) + (funcomp (ren_Tm (shift_p p)) sigma)). Defined. -Fixpoint subst_Tm (sigma_Tm : nat -> Tm) (s : Tm) {struct s} : Tm := +Fixpoint subst_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm) +(s : Tm m_Tm) {struct s} : Tm n_Tm := match s with - | VarTm s0 => sigma_Tm s0 - | Abs s0 => Abs (subst_Tm (up_Tm_Tm sigma_Tm) s0) - | App s0 s1 => App (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1) - | Pair s0 s1 => Pair (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1) - | Proj s0 s1 => Proj s0 (subst_Tm sigma_Tm s1) - | TBind s0 s1 s2 => - TBind s0 (subst_Tm sigma_Tm s1) (subst_Tm (up_Tm_Tm sigma_Tm) s2) - | Univ s0 => Univ s0 - | BVal s0 => BVal s0 - | Bool => Bool - | If s0 s1 s2 => - If (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1) (subst_Tm sigma_Tm s2) + | VarTm _ s0 => sigma_Tm s0 + | Abs _ s0 => Abs n_Tm (subst_Tm (up_Tm_Tm sigma_Tm) s0) + | App _ s0 s1 => App n_Tm (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1) + | Pair _ s0 s1 => Pair n_Tm (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1) + | Proj1 _ s0 => Proj1 n_Tm (subst_Tm sigma_Tm s0) + | Proj2 _ s0 => Proj2 n_Tm (subst_Tm sigma_Tm s0) end. -Lemma upId_Tm_Tm (sigma : nat -> Tm) (Eq : forall x, sigma x = VarTm x) : - forall x, up_Tm_Tm sigma x = VarTm x. +Lemma upId_Tm_Tm {m_Tm : nat} (sigma : fin m_Tm -> Tm m_Tm) + (Eq : forall x, sigma x = VarTm m_Tm x) : + forall x, up_Tm_Tm sigma x = VarTm (S m_Tm) x. Proof. exact (fun n => match n with - | S n' => ap (ren_Tm shift) (Eq n') - | O => eq_refl + | Some fin_n => ap (ren_Tm shift) (Eq fin_n) + | None => eq_refl end). Qed. -Fixpoint idSubst_Tm (sigma_Tm : nat -> Tm) -(Eq_Tm : forall x, sigma_Tm x = VarTm x) (s : Tm) {struct s} : +Lemma upId_list_Tm_Tm {p : nat} {m_Tm : nat} (sigma : fin m_Tm -> Tm m_Tm) + (Eq : forall x, sigma x = VarTm m_Tm x) : + forall x, up_list_Tm_Tm p sigma x = VarTm (plus p m_Tm) x. +Proof. +exact (fun n => + scons_p_eta (VarTm (plus p m_Tm)) + (fun n => ap (ren_Tm (shift_p p)) (Eq n)) (fun n => eq_refl)). +Qed. + +Fixpoint idSubst_Tm {m_Tm : nat} (sigma_Tm : fin m_Tm -> Tm m_Tm) +(Eq_Tm : forall x, sigma_Tm x = VarTm m_Tm x) (s : Tm m_Tm) {struct s} : subst_Tm sigma_Tm s = s := match s with - | VarTm s0 => Eq_Tm s0 - | Abs s0 => + | VarTm _ s0 => Eq_Tm s0 + | Abs _ s0 => congr_Abs (idSubst_Tm (up_Tm_Tm sigma_Tm) (upId_Tm_Tm _ Eq_Tm) s0) - | App s0 s1 => + | App _ s0 s1 => congr_App (idSubst_Tm sigma_Tm Eq_Tm s0) (idSubst_Tm sigma_Tm Eq_Tm s1) - | Pair s0 s1 => + | Pair _ s0 s1 => congr_Pair (idSubst_Tm sigma_Tm Eq_Tm s0) (idSubst_Tm sigma_Tm Eq_Tm s1) - | Proj s0 s1 => congr_Proj (eq_refl s0) (idSubst_Tm sigma_Tm Eq_Tm s1) - | TBind s0 s1 s2 => - congr_TBind (eq_refl s0) (idSubst_Tm sigma_Tm Eq_Tm s1) - (idSubst_Tm (up_Tm_Tm sigma_Tm) (upId_Tm_Tm _ Eq_Tm) s2) - | Univ s0 => congr_Univ (eq_refl s0) - | BVal s0 => congr_BVal (eq_refl s0) - | Bool => congr_Bool - | If s0 s1 s2 => - congr_If (idSubst_Tm sigma_Tm Eq_Tm s0) (idSubst_Tm sigma_Tm Eq_Tm s1) - (idSubst_Tm sigma_Tm Eq_Tm s2) + | Proj1 _ s0 => congr_Proj1 (idSubst_Tm sigma_Tm Eq_Tm s0) + | Proj2 _ s0 => congr_Proj2 (idSubst_Tm sigma_Tm Eq_Tm s0) end. -Lemma upExtRen_Tm_Tm (xi : nat -> nat) (zeta : nat -> nat) - (Eq : forall x, xi x = zeta x) : +Lemma upExtRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n) + (zeta : fin m -> fin n) (Eq : forall x, xi x = zeta x) : forall x, upRen_Tm_Tm xi x = upRen_Tm_Tm zeta x. Proof. -exact (fun n => match n with - | S n' => ap shift (Eq n') - | O => eq_refl - end). +exact (fun n => + match n with + | Some fin_n => ap shift (Eq fin_n) + | None => eq_refl + end). Qed. -Fixpoint extRen_Tm (xi_Tm : nat -> nat) (zeta_Tm : nat -> nat) -(Eq_Tm : forall x, xi_Tm x = zeta_Tm x) (s : Tm) {struct s} : -ren_Tm xi_Tm s = ren_Tm zeta_Tm s := +Lemma upExtRen_list_Tm_Tm {p : nat} {m : nat} {n : nat} (xi : fin m -> fin n) + (zeta : fin m -> fin n) (Eq : forall x, xi x = zeta x) : + forall x, upRen_list_Tm_Tm p xi x = upRen_list_Tm_Tm p zeta x. +Proof. +exact (fun n => + scons_p_congr (fun n => eq_refl) (fun n => ap (shift_p p) (Eq n))). +Qed. + +Fixpoint extRen_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm) +(zeta_Tm : fin m_Tm -> fin n_Tm) (Eq_Tm : forall x, xi_Tm x = zeta_Tm x) +(s : Tm m_Tm) {struct s} : ren_Tm xi_Tm s = ren_Tm zeta_Tm s := match s with - | VarTm s0 => ap (VarTm) (Eq_Tm s0) - | Abs s0 => + | VarTm _ s0 => ap (VarTm n_Tm) (Eq_Tm s0) + | Abs _ s0 => congr_Abs (extRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm) (upExtRen_Tm_Tm _ _ Eq_Tm) s0) - | App s0 s1 => + | App _ s0 s1 => congr_App (extRen_Tm xi_Tm zeta_Tm Eq_Tm s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1) - | Pair s0 s1 => + | Pair _ s0 s1 => congr_Pair (extRen_Tm xi_Tm zeta_Tm Eq_Tm s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1) - | Proj s0 s1 => congr_Proj (eq_refl s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1) - | TBind s0 s1 s2 => - congr_TBind (eq_refl s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1) - (extRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm) - (upExtRen_Tm_Tm _ _ Eq_Tm) s2) - | Univ s0 => congr_Univ (eq_refl s0) - | BVal s0 => congr_BVal (eq_refl s0) - | Bool => congr_Bool - | If s0 s1 s2 => - congr_If (extRen_Tm xi_Tm zeta_Tm Eq_Tm s0) - (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s2) + | Proj1 _ s0 => congr_Proj1 (extRen_Tm xi_Tm zeta_Tm Eq_Tm s0) + | Proj2 _ s0 => congr_Proj2 (extRen_Tm xi_Tm zeta_Tm Eq_Tm s0) end. -Lemma upExt_Tm_Tm (sigma : nat -> Tm) (tau : nat -> Tm) - (Eq : forall x, sigma x = tau x) : +Lemma upExt_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm) + (tau : fin m -> Tm n_Tm) (Eq : forall x, sigma x = tau x) : forall x, up_Tm_Tm sigma x = up_Tm_Tm tau x. Proof. exact (fun n => match n with - | S n' => ap (ren_Tm shift) (Eq n') - | O => eq_refl + | Some fin_n => ap (ren_Tm shift) (Eq fin_n) + | None => eq_refl end). Qed. -Fixpoint ext_Tm (sigma_Tm : nat -> Tm) (tau_Tm : nat -> Tm) -(Eq_Tm : forall x, sigma_Tm x = tau_Tm x) (s : Tm) {struct s} : -subst_Tm sigma_Tm s = subst_Tm tau_Tm s := +Lemma upExt_list_Tm_Tm {p : nat} {m : nat} {n_Tm : nat} + (sigma : fin m -> Tm n_Tm) (tau : fin m -> Tm n_Tm) + (Eq : forall x, sigma x = tau x) : + forall x, up_list_Tm_Tm p sigma x = up_list_Tm_Tm p tau x. +Proof. +exact (fun n => + scons_p_congr (fun n => eq_refl) + (fun n => ap (ren_Tm (shift_p p)) (Eq n))). +Qed. + +Fixpoint ext_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm) +(tau_Tm : fin m_Tm -> Tm n_Tm) (Eq_Tm : forall x, sigma_Tm x = tau_Tm x) +(s : Tm m_Tm) {struct s} : subst_Tm sigma_Tm s = subst_Tm tau_Tm s := match s with - | VarTm s0 => Eq_Tm s0 - | Abs s0 => + | VarTm _ s0 => Eq_Tm s0 + | Abs _ s0 => congr_Abs (ext_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) (upExt_Tm_Tm _ _ Eq_Tm) s0) - | App s0 s1 => + | App _ s0 s1 => congr_App (ext_Tm sigma_Tm tau_Tm Eq_Tm s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1) - | Pair s0 s1 => + | Pair _ s0 s1 => congr_Pair (ext_Tm sigma_Tm tau_Tm Eq_Tm s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1) - | Proj s0 s1 => congr_Proj (eq_refl s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1) - | TBind s0 s1 s2 => - congr_TBind (eq_refl s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1) - (ext_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) (upExt_Tm_Tm _ _ Eq_Tm) - s2) - | Univ s0 => congr_Univ (eq_refl s0) - | BVal s0 => congr_BVal (eq_refl s0) - | Bool => congr_Bool - | If s0 s1 s2 => - congr_If (ext_Tm sigma_Tm tau_Tm Eq_Tm s0) - (ext_Tm sigma_Tm tau_Tm Eq_Tm s1) (ext_Tm sigma_Tm tau_Tm Eq_Tm s2) + | Proj1 _ s0 => congr_Proj1 (ext_Tm sigma_Tm tau_Tm Eq_Tm s0) + | Proj2 _ s0 => congr_Proj2 (ext_Tm sigma_Tm tau_Tm Eq_Tm s0) end. -Lemma up_ren_ren_Tm_Tm (xi : nat -> nat) (zeta : nat -> nat) - (rho : nat -> nat) (Eq : forall x, funcomp zeta xi x = rho x) : +Lemma up_ren_ren_Tm_Tm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l) + (zeta : fin l -> fin m) (rho : fin k -> fin m) + (Eq : forall x, funcomp zeta xi x = rho x) : forall x, funcomp (upRen_Tm_Tm zeta) (upRen_Tm_Tm xi) x = upRen_Tm_Tm rho x. Proof. exact (up_ren_ren xi zeta rho Eq). Qed. -Fixpoint compRenRen_Tm (xi_Tm : nat -> nat) (zeta_Tm : nat -> nat) -(rho_Tm : nat -> nat) (Eq_Tm : forall x, funcomp zeta_Tm xi_Tm x = rho_Tm x) -(s : Tm) {struct s} : ren_Tm zeta_Tm (ren_Tm xi_Tm s) = ren_Tm rho_Tm s := +Lemma up_ren_ren_list_Tm_Tm {p : nat} {k : nat} {l : nat} {m : nat} + (xi : fin k -> fin l) (zeta : fin l -> fin m) (rho : fin k -> fin m) + (Eq : forall x, funcomp zeta xi x = rho x) : + forall x, + funcomp (upRen_list_Tm_Tm p zeta) (upRen_list_Tm_Tm p xi) x = + upRen_list_Tm_Tm p rho x. +Proof. +exact (up_ren_ren_p Eq). +Qed. + +Fixpoint compRenRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} +(xi_Tm : fin m_Tm -> fin k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm) +(rho_Tm : fin m_Tm -> fin l_Tm) +(Eq_Tm : forall x, funcomp zeta_Tm xi_Tm x = rho_Tm x) (s : Tm m_Tm) {struct + s} : ren_Tm zeta_Tm (ren_Tm xi_Tm s) = ren_Tm rho_Tm s := match s with - | VarTm s0 => ap (VarTm) (Eq_Tm s0) - | Abs s0 => + | VarTm _ s0 => ap (VarTm l_Tm) (Eq_Tm s0) + | Abs _ s0 => congr_Abs (compRenRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm) (upRen_Tm_Tm rho_Tm) (up_ren_ren _ _ _ Eq_Tm) s0) - | App s0 s1 => + | App _ s0 s1 => congr_App (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s0) (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1) - | Pair s0 s1 => + | Pair _ s0 s1 => congr_Pair (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s0) (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1) - | Proj s0 s1 => - congr_Proj (eq_refl s0) (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1) - | TBind s0 s1 s2 => - congr_TBind (eq_refl s0) (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1) - (compRenRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm) - (upRen_Tm_Tm rho_Tm) (up_ren_ren _ _ _ Eq_Tm) s2) - | Univ s0 => congr_Univ (eq_refl s0) - | BVal s0 => congr_BVal (eq_refl s0) - | Bool => congr_Bool - | If s0 s1 s2 => - congr_If (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s0) - (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1) - (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s2) + | Proj1 _ s0 => congr_Proj1 (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s0) + | Proj2 _ s0 => congr_Proj2 (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s0) end. -Lemma up_ren_subst_Tm_Tm (xi : nat -> nat) (tau : nat -> Tm) - (theta : nat -> Tm) (Eq : forall x, funcomp tau xi x = theta x) : +Lemma up_ren_subst_Tm_Tm {k : nat} {l : nat} {m_Tm : nat} + (xi : fin k -> fin l) (tau : fin l -> Tm m_Tm) (theta : fin k -> Tm m_Tm) + (Eq : forall x, funcomp tau xi x = theta x) : forall x, funcomp (up_Tm_Tm tau) (upRen_Tm_Tm xi) x = up_Tm_Tm theta x. Proof. exact (fun n => match n with - | S n' => ap (ren_Tm shift) (Eq n') - | O => eq_refl + | Some fin_n => ap (ren_Tm shift) (Eq fin_n) + | None => eq_refl end). Qed. -Fixpoint compRenSubst_Tm (xi_Tm : nat -> nat) (tau_Tm : nat -> Tm) -(theta_Tm : nat -> Tm) -(Eq_Tm : forall x, funcomp tau_Tm xi_Tm x = theta_Tm x) (s : Tm) {struct s} : -subst_Tm tau_Tm (ren_Tm xi_Tm s) = subst_Tm theta_Tm s := +Lemma up_ren_subst_list_Tm_Tm {p : nat} {k : nat} {l : nat} {m_Tm : nat} + (xi : fin k -> fin l) (tau : fin l -> Tm m_Tm) (theta : fin k -> Tm m_Tm) + (Eq : forall x, funcomp tau xi x = theta x) : + forall x, + funcomp (up_list_Tm_Tm p tau) (upRen_list_Tm_Tm p xi) x = + up_list_Tm_Tm p theta x. +Proof. +exact (fun n => + eq_trans (scons_p_comp' _ _ _ n) + (scons_p_congr (fun z => scons_p_head' _ _ z) + (fun z => + eq_trans (scons_p_tail' _ _ (xi z)) + (ap (ren_Tm (shift_p p)) (Eq z))))). +Qed. + +Fixpoint compRenSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} +(xi_Tm : fin m_Tm -> fin k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm) +(theta_Tm : fin m_Tm -> Tm l_Tm) +(Eq_Tm : forall x, funcomp tau_Tm xi_Tm x = theta_Tm x) (s : Tm m_Tm) {struct + s} : subst_Tm tau_Tm (ren_Tm xi_Tm s) = subst_Tm theta_Tm s := match s with - | VarTm s0 => Eq_Tm s0 - | Abs s0 => + | VarTm _ s0 => Eq_Tm s0 + | Abs _ s0 => congr_Abs (compRenSubst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm tau_Tm) (up_Tm_Tm theta_Tm) (up_ren_subst_Tm_Tm _ _ _ Eq_Tm) s0) - | App s0 s1 => + | App _ s0 s1 => congr_App (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s0) (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1) - | Pair s0 s1 => + | Pair _ s0 s1 => congr_Pair (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s0) (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1) - | Proj s0 s1 => - congr_Proj (eq_refl s0) - (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1) - | TBind s0 s1 s2 => - congr_TBind (eq_refl s0) - (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1) - (compRenSubst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm tau_Tm) - (up_Tm_Tm theta_Tm) (up_ren_subst_Tm_Tm _ _ _ Eq_Tm) s2) - | Univ s0 => congr_Univ (eq_refl s0) - | BVal s0 => congr_BVal (eq_refl s0) - | Bool => congr_Bool - | If s0 s1 s2 => - congr_If (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s0) - (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1) - (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s2) + | Proj1 _ s0 => + congr_Proj1 (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s0) + | Proj2 _ s0 => + congr_Proj2 (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s0) end. -Lemma up_subst_ren_Tm_Tm (sigma : nat -> Tm) (zeta_Tm : nat -> nat) - (theta : nat -> Tm) +Lemma up_subst_ren_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat} + (sigma : fin k -> Tm l_Tm) (zeta_Tm : fin l_Tm -> fin m_Tm) + (theta : fin k -> Tm m_Tm) (Eq : forall x, funcomp (ren_Tm zeta_Tm) sigma x = theta x) : forall x, funcomp (ren_Tm (upRen_Tm_Tm zeta_Tm)) (up_Tm_Tm sigma) x = @@ -840,257 +307,308 @@ Lemma up_subst_ren_Tm_Tm (sigma : nat -> Tm) (zeta_Tm : nat -> nat) Proof. exact (fun n => match n with - | S n' => + | Some fin_n => eq_trans (compRenRen_Tm shift (upRen_Tm_Tm zeta_Tm) - (funcomp shift zeta_Tm) (fun x => eq_refl) (sigma n')) + (funcomp shift zeta_Tm) (fun x => eq_refl) (sigma fin_n)) (eq_trans (eq_sym (compRenRen_Tm zeta_Tm shift (funcomp shift zeta_Tm) - (fun x => eq_refl) (sigma n'))) - (ap (ren_Tm shift) (Eq n'))) - | O => eq_refl + (fun x => eq_refl) (sigma fin_n))) + (ap (ren_Tm shift) (Eq fin_n))) + | None => eq_refl end). Qed. -Fixpoint compSubstRen_Tm (sigma_Tm : nat -> Tm) (zeta_Tm : nat -> nat) -(theta_Tm : nat -> Tm) -(Eq_Tm : forall x, funcomp (ren_Tm zeta_Tm) sigma_Tm x = theta_Tm x) -(s : Tm) {struct s} : +Lemma up_subst_ren_list_Tm_Tm {p : nat} {k : nat} {l_Tm : nat} {m_Tm : nat} + (sigma : fin k -> Tm l_Tm) (zeta_Tm : fin l_Tm -> fin m_Tm) + (theta : fin k -> Tm m_Tm) + (Eq : forall x, funcomp (ren_Tm zeta_Tm) sigma x = theta x) : + forall x, + funcomp (ren_Tm (upRen_list_Tm_Tm p zeta_Tm)) (up_list_Tm_Tm p sigma) x = + up_list_Tm_Tm p theta x. +Proof. +exact (fun n => + eq_trans (scons_p_comp' _ _ _ n) + (scons_p_congr + (fun x => ap (VarTm (plus p m_Tm)) (scons_p_head' _ _ x)) + (fun n => + eq_trans + (compRenRen_Tm (shift_p p) (upRen_list_Tm_Tm p zeta_Tm) + (funcomp (shift_p p) zeta_Tm) + (fun x => scons_p_tail' _ _ x) (sigma n)) + (eq_trans + (eq_sym + (compRenRen_Tm zeta_Tm (shift_p p) + (funcomp (shift_p p) zeta_Tm) (fun x => eq_refl) + (sigma n))) (ap (ren_Tm (shift_p p)) (Eq n)))))). +Qed. + +Fixpoint compSubstRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} +(sigma_Tm : fin m_Tm -> Tm k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm) +(theta_Tm : fin m_Tm -> Tm l_Tm) +(Eq_Tm : forall x, funcomp (ren_Tm zeta_Tm) sigma_Tm x = theta_Tm x) +(s : Tm m_Tm) {struct s} : ren_Tm zeta_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s := match s with - | VarTm s0 => Eq_Tm s0 - | Abs s0 => + | VarTm _ s0 => Eq_Tm s0 + | Abs _ s0 => congr_Abs (compSubstRen_Tm (up_Tm_Tm sigma_Tm) (upRen_Tm_Tm zeta_Tm) (up_Tm_Tm theta_Tm) (up_subst_ren_Tm_Tm _ _ _ Eq_Tm) s0) - | App s0 s1 => + | App _ s0 s1 => congr_App (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s0) (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1) - | Pair s0 s1 => + | Pair _ s0 s1 => congr_Pair (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s0) (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1) - | Proj s0 s1 => - congr_Proj (eq_refl s0) - (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1) - | TBind s0 s1 s2 => - congr_TBind (eq_refl s0) - (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1) - (compSubstRen_Tm (up_Tm_Tm sigma_Tm) (upRen_Tm_Tm zeta_Tm) - (up_Tm_Tm theta_Tm) (up_subst_ren_Tm_Tm _ _ _ Eq_Tm) s2) - | Univ s0 => congr_Univ (eq_refl s0) - | BVal s0 => congr_BVal (eq_refl s0) - | Bool => congr_Bool - | If s0 s1 s2 => - congr_If (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s0) - (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1) - (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s2) + | Proj1 _ s0 => + congr_Proj1 (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s0) + | Proj2 _ s0 => + congr_Proj2 (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s0) end. -Lemma up_subst_subst_Tm_Tm (sigma : nat -> Tm) (tau_Tm : nat -> Tm) - (theta : nat -> Tm) +Lemma up_subst_subst_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat} + (sigma : fin k -> Tm l_Tm) (tau_Tm : fin l_Tm -> Tm m_Tm) + (theta : fin k -> Tm m_Tm) (Eq : forall x, funcomp (subst_Tm tau_Tm) sigma x = theta x) : forall x, funcomp (subst_Tm (up_Tm_Tm tau_Tm)) (up_Tm_Tm sigma) x = up_Tm_Tm theta x. Proof. exact (fun n => match n with - | S n' => + | Some fin_n => eq_trans (compRenSubst_Tm shift (up_Tm_Tm tau_Tm) (funcomp (up_Tm_Tm tau_Tm) shift) (fun x => eq_refl) - (sigma n')) + (sigma fin_n)) (eq_trans (eq_sym (compSubstRen_Tm tau_Tm shift (funcomp (ren_Tm shift) tau_Tm) (fun x => eq_refl) - (sigma n'))) (ap (ren_Tm shift) (Eq n'))) - | O => eq_refl + (sigma fin_n))) (ap (ren_Tm shift) (Eq fin_n))) + | None => eq_refl end). Qed. -Fixpoint compSubstSubst_Tm (sigma_Tm : nat -> Tm) (tau_Tm : nat -> Tm) -(theta_Tm : nat -> Tm) +Lemma up_subst_subst_list_Tm_Tm {p : nat} {k : nat} {l_Tm : nat} {m_Tm : nat} + (sigma : fin k -> Tm l_Tm) (tau_Tm : fin l_Tm -> Tm m_Tm) + (theta : fin k -> Tm m_Tm) + (Eq : forall x, funcomp (subst_Tm tau_Tm) sigma x = theta x) : + forall x, + funcomp (subst_Tm (up_list_Tm_Tm p tau_Tm)) (up_list_Tm_Tm p sigma) x = + up_list_Tm_Tm p theta x. +Proof. +exact (fun n => + eq_trans + (scons_p_comp' (funcomp (VarTm (plus p l_Tm)) (zero_p p)) _ _ n) + (scons_p_congr + (fun x => scons_p_head' _ (fun z => ren_Tm (shift_p p) _) x) + (fun n => + eq_trans + (compRenSubst_Tm (shift_p p) (up_list_Tm_Tm p tau_Tm) + (funcomp (up_list_Tm_Tm p tau_Tm) (shift_p p)) + (fun x => eq_refl) (sigma n)) + (eq_trans + (eq_sym + (compSubstRen_Tm tau_Tm (shift_p p) _ + (fun x => eq_sym (scons_p_tail' _ _ x)) (sigma n))) + (ap (ren_Tm (shift_p p)) (Eq n)))))). +Qed. + +Fixpoint compSubstSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} +(sigma_Tm : fin m_Tm -> Tm k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm) +(theta_Tm : fin m_Tm -> Tm l_Tm) (Eq_Tm : forall x, funcomp (subst_Tm tau_Tm) sigma_Tm x = theta_Tm x) -(s : Tm) {struct s} : +(s : Tm m_Tm) {struct s} : subst_Tm tau_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s := match s with - | VarTm s0 => Eq_Tm s0 - | Abs s0 => + | VarTm _ s0 => Eq_Tm s0 + | Abs _ s0 => congr_Abs (compSubstSubst_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) (up_Tm_Tm theta_Tm) (up_subst_subst_Tm_Tm _ _ _ Eq_Tm) s0) - | App s0 s1 => + | App _ s0 s1 => congr_App (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s0) (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1) - | Pair s0 s1 => + | Pair _ s0 s1 => congr_Pair (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s0) (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1) - | Proj s0 s1 => - congr_Proj (eq_refl s0) - (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1) - | TBind s0 s1 s2 => - congr_TBind (eq_refl s0) - (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1) - (compSubstSubst_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) - (up_Tm_Tm theta_Tm) (up_subst_subst_Tm_Tm _ _ _ Eq_Tm) s2) - | Univ s0 => congr_Univ (eq_refl s0) - | BVal s0 => congr_BVal (eq_refl s0) - | Bool => congr_Bool - | If s0 s1 s2 => - congr_If (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s0) - (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1) - (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s2) + | Proj1 _ s0 => + congr_Proj1 (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s0) + | Proj2 _ s0 => + congr_Proj2 (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s0) end. -Lemma renRen_Tm (xi_Tm : nat -> nat) (zeta_Tm : nat -> nat) (s : Tm) : +Lemma renRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} + (xi_Tm : fin m_Tm -> fin k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm) + (s : Tm m_Tm) : ren_Tm zeta_Tm (ren_Tm xi_Tm s) = ren_Tm (funcomp zeta_Tm xi_Tm) s. Proof. exact (compRenRen_Tm xi_Tm zeta_Tm _ (fun n => eq_refl) s). Qed. -Lemma renRen'_Tm_pointwise (xi_Tm : nat -> nat) (zeta_Tm : nat -> nat) : +Lemma renRen'_Tm_pointwise {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} + (xi_Tm : fin m_Tm -> fin k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm) : pointwise_relation _ eq (funcomp (ren_Tm zeta_Tm) (ren_Tm xi_Tm)) (ren_Tm (funcomp zeta_Tm xi_Tm)). Proof. exact (fun s => compRenRen_Tm xi_Tm zeta_Tm _ (fun n => eq_refl) s). Qed. -Lemma renSubst_Tm (xi_Tm : nat -> nat) (tau_Tm : nat -> Tm) (s : Tm) : - subst_Tm tau_Tm (ren_Tm xi_Tm s) = subst_Tm (funcomp tau_Tm xi_Tm) s. +Lemma renSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} + (xi_Tm : fin m_Tm -> fin k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm) (s : Tm m_Tm) + : subst_Tm tau_Tm (ren_Tm xi_Tm s) = subst_Tm (funcomp tau_Tm xi_Tm) s. Proof. exact (compRenSubst_Tm xi_Tm tau_Tm _ (fun n => eq_refl) s). Qed. -Lemma renSubst_Tm_pointwise (xi_Tm : nat -> nat) (tau_Tm : nat -> Tm) : +Lemma renSubst_Tm_pointwise {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} + (xi_Tm : fin m_Tm -> fin k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm) : pointwise_relation _ eq (funcomp (subst_Tm tau_Tm) (ren_Tm xi_Tm)) (subst_Tm (funcomp tau_Tm xi_Tm)). Proof. exact (fun s => compRenSubst_Tm xi_Tm tau_Tm _ (fun n => eq_refl) s). Qed. -Lemma substRen_Tm (sigma_Tm : nat -> Tm) (zeta_Tm : nat -> nat) (s : Tm) : +Lemma substRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} + (sigma_Tm : fin m_Tm -> Tm k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm) + (s : Tm m_Tm) : ren_Tm zeta_Tm (subst_Tm sigma_Tm s) = subst_Tm (funcomp (ren_Tm zeta_Tm) sigma_Tm) s. Proof. exact (compSubstRen_Tm sigma_Tm zeta_Tm _ (fun n => eq_refl) s). Qed. -Lemma substRen_Tm_pointwise (sigma_Tm : nat -> Tm) (zeta_Tm : nat -> nat) : +Lemma substRen_Tm_pointwise {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} + (sigma_Tm : fin m_Tm -> Tm k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm) : pointwise_relation _ eq (funcomp (ren_Tm zeta_Tm) (subst_Tm sigma_Tm)) (subst_Tm (funcomp (ren_Tm zeta_Tm) sigma_Tm)). Proof. exact (fun s => compSubstRen_Tm sigma_Tm zeta_Tm _ (fun n => eq_refl) s). Qed. -Lemma substSubst_Tm (sigma_Tm : nat -> Tm) (tau_Tm : nat -> Tm) (s : Tm) : +Lemma substSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} + (sigma_Tm : fin m_Tm -> Tm k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm) + (s : Tm m_Tm) : subst_Tm tau_Tm (subst_Tm sigma_Tm s) = subst_Tm (funcomp (subst_Tm tau_Tm) sigma_Tm) s. Proof. exact (compSubstSubst_Tm sigma_Tm tau_Tm _ (fun n => eq_refl) s). Qed. -Lemma substSubst_Tm_pointwise (sigma_Tm : nat -> Tm) (tau_Tm : nat -> Tm) : +Lemma substSubst_Tm_pointwise {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} + (sigma_Tm : fin m_Tm -> Tm k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm) : pointwise_relation _ eq (funcomp (subst_Tm tau_Tm) (subst_Tm sigma_Tm)) (subst_Tm (funcomp (subst_Tm tau_Tm) sigma_Tm)). Proof. exact (fun s => compSubstSubst_Tm sigma_Tm tau_Tm _ (fun n => eq_refl) s). Qed. -Lemma rinstInst_up_Tm_Tm (xi : nat -> nat) (sigma : nat -> Tm) - (Eq : forall x, funcomp (VarTm) xi x = sigma x) : - forall x, funcomp (VarTm) (upRen_Tm_Tm xi) x = up_Tm_Tm sigma x. +Lemma rinstInst_up_Tm_Tm {m : nat} {n_Tm : nat} (xi : fin m -> fin n_Tm) + (sigma : fin m -> Tm n_Tm) + (Eq : forall x, funcomp (VarTm n_Tm) xi x = sigma x) : + forall x, funcomp (VarTm (S n_Tm)) (upRen_Tm_Tm xi) x = up_Tm_Tm sigma x. Proof. exact (fun n => match n with - | S n' => ap (ren_Tm shift) (Eq n') - | O => eq_refl + | Some fin_n => ap (ren_Tm shift) (Eq fin_n) + | None => eq_refl end). Qed. -Fixpoint rinst_inst_Tm (xi_Tm : nat -> nat) (sigma_Tm : nat -> Tm) -(Eq_Tm : forall x, funcomp (VarTm) xi_Tm x = sigma_Tm x) (s : Tm) {struct s} - : ren_Tm xi_Tm s = subst_Tm sigma_Tm s := +Lemma rinstInst_up_list_Tm_Tm {p : nat} {m : nat} {n_Tm : nat} + (xi : fin m -> fin n_Tm) (sigma : fin m -> Tm n_Tm) + (Eq : forall x, funcomp (VarTm n_Tm) xi x = sigma x) : + forall x, + funcomp (VarTm (plus p n_Tm)) (upRen_list_Tm_Tm p xi) x = + up_list_Tm_Tm p sigma x. +Proof. +exact (fun n => + eq_trans (scons_p_comp' _ _ (VarTm (plus p n_Tm)) n) + (scons_p_congr (fun z => eq_refl) + (fun n => ap (ren_Tm (shift_p p)) (Eq n)))). +Qed. + +Fixpoint rinst_inst_Tm {m_Tm : nat} {n_Tm : nat} +(xi_Tm : fin m_Tm -> fin n_Tm) (sigma_Tm : fin m_Tm -> Tm n_Tm) +(Eq_Tm : forall x, funcomp (VarTm n_Tm) xi_Tm x = sigma_Tm x) (s : Tm m_Tm) +{struct s} : ren_Tm xi_Tm s = subst_Tm sigma_Tm s := match s with - | VarTm s0 => Eq_Tm s0 - | Abs s0 => + | VarTm _ s0 => Eq_Tm s0 + | Abs _ s0 => congr_Abs (rinst_inst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm sigma_Tm) (rinstInst_up_Tm_Tm _ _ Eq_Tm) s0) - | App s0 s1 => + | App _ s0 s1 => congr_App (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s0) (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1) - | Pair s0 s1 => + | Pair _ s0 s1 => congr_Pair (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s0) (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1) - | Proj s0 s1 => - congr_Proj (eq_refl s0) (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1) - | TBind s0 s1 s2 => - congr_TBind (eq_refl s0) (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1) - (rinst_inst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm sigma_Tm) - (rinstInst_up_Tm_Tm _ _ Eq_Tm) s2) - | Univ s0 => congr_Univ (eq_refl s0) - | BVal s0 => congr_BVal (eq_refl s0) - | Bool => congr_Bool - | If s0 s1 s2 => - congr_If (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s0) - (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1) - (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s2) + | Proj1 _ s0 => congr_Proj1 (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s0) + | Proj2 _ s0 => congr_Proj2 (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s0) end. -Lemma rinstInst'_Tm (xi_Tm : nat -> nat) (s : Tm) : - ren_Tm xi_Tm s = subst_Tm (funcomp (VarTm) xi_Tm) s. +Lemma rinstInst'_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm) + (s : Tm m_Tm) : ren_Tm xi_Tm s = subst_Tm (funcomp (VarTm n_Tm) xi_Tm) s. Proof. exact (rinst_inst_Tm xi_Tm _ (fun n => eq_refl) s). Qed. -Lemma rinstInst'_Tm_pointwise (xi_Tm : nat -> nat) : - pointwise_relation _ eq (ren_Tm xi_Tm) (subst_Tm (funcomp (VarTm) xi_Tm)). +Lemma rinstInst'_Tm_pointwise {m_Tm : nat} {n_Tm : nat} + (xi_Tm : fin m_Tm -> fin n_Tm) : + pointwise_relation _ eq (ren_Tm xi_Tm) + (subst_Tm (funcomp (VarTm n_Tm) xi_Tm)). Proof. exact (fun s => rinst_inst_Tm xi_Tm _ (fun n => eq_refl) s). Qed. -Lemma instId'_Tm (s : Tm) : subst_Tm (VarTm) s = s. +Lemma instId'_Tm {m_Tm : nat} (s : Tm m_Tm) : subst_Tm (VarTm m_Tm) s = s. Proof. -exact (idSubst_Tm (VarTm) (fun n => eq_refl) s). +exact (idSubst_Tm (VarTm m_Tm) (fun n => eq_refl) s). Qed. -Lemma instId'_Tm_pointwise : pointwise_relation _ eq (subst_Tm (VarTm)) id. +Lemma instId'_Tm_pointwise {m_Tm : nat} : + pointwise_relation _ eq (subst_Tm (VarTm m_Tm)) id. Proof. -exact (fun s => idSubst_Tm (VarTm) (fun n => eq_refl) s). +exact (fun s => idSubst_Tm (VarTm m_Tm) (fun n => eq_refl) s). Qed. -Lemma rinstId'_Tm (s : Tm) : ren_Tm id s = s. +Lemma rinstId'_Tm {m_Tm : nat} (s : Tm m_Tm) : ren_Tm id s = s. Proof. exact (eq_ind_r (fun t => t = s) (instId'_Tm s) (rinstInst'_Tm id s)). Qed. -Lemma rinstId'_Tm_pointwise : pointwise_relation _ eq (@ren_Tm id) id. +Lemma rinstId'_Tm_pointwise {m_Tm : nat} : + pointwise_relation _ eq (@ren_Tm m_Tm m_Tm id) id. Proof. exact (fun s => eq_ind_r (fun t => t = s) (instId'_Tm s) (rinstInst'_Tm id s)). Qed. -Lemma varL'_Tm (sigma_Tm : nat -> Tm) (x : nat) : - subst_Tm sigma_Tm (VarTm x) = sigma_Tm x. +Lemma varL'_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm) + (x : fin m_Tm) : subst_Tm sigma_Tm (VarTm m_Tm x) = sigma_Tm x. Proof. exact (eq_refl). Qed. -Lemma varL'_Tm_pointwise (sigma_Tm : nat -> Tm) : - pointwise_relation _ eq (funcomp (subst_Tm sigma_Tm) (VarTm)) sigma_Tm. +Lemma varL'_Tm_pointwise {m_Tm : nat} {n_Tm : nat} + (sigma_Tm : fin m_Tm -> Tm n_Tm) : + pointwise_relation _ eq (funcomp (subst_Tm sigma_Tm) (VarTm m_Tm)) sigma_Tm. Proof. exact (fun x => eq_refl). Qed. -Lemma varLRen'_Tm (xi_Tm : nat -> nat) (x : nat) : - ren_Tm xi_Tm (VarTm x) = VarTm (xi_Tm x). +Lemma varLRen'_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm) + (x : fin m_Tm) : ren_Tm xi_Tm (VarTm m_Tm x) = VarTm n_Tm (xi_Tm x). Proof. exact (eq_refl). Qed. -Lemma varLRen'_Tm_pointwise (xi_Tm : nat -> nat) : - pointwise_relation _ eq (funcomp (ren_Tm xi_Tm) (VarTm)) - (funcomp (VarTm) xi_Tm). +Lemma varLRen'_Tm_pointwise {m_Tm : nat} {n_Tm : nat} + (xi_Tm : fin m_Tm -> fin n_Tm) : + pointwise_relation _ eq (funcomp (ren_Tm xi_Tm) (VarTm m_Tm)) + (funcomp (VarTm n_Tm) xi_Tm). Proof. exact (fun x => eq_refl). Qed. @@ -1098,25 +616,17 @@ Qed. Class Up_Tm X Y := up_Tm : X -> Y. -Class Up_PTm X Y := - up_PTm : X -> Y. - -#[global] Instance Subst_Tm : (Subst1 _ _ _) := @subst_Tm. - -#[global] Instance Up_Tm_Tm : (Up_Tm _ _) := @up_Tm_Tm. - -#[global] Instance Ren_Tm : (Ren1 _ _ _) := @ren_Tm. - -#[global] Instance VarInstance_Tm : (Var _ _) := @VarTm. - -#[global] Instance Subst_PTm : (Subst1 _ _ _) := @subst_PTm. - -#[global] Instance Up_PTm_PTm : (Up_PTm _ _) := @up_PTm_PTm. - -#[global] Instance Ren_PTm : (Ren1 _ _ _) := @ren_PTm. +#[global] +Instance Subst_Tm {m_Tm n_Tm : nat}: (Subst1 _ _ _) := (@subst_Tm m_Tm n_Tm). #[global] -Instance VarInstance_PTm : (Var _ _) := @VarPTm. +Instance Up_Tm_Tm {m n_Tm : nat}: (Up_Tm _ _) := (@up_Tm_Tm m n_Tm). + +#[global] +Instance Ren_Tm {m_Tm n_Tm : nat}: (Ren1 _ _ _) := (@ren_Tm m_Tm n_Tm). + +#[global] +Instance VarInstance_Tm {n_Tm : nat}: (Var _ _) := (@VarTm n_Tm). Notation "[ sigma_Tm ]" := (subst_Tm sigma_Tm) ( at level 1, left associativity, only printing) : fscope. @@ -1142,34 +652,10 @@ Notation "x '__Tm'" := (@ids _ _ VarInstance_Tm x) Notation "x '__Tm'" := (VarTm x) ( at level 5, format "x __Tm") : subst_scope. -Notation "[ sigma_PTm ]" := (subst_PTm sigma_PTm) -( at level 1, left associativity, only printing) : fscope. - -Notation "s [ sigma_PTm ]" := (subst_PTm sigma_PTm s) -( at level 7, left associativity, only printing) : subst_scope. - -Notation "↑__PTm" := up_PTm (only printing) : subst_scope. - -Notation "↑__PTm" := up_PTm_PTm (only printing) : subst_scope. - -Notation "⟨ xi_PTm ⟩" := (ren_PTm xi_PTm) -( at level 1, left associativity, only printing) : fscope. - -Notation "s ⟨ xi_PTm ⟩" := (ren_PTm xi_PTm s) -( at level 7, left associativity, only printing) : subst_scope. - -Notation "'var'" := VarPTm ( at level 1, only printing) : subst_scope. - -Notation "x '__PTm'" := (@ids _ _ VarInstance_PTm x) -( at level 5, format "x __PTm", only printing) : subst_scope. - -Notation "x '__PTm'" := (VarPTm x) ( at level 5, format "x __PTm") : -subst_scope. - #[global] -Instance subst_Tm_morphism : +Instance subst_Tm_morphism {m_Tm : nat} {n_Tm : nat}: (Proper (respectful (pointwise_relation _ eq) (respectful eq eq)) - (@subst_Tm)). + (@subst_Tm m_Tm n_Tm)). Proof. exact (fun f_Tm g_Tm Eq_Tm s t Eq_st => eq_ind s (fun t' => subst_Tm f_Tm s = subst_Tm g_Tm t') @@ -1177,16 +663,17 @@ exact (fun f_Tm g_Tm Eq_Tm s t Eq_st => Qed. #[global] -Instance subst_Tm_morphism2 : +Instance subst_Tm_morphism2 {m_Tm : nat} {n_Tm : nat}: (Proper (respectful (pointwise_relation _ eq) (pointwise_relation _ eq)) - (@subst_Tm)). + (@subst_Tm m_Tm n_Tm)). Proof. exact (fun f_Tm g_Tm Eq_Tm s => ext_Tm f_Tm g_Tm Eq_Tm s). Qed. #[global] -Instance ren_Tm_morphism : - (Proper (respectful (pointwise_relation _ eq) (respectful eq eq)) (@ren_Tm)). +Instance ren_Tm_morphism {m_Tm : nat} {n_Tm : nat}: + (Proper (respectful (pointwise_relation _ eq) (respectful eq eq)) + (@ren_Tm m_Tm n_Tm)). Proof. exact (fun f_Tm g_Tm Eq_Tm s t Eq_st => eq_ind s (fun t' => ren_Tm f_Tm s = ren_Tm g_Tm t') @@ -1194,62 +681,22 @@ exact (fun f_Tm g_Tm Eq_Tm s t Eq_st => Qed. #[global] -Instance ren_Tm_morphism2 : +Instance ren_Tm_morphism2 {m_Tm : nat} {n_Tm : nat}: (Proper (respectful (pointwise_relation _ eq) (pointwise_relation _ eq)) - (@ren_Tm)). + (@ren_Tm m_Tm n_Tm)). Proof. exact (fun f_Tm g_Tm Eq_Tm s => extRen_Tm f_Tm g_Tm Eq_Tm s). Qed. -#[global] -Instance subst_PTm_morphism : - (Proper (respectful (pointwise_relation _ eq) (respectful eq eq)) - (@subst_PTm)). -Proof. -exact (fun f_PTm g_PTm Eq_PTm s t Eq_st => - eq_ind s (fun t' => subst_PTm f_PTm s = subst_PTm g_PTm t') - (ext_PTm f_PTm g_PTm Eq_PTm s) t Eq_st). -Qed. - -#[global] -Instance subst_PTm_morphism2 : - (Proper (respectful (pointwise_relation _ eq) (pointwise_relation _ eq)) - (@subst_PTm)). -Proof. -exact (fun f_PTm g_PTm Eq_PTm s => ext_PTm f_PTm g_PTm Eq_PTm s). -Qed. - -#[global] -Instance ren_PTm_morphism : - (Proper (respectful (pointwise_relation _ eq) (respectful eq eq)) (@ren_PTm)). -Proof. -exact (fun f_PTm g_PTm Eq_PTm s t Eq_st => - eq_ind s (fun t' => ren_PTm f_PTm s = ren_PTm g_PTm t') - (extRen_PTm f_PTm g_PTm Eq_PTm s) t Eq_st). -Qed. - -#[global] -Instance ren_PTm_morphism2 : - (Proper (respectful (pointwise_relation _ eq) (pointwise_relation _ eq)) - (@ren_PTm)). -Proof. -exact (fun f_PTm g_PTm Eq_PTm s => extRen_PTm f_PTm g_PTm Eq_PTm s). -Qed. - Ltac auto_unfold := repeat - unfold VarInstance_PTm, Var, ids, Ren_PTm, Ren1, ren1, - Up_PTm_PTm, Up_PTm, up_PTm, Subst_PTm, Subst1, subst1, - VarInstance_Tm, Var, ids, Ren_Tm, Ren1, ren1, Up_Tm_Tm, - Up_Tm, up_Tm, Subst_Tm, Subst1, subst1. + unfold VarInstance_Tm, Var, ids, Ren_Tm, Ren1, ren1, + Up_Tm_Tm, Up_Tm, up_Tm, Subst_Tm, Subst1, subst1. Tactic Notation "auto_unfold" "in" "*" := repeat - unfold VarInstance_PTm, Var, ids, - Ren_PTm, Ren1, ren1, Up_PTm_PTm, - Up_PTm, up_PTm, Subst_PTm, - Subst1, subst1, VarInstance_Tm, - Var, ids, Ren_Tm, Ren1, ren1, - Up_Tm_Tm, Up_Tm, up_Tm, Subst_Tm, - Subst1, subst1 in *. + unfold VarInstance_Tm, Var, ids, + Ren_Tm, Ren1, ren1, Up_Tm_Tm, + Up_Tm, up_Tm, Subst_Tm, Subst1, + subst1 in *. Ltac asimpl' := repeat (first [ progress setoid_rewrite substSubst_Tm_pointwise @@ -1260,14 +707,6 @@ Ltac asimpl' := repeat (first | progress setoid_rewrite renSubst_Tm | progress setoid_rewrite renRen'_Tm_pointwise | progress setoid_rewrite renRen_Tm - | progress setoid_rewrite substSubst_PTm_pointwise - | progress setoid_rewrite substSubst_PTm - | progress setoid_rewrite substRen_PTm_pointwise - | progress setoid_rewrite substRen_PTm - | progress setoid_rewrite renSubst_PTm_pointwise - | progress setoid_rewrite renSubst_PTm - | progress setoid_rewrite renRen'_PTm_pointwise - | progress setoid_rewrite renRen_PTm | progress setoid_rewrite varLRen'_Tm_pointwise | progress setoid_rewrite varLRen'_Tm | progress setoid_rewrite varL'_Tm_pointwise @@ -1276,55 +715,50 @@ Ltac asimpl' := repeat (first | progress setoid_rewrite rinstId'_Tm | progress setoid_rewrite instId'_Tm_pointwise | progress setoid_rewrite instId'_Tm - | progress setoid_rewrite varLRen'_PTm_pointwise - | progress setoid_rewrite varLRen'_PTm - | progress setoid_rewrite varL'_PTm_pointwise - | progress setoid_rewrite varL'_PTm - | progress setoid_rewrite rinstId'_PTm_pointwise - | progress setoid_rewrite rinstId'_PTm - | progress setoid_rewrite instId'_PTm_pointwise - | progress setoid_rewrite instId'_PTm | progress - unfold up_Tm_Tm, upRen_Tm_Tm, up_PTm_PTm, upRen_PTm_PTm, - up_ren - | progress cbn[subst_Tm ren_Tm subst_PTm ren_PTm] + unfold up_list_Tm_Tm, up_Tm_Tm, upRen_list_Tm_Tm, + upRen_Tm_Tm, up_ren + | progress cbn[subst_Tm ren_Tm] | progress fsimpl ]). Ltac asimpl := check_no_evars; repeat - unfold VarInstance_PTm, Var, ids, Ren_PTm, Ren1, ren1, - Up_PTm_PTm, Up_PTm, up_PTm, Subst_PTm, Subst1, subst1, - VarInstance_Tm, Var, ids, Ren_Tm, Ren1, ren1, Up_Tm_Tm, - Up_Tm, up_Tm, Subst_Tm, Subst1, subst1 in *; asimpl'; - minimize. + unfold VarInstance_Tm, Var, ids, Ren_Tm, Ren1, ren1, + Up_Tm_Tm, Up_Tm, up_Tm, Subst_Tm, Subst1, subst1 in *; + asimpl'; minimize. Tactic Notation "asimpl" "in" hyp(J) := revert J; asimpl; intros J. Tactic Notation "auto_case" := auto_case ltac:(asimpl; cbn; eauto). Ltac substify := auto_unfold; try setoid_rewrite rinstInst'_Tm_pointwise; - try setoid_rewrite rinstInst'_Tm; - try setoid_rewrite rinstInst'_PTm_pointwise; - try setoid_rewrite rinstInst'_PTm. + try setoid_rewrite rinstInst'_Tm. Ltac renamify := auto_unfold; try setoid_rewrite_left rinstInst'_Tm_pointwise; - try setoid_rewrite_left rinstInst'_Tm; - try setoid_rewrite_left rinstInst'_PTm_pointwise; - try setoid_rewrite_left rinstInst'_PTm. + try setoid_rewrite_left rinstInst'_Tm. End Core. Module Extra. -Import Core. +Import +Core. -#[global] Hint Opaque subst_Tm: rewrite. +Arguments VarTm {n_Tm}. -#[global] Hint Opaque ren_Tm: rewrite. +Arguments Proj2 {n_Tm}. -#[global] Hint Opaque subst_PTm: rewrite. +Arguments Proj1 {n_Tm}. -#[global] Hint Opaque ren_PTm: rewrite. +Arguments Pair {n_Tm}. + +Arguments App {n_Tm}. + +Arguments Abs {n_Tm}. + +#[global]Hint Opaque subst_Tm: rewrite. + +#[global]Hint Opaque ren_Tm: rewrite. End Extra. diff --git a/theories/Autosubst2/unscoped.v b/theories/Autosubst2/unscoped.v deleted file mode 100644 index 55f8172..0000000 --- a/theories/Autosubst2/unscoped.v +++ /dev/null @@ -1,213 +0,0 @@ -(** * Autosubst Header for Unnamed Syntax - -Version: December 11, 2019. - *) - -(* Adrian: - I changed this library a bit to work better with my generated code. - 1. I use nat directly instead of defining fin to be nat and using Some/None as S/O - 2. I removed the "s, sigma" notation for scons because it interacts with dependent function types "forall x, A"*) -Require Import core. -Require Import Setoid Morphisms Relation_Definitions. - -Definition ap {X Y} (f : X -> Y) {x y : X} (p : x = y) : f x = f y := - match p with eq_refl => eq_refl end. - -Definition apc {X Y} {f g : X -> Y} {x y : X} (p : f = g) (q : x = y) : f x = g y := - match q with eq_refl => match p with eq_refl => eq_refl end end. - -(** ** Primitives of the Sigma Calculus. *) - -Definition shift := S. - -Definition var_zero := 0. - -Definition id {X} := @Datatypes.id X. - -Definition scons {X: Type} (x : X) (xi : nat -> X) := - fun n => match n with - | 0 => x - | S n => xi n - end. - -#[ export ] -Hint Opaque scons : rewrite. - -(** ** Type Class Instances for Notation -Required to make notation work. *) - -(** *** Type classes for renamings. *) - -Class Ren1 (X1 : Type) (Y Z : Type) := - ren1 : X1 -> Y -> Z. - -Class Ren2 (X1 X2 : Type) (Y Z : Type) := - ren2 : X1 -> X2 -> Y -> Z. - -Class Ren3 (X1 X2 X3 : Type) (Y Z : Type) := - ren3 : X1 -> X2 -> X3 -> Y -> Z. - -Class Ren4 (X1 X2 X3 X4 : Type) (Y Z : Type) := - ren4 : X1 -> X2 -> X3 -> X4 -> Y -> Z. - -Class Ren5 (X1 X2 X3 X4 X5 : Type) (Y Z : Type) := - ren5 : X1 -> X2 -> X3 -> X4 -> X5 -> Y -> Z. - -Module RenNotations. - Notation "s ⟨ xi1 ⟩" := (ren1 xi1 s) (at level 7, left associativity, format "s ⟨ xi1 ⟩") : subst_scope. - - Notation "s ⟨ xi1 ; xi2 ⟩" := (ren2 xi1 xi2 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ⟩") : subst_scope. - - Notation "s ⟨ xi1 ; xi2 ; xi3 ⟩" := (ren3 xi1 xi2 xi3 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ; xi3 ⟩") : subst_scope. - - Notation "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ⟩" := (ren4 xi1 xi2 xi3 xi4 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ⟩") : subst_scope. - - Notation "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ; xi5 ⟩" := (ren5 xi1 xi2 xi3 xi4 xi5 s) (at level 7, left associativity, format "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ; xi5 ⟩") : subst_scope. - - Notation "⟨ xi ⟩" := (ren1 xi) (at level 1, left associativity, format "⟨ xi ⟩") : fscope. - - Notation "⟨ xi1 ; xi2 ⟩" := (ren2 xi1 xi2) (at level 1, left associativity, format "⟨ xi1 ; xi2 ⟩") : fscope. -End RenNotations. - -(** *** Type Classes for Substiution *) - -Class Subst1 (X1 : Type) (Y Z: Type) := - subst1 : X1 -> Y -> Z. - -Class Subst2 (X1 X2 : Type) (Y Z: Type) := - subst2 : X1 -> X2 -> Y -> Z. - -Class Subst3 (X1 X2 X3 : Type) (Y Z: Type) := - subst3 : X1 -> X2 -> X3 -> Y -> Z. - -Class Subst4 (X1 X2 X3 X4: Type) (Y Z: Type) := - subst4 : X1 -> X2 -> X3 -> X4 -> Y -> Z. - -Class Subst5 (X1 X2 X3 X4 X5 : Type) (Y Z: Type) := - subst5 : X1 -> X2 -> X3 -> X4 -> X5 -> Y -> Z. - -Module SubstNotations. - Notation "s [ sigma ]" := (subst1 sigma s) (at level 7, left associativity, format "s '/' [ sigma ]") : subst_scope. - - Notation "s [ sigma ; tau ]" := (subst2 sigma tau s) (at level 7, left associativity, format "s '/' [ sigma ; '/' tau ]") : subst_scope. -End SubstNotations. - -(** *** Type Class for Variables *) - -Class Var X Y := - ids : X -> Y. - -#[export] Instance idsRen : Var nat nat := id. - -(** ** Proofs for the substitution primitives. *) - -Arguments funcomp {X Y Z} (g)%fscope (f)%fscope. - -Module CombineNotations. - Notation "f >> g" := (funcomp g f) (at level 50) : fscope. - - Notation "s .: sigma" := (scons s sigma) (at level 55, sigma at next level, right associativity) : subst_scope. - - #[ global ] - Open Scope fscope. - #[ global ] - Open Scope subst_scope. -End CombineNotations. - -Import CombineNotations. - - -(** A generic lifting of a renaming. *) -Definition up_ren (xi : nat -> nat) := - 0 .: (xi >> S). - -(** A generic proof that lifting of renamings composes. *) -Lemma up_ren_ren (xi: nat -> nat) (zeta : nat -> nat) (rho: nat -> nat) (E: forall x, (xi >> zeta) x = rho x) : - forall x, (up_ren xi >> up_ren zeta) x = up_ren rho x. -Proof. - intros [|x]. - - reflexivity. - - unfold up_ren. cbn. unfold funcomp. f_equal. apply E. -Qed. - -(** Eta laws. *) -Lemma scons_eta' {T} (f : nat -> T) : - pointwise_relation _ eq (f var_zero .: (funcomp f shift)) f. -Proof. intros x. destruct x; reflexivity. Qed. - -Lemma scons_eta_id' : - pointwise_relation _ eq (var_zero .: shift) id. -Proof. intros x. destruct x; reflexivity. Qed. - -Lemma scons_comp' (T: Type) {U} (s: T) (sigma: nat -> T) (tau: T -> U) : - pointwise_relation _ eq (funcomp tau (s .: sigma)) ((tau s) .: (funcomp tau sigma)). -Proof. intros x. destruct x; reflexivity. Qed. - -(* Morphism for Setoid Rewriting. The only morphism that can be defined statically. *) -#[export] Instance scons_morphism {X: Type} : - Proper (eq ==> pointwise_relation _ eq ==> pointwise_relation _ eq) (@scons X). -Proof. - intros ? t -> sigma tau H. - intros [|x]. - cbn. reflexivity. - apply H. -Qed. - -#[export] Instance scons_morphism2 {X: Type} : - Proper (eq ==> pointwise_relation _ eq ==> eq ==> eq) (@scons X). -Proof. - intros ? t -> sigma tau H ? x ->. - destruct x as [|x]. - cbn. reflexivity. - apply H. -Qed. - -(** ** Generic lifting of an allfv predicate *) -Definition up_allfv (p: nat -> Prop) : nat -> Prop := scons True p. - -(** ** Notations for unscoped syntax *) -Module UnscopedNotations. - Include RenNotations. - Include SubstNotations. - Include CombineNotations. - - (* Notation "s , sigma" := (scons s sigma) (at level 60, format "s , sigma", right associativity) : subst_scope. *) - - Notation "s '..'" := (scons s ids) (at level 1, format "s ..") : subst_scope. - - Notation "↑" := (shift) : subst_scope. - - #[global] - Open Scope fscope. - #[global] - Open Scope subst_scope. -End UnscopedNotations. - -(** ** Tactics for unscoped syntax *) - -(** Automatically does a case analysis on a natural number, useful for proofs with context renamings/context morphisms. *) -Tactic Notation "auto_case" tactic(t) := (match goal with - | [|- forall (i : nat), _] => intros []; t - end). - - -(** Generic fsimpl tactic: simplifies the above primitives in a goal. *) -Ltac fsimpl := - repeat match goal with - | [|- context[id >> ?f]] => change (id >> f) with f (* AsimplCompIdL *) - | [|- context[?f >> id]] => change (f >> id) with f (* AsimplCompIdR *) - | [|- context [id ?s]] => change (id s) with s - | [|- context[(?f >> ?g) >> ?h]] => change ((f >> g) >> h) with (f >> (g >> h)) - | [|- context[(?v .: ?g) var_zero]] => change ((v .: g) var_zero) with v - | [|- context[(?v .: ?g) 0]] => change ((v .: g) 0) with v - | [|- context[(?v .: ?g) (S ?n)]] => change ((v .: g) (S n)) with (g n) - | [|- context[?f >> (?x .: ?g)]] => change (f >> (x .: g)) with g (* f should evaluate to shift *) - | [|- context[var_zero]] => change var_zero with 0 - | [|- context[?x2 .: (funcomp ?f shift)]] => change (scons x2 (funcomp f shift)) with (scons (f var_zero) (funcomp f shift)); setoid_rewrite (@scons_eta' _ _ f) - | [|- context[?f var_zero .: ?g]] => change (scons (f var_zero) g) with (scons (f var_zero) (funcomp f shift)); rewrite scons_eta' - | [|- _ = ?h (?f ?s)] => change (h (f s)) with ((f >> h) s) - | [|- ?h (?f ?s) = _] => change (h (f s)) with ((f >> h) s) - (* DONE had to put an underscore as the last argument to scons. This might be an argument against unfolding funcomp *) - | [|- context[funcomp _ (scons _ _)]] => setoid_rewrite scons_comp'; eta_reduce - | [|- context[scons var_zero shift]] => setoid_rewrite scons_eta_id'; eta_reduce - end. \ No newline at end of file diff --git a/theories/compile.v b/theories/compile.v deleted file mode 100644 index e481d71..0000000 --- a/theories/compile.v +++ /dev/null @@ -1,186 +0,0 @@ -Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax fp_red. -Require Import ssreflect ssrbool. -From Hammer Require Import Tactics. -From stdpp Require Import relations (rtc(..)). - -Module Compile. - Fixpoint F {n} (a : Tm n) : Tm n := - match a with - | TBind p A B => Pair (Pair (Const p) (F A)) (Abs (F B)) - | Const k => Const k - | Univ i => Univ i - | Abs a => Abs (F a) - | App a b => App (F a) (F b) - | VarTm i => VarTm i - | Pair a b => Pair (F a) (F b) - | Proj t a => Proj t (F a) - | Bot => Bot - | If a b c => App (App (F a) (F b)) (F c) - | BVal b => if b then (Abs (Abs (VarTm (shift var_zero)))) else (Abs (Abs (VarTm var_zero))) - | Bool => Bool - end. - - Lemma renaming n m (a : Tm n) (ξ : fin n -> fin m) : - F (ren_Tm ξ a)= ren_Tm ξ (F a). - Proof. move : m ξ. elim : n / a => //=; hauto lq:on. Qed. - - #[local]Hint Rewrite Compile.renaming : compile. - - Lemma morphing n m (a : Tm n) (ρ0 ρ1 : fin n -> Tm m) : - (forall i, ρ0 i = F (ρ1 i)) -> - subst_Tm ρ0 (F a) = F (subst_Tm ρ1 a). - Proof. - move : m ρ0 ρ1. elim : n / a => n//=. - - hauto lq:on inv:option rew:db:compile unfold:funcomp. - - hauto lq:on rew:off. - - hauto lq:on rew:off. - - hauto lq:on. - - hauto lq:on inv:option rew:db:compile unfold:funcomp. - - hauto lq:on rew:off. - - hauto lq:on rew:off. - Qed. - - Lemma substing n b (a : Tm (S n)) : - subst_Tm (scons (F b) VarTm) (F a) = F (subst_Tm (scons b VarTm) a). - Proof. - apply morphing. - case => //=. - Qed. - -End Compile. - -#[export] Hint Rewrite Compile.renaming Compile.morphing : compile. - - -Module Join. - Definition R {n} (a b : Tm n) := join (Compile.F a) (Compile.F b). - - Lemma BindInj n p0 p1 (A0 A1 : Tm n) B0 B1 : - R (TBind p0 A0 B0) (TBind p1 A1 B1) -> p0 = p1 /\ R A0 A1 /\ R B0 B1. - Proof. - rewrite /R /= !join_pair_inj. - move => [[/join_const_inj h0 h1] h2]. - apply abs_eq in h2. - evar (t : Tm (S n)). - have : join (App (ren_Tm shift (Abs (Compile.F B1))) (VarTm var_zero)) t by - apply Join.FromPar; apply Par.AppAbs; auto using Par.refl. - subst t. rewrite -/ren_Tm. - move : h2. move /join_transitive => /[apply]. asimpl => h2. - tauto. - Qed. - - Lemma UnivInj n i j : R (Univ i : Tm n) (Univ j) -> i = j. - Proof. hauto l:on use:join_univ_inj. Qed. - - Lemma transitive n (a b c : Tm n) : - R a b -> R b c -> R a c. - Proof. hauto l:on use:join_transitive unfold:R. Qed. - - Lemma symmetric n (a b : Tm n) : - R a b -> R b a. - Proof. hauto l:on use:join_symmetric. Qed. - - Lemma reflexive n (a : Tm n) : - R a a. - Proof. hauto l:on use:join_refl. Qed. - - Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) : - R a b -> R (subst_Tm ρ a) (subst_Tm ρ b). - Proof. - rewrite /R. - rewrite /join. - move => [C [h0 h1]]. - repeat (rewrite <- Compile.morphing with (ρ0 := funcomp Compile.F ρ); last by reflexivity). - hauto lq:on use:join_substing. - Qed. - -End Join. - -Module Equiv. - Inductive R {n} : Tm n -> Tm n -> Prop := - (***************** Beta ***********************) - | AppAbs a b : - R (App (Abs a) b) (subst_Tm (scons b VarTm) a) - | ProjPair p a b : - R (Proj p (Pair a b)) (if p is PL then a else b) - - (****************** Eta ***********************) - | AppEta a : - R a (Abs (App (ren_Tm shift a) (VarTm var_zero))) - | PairEta a : - R a (Pair (Proj PL a) (Proj PR a)) - - (*************** Congruence ********************) - | Var i : R (VarTm i) (VarTm i) - | AbsCong a b : - R a b -> - R (Abs a) (Abs b) - | AppCong a0 a1 b0 b1 : - R a0 a1 -> - R b0 b1 -> - R (App a0 b0) (App a1 b1) - | PairCong a0 a1 b0 b1 : - R a0 a1 -> - R b0 b1 -> - R (Pair a0 b0) (Pair a1 b1) - | ProjCong p a0 a1 : - R a0 a1 -> - R (Proj p a0) (Proj p a1) - | BindCong p A0 A1 B0 B1: - R A0 A1 -> - R B0 B1 -> - R (TBind p A0 B0) (TBind p A1 B1) - | UnivCong i : - R (Univ i) (Univ i). -End Equiv. - -Module EquivJoin. - Lemma FromEquiv n (a b : Tm n) : Equiv.R a b -> Join.R a b. - Proof. - move => h. elim : n a b /h => n. - - move => a b. - rewrite /Join.R /join /=. - eexists. split. apply relations.rtc_once. - apply Par.AppAbs; auto using Par.refl. - rewrite Compile.substing. - apply relations.rtc_refl. - - move => p a b. - apply Join.FromPar. - simpl. apply : Par.ProjPair'; auto using Par.refl. - case : p => //=. - - move => a. apply Join.FromPar => /=. - apply : Par.AppEta'; auto using Par.refl. - by autorewrite with compile. - - move => a. apply Join.FromPar => /=. - apply : Par.PairEta; auto using Par.refl. - - hauto l:on use:Join.FromPar, Par.Var. - - hauto lq:on use:Join.AbsCong. - - qauto l:on use:Join.AppCong. - - qauto l:on use:Join.PairCong. - - qauto use:Join.ProjCong. - - rewrite /Join.R => p A0 A1 B0 B1 _ hA _ hB /=. - sfirstorder use:Join.PairCong,Join.AbsCong,Join.FromPar,Par.ConstCong. - - hauto l:on. - Qed. -End EquivJoin. - -Lemma compile_rpar n (a b : Tm n) : RPar'.R a b -> RPar'.R (Compile.F a) (Compile.F b). -Proof. - move => h. elim : n a b /h. - - move => n a0 a1 b0 b1 ha iha hb ihb /=. - rewrite -Compile.substing. - apply RPar'.AppAbs => //. - - hauto q:on use:RPar'.ProjPair'. - - qauto ctrs:RPar'.R. - - hauto lq:on ctrs:RPar'.R. - - hauto lq:on ctrs:RPar'.R. - - hauto lq:on ctrs:RPar'.R. - - hauto lq:on ctrs:RPar'.R. - - hauto lq:on ctrs:RPar'.R. - - hauto lq:on ctrs:RPar'.R. - - hauto lq:on ctrs:RPar'.R. - - hauto lq:on ctrs:RPar'.R. -Qed. - -Lemma compile_rpars n (a b : Tm n) : rtc RPar'.R a b -> rtc RPar'.R (Compile.F a) (Compile.F b). -Proof. induction 1; hauto lq:on ctrs:rtc use:compile_rpar. Qed. diff --git a/theories/diagram.txt b/theories/diagram.txt deleted file mode 100644 index ab47f3d..0000000 --- a/theories/diagram.txt +++ /dev/null @@ -1,42 +0,0 @@ -a0 --> a3 -| | -| | -v v -a1 --> a4 - - -(Abs a3) >>* a2 - -Abs a0 ----> (Abs a3) >>* a2 - | | - | | -Abs a1 ----> (Abs a4) >>* - - -a0 >> a1 -| | -| | -v v -b0 >> b1 - - -prov x (x, x) - -prov x b - - -a => b - -prov x a - -prov y b - -prov x c -prov y c - -extract c = x -extract c = y - -prov x b - -pr diff --git a/theories/fp_red.v b/theories/fp_red.v index f9abc08..f71a9d0 100644 --- a/theories/fp_red.v +++ b/theories/fp_red.v @@ -1,1229 +1,301 @@ -From Ltac2 Require Ltac2. -Import Ltac2.Notations. -Import Ltac2.Control. -Require Import ssreflect ssrbool. -Require Import FunInd. -Require Import Arith.Wf_nat. -Require Import Psatz. -From stdpp Require Import relations (rtc (..), rtc_once, rtc_r). +Require Import ssreflect. +From stdpp Require Import relations (rtc (..)). From Hammer Require Import Tactics. -Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax. - -Ltac2 spec_refl () := - List.iter - (fun a => match a with - | (i, _, _) => - let h := Control.hyp i in - try (specialize $h with (1 := eq_refl)) - end) (Control.hyps ()). - -Ltac spec_refl := ltac2:(spec_refl ()). +Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax. (* Trying my best to not write C style module_funcname *) Module Par. - Inductive R : PTm -> PTm -> Prop := + Inductive R {n} : Tm n -> Tm n -> Prop := (***************** Beta ***********************) | AppAbs a0 a1 b0 b1 : R a0 a1 -> R b0 b1 -> - R (PApp (PAbs a0) b0) (subst_PTm (scons b1 VarPTm) a1) + R (App (Abs a0) b0) (subst_Tm (scons b1 VarTm) a1) | AppPair a0 a1 b0 b1 c0 c1: R a0 a1 -> R b0 b1 -> R c0 c1 -> - R (PApp (PPair a0 b0) c0) (PPair (PApp a1 c1) (PApp b1 c1)) - | ProjAbs p a0 a1 : + R (App (Pair a0 b0) c0) (Pair (App a1 c1) (App b1 c1)) + | Proj1Abs a0 a1 : R a0 a1 -> - R (PProj p (PAbs a0)) (PAbs (PProj p a1)) - | ProjPair p a0 a1 b0 b1 : + R (Proj1 (Abs a0)) (Abs (Proj1 a0)) + | Proj1Pair a0 a1 b : R a0 a1 -> - R b0 b1 -> - R (PProj p (PPair a0 b0)) (if p is PL then a1 else b1) + R (Proj1 (Pair a0 b)) a1 + | Proj2Abs a0 a1 : + R a0 a1 -> + R (Proj2 (Abs a0)) (Abs (Proj2 a0)) + | Proj2Pair a0 a1 b : + R a0 a1 -> + R (Proj2 (Pair a0 b)) a1 (****************** Eta ***********************) | AppEta a0 a1 : R a0 a1 -> - R a0 (PAbs (PApp (ren_PTm shift a1) (VarPTm var_zero))) + R a0 (Abs (App (ren_Tm shift a1) (VarTm var_zero))) | PairEta a0 a1 : R a0 a1 -> - R a0 (PPair (PProj PL a1) (PProj PR a1)) + R a0 (Pair (Proj1 a1) (Proj2 a1)) (*************** Congruence ********************) - | Var i : R (VarPTm i) (VarPTm i) + | Var i : R (VarTm i) (VarTm i) | AbsCong a0 a1 : R a0 a1 -> - R (PAbs a0) (PAbs a1) + R (Abs a0) (Abs a1) | AppCong a0 a1 b0 b1 : R a0 a1 -> R b0 b1 -> - R (PApp a0 b0) (PApp a1 b1) + R (App a0 b0) (App a1 b1) | PairCong a0 a1 b0 b1 : R a0 a1 -> R b0 b1 -> - R (PPair a0 b0) (PPair a1 b1) - | ProjCong p a0 a1 : + R (Pair a0 b0) (Pair a1 b1) + | Proj1Cong a0 a1 : R a0 a1 -> - R (PProj p a0) (PProj p a1) - | ConstCong k : - R (PConst k) (PConst k) - | Univ i : - R (PUniv i) (PUniv i) - | Bot : - R PBot PBot. - - Lemma refl (a : PTm) : R a a. - elim : a; hauto ctrs:R. - Qed. - - Lemma AppAbs' a0 a1 (b0 b1 t : PTm) : - t = subst_PTm (scons b1 VarPTm) a1 -> + R (Proj1 a0) (Proj1 a1) + | Proj2Cong a0 a1 : R a0 a1 -> - R b0 b1 -> - R (PApp (PAbs a0) b0) t. - Proof. move => ->. apply AppAbs. Qed. - - Lemma ProjPair' p (a0 a1 b0 b1 : PTm) t : - t = (if p is PL then a1 else b1) -> - R a0 a1 -> - R b0 b1 -> - R (PProj p (PPair a0 b0)) t. - Proof. move => > ->. apply ProjPair. Qed. - - Lemma AppEta' (a0 a1 b : PTm) : - b = (PAbs (PApp (ren_PTm shift a1) (VarPTm var_zero))) -> - R a0 a1 -> - R a0 b. - Proof. move => ->; apply AppEta. Qed. - - Lemma renaming (a b : PTm) (ξ : nat -> nat) : - R a b -> R (ren_PTm ξ a) (ren_PTm ξ b). - Proof. - move => h. move : ξ. - elim : a b /h. - move => *; apply : AppAbs'; eauto; by asimpl. - all : match goal with - | [ |- context[var_zero]] => move => *; apply : AppEta'; eauto; by asimpl - | _ => qauto ctrs:R use:ProjPair' - end. - Qed. - - - Lemma morphing (a b : PTm) (ρ0 ρ1 : nat -> PTm) : - (forall i, R (ρ0 i) (ρ1 i)) -> - R a b -> R (subst_PTm ρ0 a) (subst_PTm ρ1 b). - Proof. - move => + h. move : ρ0 ρ1. elim : a b/h. - - move => a0 a1 b0 b1 ha iha hb ihb ρ0 ρ1 hρ /=. - eapply AppAbs' with (a1 := subst_PTm (up_PTm_PTm ρ1) a1); eauto. - by asimpl. - hauto l:on use:renaming inv:nat. - - hauto lq:on rew:off ctrs:R. - - hauto l:on inv:nat use:renaming ctrs:R. - - hauto lq:on use:ProjPair'. - - move => a0 a1 ha iha ρ0 ρ1 hρ /=. - apply : AppEta'; eauto. by asimpl. - - hauto lq:on ctrs:R. - - sfirstorder. - - hauto l:on inv:nat ctrs:R use:renaming. - - hauto q:on ctrs:R. - - qauto l:on ctrs:R. - - qauto l:on ctrs:R. - - hauto l:on inv:option ctrs:R use:renaming. - - qauto l:on ctrs:R. - - qauto l:on ctrs:R. - Qed. - - Lemma substing (a b : PTm) (ρ : nat -> PTm) : - R a b -> R (subst_PTm ρ a) (subst_PTm ρ b). - Proof. hauto l:on use:morphing, refl. Qed. - - Lemma antirenaming (a : PTm) (b : PTm) (ξ : nat -> nat) : - R (ren_PTm ξ a) b -> exists b0, R a b0 /\ ren_PTm ξ b0 = b. - Proof. - move E : (ren_PTm ξ a) => u h. - move : ξ a E. elim : u b/h. - - move => a0 a1 b0 b1 ha iha hb ihb ξ []//=. - move => c c0 [+ ?]. subst. - case : c => //=. - move => c [?]. subst. - spec_refl. - move : iha => [c1][ih0]?. subst. - move : ihb => [c2][ih1]?. subst. - eexists. split. - apply AppAbs; eauto. - by asimpl. - - move => a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc ξ []//=. - move => []//= t t0 t1 [*]. subst. - spec_refl. - move : iha => [? [*]]. - move : ihb => [? [*]]. - move : ihc => [? [*]]. - eexists. split. - apply AppPair; hauto. subst. - by asimpl. - - move => p a0 a1 ha iha ξ []//= p0 []//= t [*]. subst. - spec_refl. move : iha => [b0 [? ?]]. subst. - eexists. split. apply ProjAbs; eauto. by asimpl. - - move => p a0 a1 b0 b1 ha iha hb ihb ξ []//= p0 []//= t t0[*]. - subst. spec_refl. - move : iha => [b0 [? ?]]. - move : ihb => [c0 [? ?]]. subst. - eexists. split. by eauto using ProjPair. - hauto q:on. - - move => a0 a1 ha iha ξ a ?. subst. - spec_refl. move : iha => [a0 [? ?]]. subst. - eexists. split. apply AppEta; eauto. - by asimpl. - - move => a0 a1 ha iha ξ a ?. subst. - spec_refl. move : iha => [b0 [? ?]]. subst. - eexists. split. apply PairEta; eauto. - by asimpl. - - move => i ξ []//=. - hauto l:on. - - move => a0 a1 ha iha ξ []//= t [*]. subst. - spec_refl. - move :iha => [b0 [? ?]]. subst. - eexists. split. by apply AbsCong; eauto. - done. - - move => a0 a1 b0 b1 ha iha hb ihb ξ []//= t t0 [*]. subst. - spec_refl. - move : iha => [b0 [? ?]]. subst. - move : ihb => [c0 [? ?]]. subst. - eexists. split. by apply AppCong; eauto. - done. - - move => a0 a1 b0 b1 ha iha hb ihb ξ []//= t t0[*]. subst. - spec_refl. - move : iha => [b0 [? ?]]. subst. - move : ihb => [c0 [? ?]]. subst. - eexists. split=>/=. by apply PairCong; eauto. - done. - - move => p a0 a1 ha iha ξ []//= p0 t [*]. subst. - spec_refl. - move : iha => [b0 [? ?]]. subst. - eexists. split. by apply ProjCong; eauto. - done. - - hauto q:on inv:PTm ctrs:R. - - hauto q:on inv:PTm ctrs:R. - - hauto q:on inv:PTm ctrs:R. - Qed. - + R (Proj2 a0) (Proj2 a1). End Par. -Module Pars. - Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) : - rtc Par.R a b -> rtc Par.R (ren_PTm ξ a) (ren_PTm ξ b). - Proof. - induction 1; hauto lq:on ctrs:rtc use:Par.renaming. - Qed. - - Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) : - rtc Par.R a b -> - rtc Par.R (subst_PTm ρ a) (subst_PTm ρ b). - induction 1; hauto l:on ctrs:rtc use:Par.substing. - Qed. - - Lemma antirenaming n m (a : PTm n) (b : PTm m) (ξ : fin n -> fin m) : - rtc Par.R (ren_PTm ξ a) b -> exists b0, rtc Par.R a b0 /\ ren_PTm ξ b0 = b. - Proof. - move E :(ren_PTm ξ a) => u h. - move : a E. - elim : u b /h. - - sfirstorder. - - move => a b c h0 h1 ih1 a0 ?. subst. - move /Par.antirenaming : h0. - move => [b0 [h2 ?]]. subst. - hauto lq:on rew:off ctrs:rtc. - Qed. - - #[local]Ltac solve_s_rec := - move => *; eapply rtc_l; eauto; - hauto lq:on ctrs:Par.R use:Par.refl. - - #[local]Ltac solve_s := - repeat (induction 1; last by solve_s_rec); apply rtc_refl. - - Lemma ProjCong n p (a0 a1 : PTm n) : - rtc Par.R a0 a1 -> - rtc Par.R (PProj p a0) (PProj p a1). - Proof. solve_s. Qed. - - Lemma PairCong n (a0 a1 b0 b1 : PTm n) : - rtc Par.R a0 a1 -> - rtc Par.R b0 b1 -> - rtc Par.R (PPair a0 b0) (PPair a1 b1). - Proof. solve_s. Qed. - - Lemma AppCong n (a0 a1 b0 b1 : PTm n) : - rtc Par.R a0 a1 -> - rtc Par.R b0 b1 -> - rtc Par.R (PApp a0 b0) (PApp a1 b1). - Proof. solve_s. Qed. - - Lemma AbsCong n (a b : PTm (S n)) : - rtc Par.R a b -> - rtc Par.R (PAbs a) (PAbs b). - Proof. solve_s. Qed. - -End Pars. - -Definition var_or_const {n} (a : PTm n) := - match a with - | VarPTm _ => true - | PBot => true - | _ => false - end. - - (***************** Beta rules only ***********************) Module RPar. - Inductive R {n} : PTm n -> PTm n -> Prop := + Inductive R {n} : Tm n -> Tm n -> Prop := (***************** Beta ***********************) | AppAbs a0 a1 b0 b1 : R a0 a1 -> R b0 b1 -> - R (PApp (PAbs a0) b0) (subst_PTm (scons b1 VarPTm) a1) + R (App (Abs a0) b0) (subst_Tm (scons b1 VarTm) a1) | AppPair a0 a1 b0 b1 c0 c1: R a0 a1 -> R b0 b1 -> R c0 c1 -> - R (PApp (PPair a0 b0) c0) (PPair (PApp a1 c1) (PApp b1 c1)) - | ProjAbs p a0 a1 : + R (App (Pair a0 b0) c0) (Pair (App a1 c1) (App b1 c1)) + | Proj1Abs a0 a1 : R a0 a1 -> - R (PProj p (PAbs a0)) (PAbs (PProj p a1)) - | ProjPair p a0 a1 b0 b1 : + R (Proj1 (Abs a0)) (Abs (Proj1 a0)) + | Proj1Pair a0 a1 b : R a0 a1 -> - R b0 b1 -> - R (PProj p (PPair a0 b0)) (if p is PL then a1 else b1) - + R (Proj1 (Pair a0 b)) a1 + | Proj2Abs a0 a1 : + R a0 a1 -> + R (Proj2 (Abs a0)) (Abs (Proj2 a0)) + | Proj2Pair a0 a1 b : + R a0 a1 -> + R (Proj2 (Pair a0 b)) a1 (*************** Congruence ********************) - | Var i : R (VarPTm i) (VarPTm i) + | Var i : R (VarTm i) (VarTm i) | AbsCong a0 a1 : R a0 a1 -> - R (PAbs a0) (PAbs a1) + R (Abs a0) (Abs a1) | AppCong a0 a1 b0 b1 : R a0 a1 -> R b0 b1 -> - R (PApp a0 b0) (PApp a1 b1) + R (App a0 b0) (App a1 b1) | PairCong a0 a1 b0 b1 : R a0 a1 -> R b0 b1 -> - R (PPair a0 b0) (PPair a1 b1) - | ProjCong p a0 a1 : + R (Pair a0 b0) (Pair a1 b1) + | Proj1Cong a0 a1 : R a0 a1 -> - R (PProj p a0) (PProj p a1) - | ConstCong k : - R (PConst k) (PConst k) - | Univ i : - R (PUniv i) (PUniv i) - | Bot : - R PBot PBot. + R (Proj1 a0) (Proj1 a1) + | Proj2Cong a0 a1 : + R a0 a1 -> + R (Proj2 a0) (Proj2 a1). - Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop. + Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop. - Lemma refl n (a : PTm n) : R a a. + Lemma refl n (a : Tm n) : R a a. Proof. induction a; hauto lq:on ctrs:R. Qed. - Lemma AppAbs' n a0 a1 (b0 b1 t : PTm n) : - t = subst_PTm (scons b1 VarPTm) a1 -> + Lemma AppAbs' n a0 a1 (b0 b1 t : Tm n) : + t = subst_Tm (scons b1 VarTm) a1 -> R a0 a1 -> R b0 b1 -> - R (PApp (PAbs a0) b0) t. + R (App (Abs a0) b0) t. Proof. move => ->. apply AppAbs. Qed. - Lemma ProjPair' n p (a0 a1 b0 b1 : PTm n) t : - t = (if p is PL then a1 else b1) -> - R a0 a1 -> - R b0 b1 -> - R (PProj p (PPair a0 b0)) t. - Proof. move => > ->. apply ProjPair. Qed. - - Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) : - R a b -> R (ren_PTm ξ a) (ren_PTm ξ b). + Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) : + R a b -> R (ren_Tm ξ a) (ren_Tm ξ b). Proof. move => h. move : m ξ. elim : n a b /h. move => *; apply : AppAbs'; eauto; by asimpl. - all : qauto ctrs:R use:ProjPair'. - Qed. - - Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> PTm m) (ξ : fin m -> fin p) : - (forall i, R (ρ0 i) (ρ1 i)) -> - (forall i, R ((funcomp (ren_PTm ξ) ρ0) i) ((funcomp (ren_PTm ξ) ρ1) i)). - Proof. eauto using renaming. Qed. - - Lemma morphing_ext n m (ρ0 ρ1 : fin n -> PTm m) a b : - R a b -> - (forall i, R (ρ0 i) (ρ1 i)) -> - (forall i, R ((scons a ρ0) i) ((scons b ρ1) i)). - Proof. hauto q:on inv:nat. Qed. - - Lemma morphing_up n m (ρ0 ρ1 : fin n -> PTm m) : - (forall i, R (ρ0 i) (ρ1 i)) -> - (forall i, R (up_PTm_PTm ρ0 i) (up_PTm_PTm ρ1 i)). - Proof. hauto l:on ctrs:R use:morphing_ext, morphing_ren unfold:up_PTm_PTm. Qed. - - Lemma morphing n m (a b : PTm n) (ρ0 ρ1 : fin n -> PTm m) : - (forall i, R (ρ0 i) (ρ1 i)) -> - R a b -> R (subst_PTm ρ0 a) (subst_PTm ρ1 b). - Proof. - move => + h. move : m ρ0 ρ1. - elim : n a b /h. - - move => *. - apply : AppAbs'; eauto using morphing_up. - by asimpl. - - hauto lq:on ctrs:R. - - hauto lq:on ctrs:R use:morphing_up. - - hauto lq:on ctrs:R use:ProjPair' use:morphing_up. - - hauto lq:on ctrs:R use:morphing_up. - - hauto lq:on ctrs:R use:morphing_up. - - hauto lq:on ctrs:R use:morphing_up. - - hauto lq:on ctrs:R. - - hauto lq:on ctrs:R. - - hauto lq:on ctrs:R use:morphing_up. - - hauto lq:on ctrs:R. - - hauto lq:on ctrs:R. - Qed. - - Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) : - R a b -> - R (subst_PTm ρ a) (subst_PTm ρ b). - Proof. hauto l:on use:morphing, refl. Qed. - - Lemma cong n (a b : PTm (S n)) c d : - R a b -> - R c d -> - R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons d VarPTm) b). - Proof. - move => h0 h1. apply morphing => //=. - qauto l:on ctrs:R inv:nat. - Qed. - - Lemma var_or_const_imp {n} (a b : PTm n) : - var_or_const a -> - a = b -> ~~ var_or_const b -> False. - Proof. - hauto lq:on inv:PTm. - Qed. - - Lemma var_or_const_up n m (ρ : fin n -> PTm m) : - (forall i, var_or_const (ρ i)) -> - (forall i, var_or_const (up_PTm_PTm ρ i)). - Proof. - move => h /= [i|]. - - asimpl. - move /(_ i) in h. - rewrite /funcomp. - move : (ρ i) h. - case => //=. - - sfirstorder. - Qed. - - Local Ltac antiimp := qauto l:on use:var_or_const_imp. - - Lemma antirenaming n m (a : PTm n) (b : PTm m) (ρ : fin n -> PTm m) : - (forall i, var_or_const (ρ i)) -> - R (subst_PTm ρ a) b -> exists b0, R a b0 /\ subst_PTm ρ b0 = b. - Proof. - move E : (subst_PTm ρ a) => u hρ h. - move : n ρ hρ a E. elim : m u b/h. - - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; - first by antiimp. - move => c c0 [+ ?]. subst. - case : c => //=; first by antiimp. - move => c [?]. subst. - spec_refl. - have /var_or_const_up hρ' := hρ. - move : iha hρ' => /[apply] iha. - move : ihb hρ => /[apply] ihb. - spec_refl. - move : iha => [c1][ih0]?. subst. - move : ihb => [c2][ih1]?. subst. - eexists. split. - apply AppAbs; eauto. - by asimpl. - - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc m ρ hρ. - move => []//=; - first by antiimp. - move => []//=; first by antiimp. - move => t t0 t1 [*]. subst. - have {}/iha := hρ => iha. - have {}/ihb := hρ => ihb. - have {}/ihc := hρ => ihc. - spec_refl. - move : iha => [? [*]]. - move : ihb => [? [*]]. - move : ihc => [? [*]]. - eexists. split. - apply AppPair; hauto. subst. - by asimpl. - - move => n p a0 a1 ha iha m ρ hρ []//=; - first by antiimp. - move => p0 []//= t [*]; first by antiimp. subst. - have /var_or_const_up {}/iha := hρ => iha. - spec_refl. move : iha => [b0 [? ?]]. subst. - eexists. split. apply ProjAbs; eauto. by asimpl. - - move => n p a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; - first by antiimp. - move => p0 []//=; first by antiimp. move => t t0[*]. - subst. - have {}/iha := (hρ) => iha. - have {}/ihb := (hρ) => ihb. - spec_refl. - move : iha => [b0 [? ?]]. - move : ihb => [c0 [? ?]]. subst. - eexists. split. by eauto using ProjPair. - hauto q:on. - - move => n i m ρ hρ []//=. - hauto l:on. - - move => n a0 a1 ha iha m ρ hρ []//=; first by antiimp. - move => t [*]. subst. - have /var_or_const_up {}/iha := hρ => iha. - spec_refl. - move :iha => [b0 [? ?]]. subst. - eexists. split. by apply AbsCong; eauto. - by asimpl. - - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; - first by antiimp. - move => t t0 [*]. subst. - have {}/iha := (hρ) => iha. - have {}/ihb := (hρ) => ihb. - spec_refl. - move : iha => [b0 [? ?]]. subst. - move : ihb => [c0 [? ?]]. subst. - eexists. split. by apply AppCong; eauto. - done. - - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; - first by antiimp. - move => t t0[*]. subst. - have {}/iha := (hρ) => iha. - have {}/ihb := (hρ) => ihb. - spec_refl. - move : iha => [b0 [? ?]]. subst. - move : ihb => [c0 [? ?]]. subst. - eexists. split. by apply PairCong; eauto. - by asimpl. - - move => n p a0 a1 ha iha m ρ hρ []//=; - first by antiimp. - move => p0 t [*]. subst. - have {}/iha := (hρ) => iha. - spec_refl. - move : iha => [b0 [? ?]]. subst. - eexists. split. apply ProjCong; eauto. reflexivity. - - hauto q:on ctrs:R inv:PTm. - - hauto q:on ctrs:R inv:PTm. - - hauto q:on ctrs:R inv:PTm. - Qed. -End RPar. - -(***************** Beta rules only ***********************) -Module RPar'. - Inductive R {n} : PTm n -> PTm n -> Prop := - (***************** Beta ***********************) - | AppAbs a0 a1 b0 b1 : - R a0 a1 -> - R b0 b1 -> - R (PApp (PAbs a0) b0) (subst_PTm (scons b1 VarPTm) a1) - | ProjPair p a0 a1 b0 b1 : - R a0 a1 -> - R b0 b1 -> - R (PProj p (PPair a0 b0)) (if p is PL then a1 else b1) - - - (*************** Congruence ********************) - | Var i : R (VarPTm i) (VarPTm i) - | AbsCong a0 a1 : - R a0 a1 -> - R (PAbs a0) (PAbs a1) - | AppCong a0 a1 b0 b1 : - R a0 a1 -> - R b0 b1 -> - R (PApp a0 b0) (PApp a1 b1) - | PairCong a0 a1 b0 b1 : - R a0 a1 -> - R b0 b1 -> - R (PPair a0 b0) (PPair a1 b1) - | ProjCong p a0 a1 : - R a0 a1 -> - R (PProj p a0) (PProj p a1) - | ConstCong k : - R (PConst k) (PConst k) - | UnivCong i : - R (PUniv i) (PUniv i) - | BotCong : - R PBot PBot. - - Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop. - - Lemma refl n (a : PTm n) : R a a. - Proof. - induction a; hauto lq:on ctrs:R. - Qed. - - Lemma AppAbs' n a0 a1 (b0 b1 t : PTm n) : - t = subst_PTm (scons b1 VarPTm) a1 -> - R a0 a1 -> - R b0 b1 -> - R (PApp (PAbs a0) b0) t. - Proof. move => ->. apply AppAbs. Qed. - - Lemma ProjPair' n p (a0 a1 b0 b1 : PTm n) t : - t = (if p is PL then a1 else b1) -> - R a0 a1 -> - R b0 b1 -> - R (PProj p (PPair a0 b0)) t. - Proof. move => > ->. apply ProjPair. Qed. - - Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) : - R a b -> R (ren_PTm ξ a) (ren_PTm ξ b). - Proof. - move => h. move : m ξ. - elim : n a b /h. - move => *; apply : AppAbs'; eauto; by asimpl. - all : qauto ctrs:R use:ProjPair'. - Qed. - - Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> PTm m) (ξ : fin m -> fin p) : - (forall i, R (ρ0 i) (ρ1 i)) -> - (forall i, R ((funcomp (ren_PTm ξ) ρ0) i) ((funcomp (ren_PTm ξ) ρ1) i)). - Proof. eauto using renaming. Qed. - - Lemma morphing_ext n m (ρ0 ρ1 : fin n -> PTm m) a b : - R a b -> - (forall i, R (ρ0 i) (ρ1 i)) -> - (forall i, R ((scons a ρ0) i) ((scons b ρ1) i)). - Proof. hauto q:on inv:nat. Qed. - - Lemma morphing_up n m (ρ0 ρ1 : fin n -> PTm m) : - (forall i, R (ρ0 i) (ρ1 i)) -> - (forall i, R (up_PTm_PTm ρ0 i) (up_PTm_PTm ρ1 i)). - Proof. hauto l:on ctrs:R use:morphing_ext, morphing_ren unfold:up_PTm_PTm. Qed. - - Lemma morphing n m (a b : PTm n) (ρ0 ρ1 : fin n -> PTm m) : - (forall i, R (ρ0 i) (ρ1 i)) -> - R a b -> R (subst_PTm ρ0 a) (subst_PTm ρ1 b). - Proof. - move => + h. move : m ρ0 ρ1. - elim : n a b /h. - - move => *. - apply : AppAbs'; eauto using morphing_up. - by asimpl. - - hauto lq:on ctrs:R use:ProjPair' use:morphing_up. - - hauto lq:on ctrs:R use:morphing_up. - - hauto lq:on ctrs:R use:morphing_up. - - hauto lq:on ctrs:R use:morphing_up. - - hauto lq:on ctrs:R. - - hauto lq:on ctrs:R. - - hauto l:on ctrs:R use:morphing_up. - - hauto lq:on ctrs:R. - - hauto lq:on ctrs:R. - Qed. - - Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) : - R a b -> - R (subst_PTm ρ a) (subst_PTm ρ b). - Proof. hauto l:on use:morphing, refl. Qed. - - Lemma cong n (a b : PTm (S n)) c d : - R a b -> - R c d -> - R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons d VarPTm) b). - Proof. - move => h0 h1. apply morphing => //=. - qauto l:on ctrs:R inv:nat. - Qed. - - Lemma var_or_const_imp {n} (a b : PTm n) : - var_or_const a -> - a = b -> ~~ var_or_const b -> False. - Proof. - hauto lq:on inv:PTm. - Qed. - - Lemma var_or_const_up n m (ρ : fin n -> PTm m) : - (forall i, var_or_const (ρ i)) -> - (forall i, var_or_const (up_PTm_PTm ρ i)). - Proof. - move => h /= [i|]. - - asimpl. - move /(_ i) in h. - rewrite /funcomp. - move : (ρ i) h. - case => //=. - - sfirstorder. - Qed. - - Local Ltac antiimp := qauto l:on use:var_or_const_imp. - - Lemma antirenaming n m (a : PTm n) (b : PTm m) (ρ : fin n -> PTm m) : - (forall i, var_or_const (ρ i)) -> - R (subst_PTm ρ a) b -> exists b0, R a b0 /\ subst_PTm ρ b0 = b. - Proof. - move E : (subst_PTm ρ a) => u hρ h. - move : n ρ hρ a E. elim : m u b/h. - - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; - first by antiimp. - move => c c0 [+ ?]. subst. - case : c => //=; first by antiimp. - move => c [?]. subst. - spec_refl. - have /var_or_const_up hρ' := hρ. - move : iha hρ' => /[apply] iha. - move : ihb hρ => /[apply] ihb. - spec_refl. - move : iha => [c1][ih0]?. subst. - move : ihb => [c2][ih1]?. subst. - eexists. split. - apply AppAbs; eauto. - by asimpl. - - move => n p a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; - first by antiimp. - move => p0 []//=; first by antiimp. move => t t0[*]. - subst. - have {}/iha := (hρ) => iha. - have {}/ihb := (hρ) => ihb. - spec_refl. - move : iha => [b0 [? ?]]. - move : ihb => [c0 [? ?]]. subst. - eexists. split. by eauto using ProjPair. - hauto q:on. - - move => n i m ρ hρ []//=. - hauto l:on. - - move => n a0 a1 ha iha m ρ hρ []//=; first by antiimp. - move => t [*]. subst. - have /var_or_const_up {}/iha := hρ => iha. - spec_refl. - move :iha => [b0 [? ?]]. subst. - eexists. split. by apply AbsCong; eauto. - by asimpl. - - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; - first by antiimp. - move => t t0 [*]. subst. - have {}/iha := (hρ) => iha. - have {}/ihb := (hρ) => ihb. - spec_refl. - move : iha => [b0 [? ?]]. subst. - move : ihb => [c0 [? ?]]. subst. - eexists. split. by apply AppCong; eauto. - done. - - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; - first by antiimp. - move => t t0[*]. subst. - have {}/iha := (hρ) => iha. - have {}/ihb := (hρ) => ihb. - spec_refl. - move : iha => [b0 [? ?]]. subst. - move : ihb => [c0 [? ?]]. subst. - eexists. split. by apply PairCong; eauto. - by asimpl. - - move => n p a0 a1 ha iha m ρ hρ []//=; - first by antiimp. - move => p0 t [*]. subst. - have {}/iha := (hρ) => iha. - spec_refl. - move : iha => [b0 [? ?]]. subst. - eexists. split. apply ProjCong; eauto. reflexivity. - - hauto q:on ctrs:R inv:PTm. - - move => n i n0 ρ hρ []//=; first by antiimp. - hauto l:on. - - hauto q:on inv:PTm ctrs:R. - Qed. -End RPar'. - -Module ERed. - Inductive R {n} : PTm n -> PTm n -> Prop := - (****************** Eta ***********************) - | AppEta a : - R a (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero))) - | PairEta a : - R a (PPair (PProj PL a) (PProj PR a)) - - (*************** Congruence ********************) - | AbsCong a0 a1 : - R a0 a1 -> - R (PAbs a0) (PAbs a1) - | AppCong0 a0 a1 b : - R a0 a1 -> - R (PApp a0 b) (PApp a1 b) - | AppCong1 a b0 b1 : - R b0 b1 -> - R (PApp a b0) (PApp a b1) - | PairCong0 a0 a1 b : - R a0 a1 -> - R (PPair a0 b) (PPair a1 b) - | PairCong1 a b0 b1 : - R b0 b1 -> - R (PPair a b0) (PPair a b1) - | ProjCong p a0 a1 : - R a0 a1 -> - R (PProj p a0) (PProj p a1). - - Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop. - - Lemma AppEta' n a (u : PTm n) : - u = (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero))) -> - R a u. - Proof. move => ->. apply AppEta. Qed. - - Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) : - R a b -> R (ren_PTm ξ a) (ren_PTm ξ b). - Proof. - move => h. move : m ξ. - elim : n a b /h. - - move => n a m ξ. - apply AppEta'. by asimpl. all : qauto ctrs:R. Qed. - Lemma substing n m (a : PTm n) b (ρ : fin n -> PTm m) : - R a b -> - R (subst_PTm ρ a) (subst_PTm ρ b). - Proof. - move => h. move : m ρ. elim : n a b / h => n. - move => a m ρ /=. - apply : AppEta'; eauto. by asimpl. - all : hauto ctrs:R inv:nat use:renaming. - Qed. - -End ERed. - -Module EReds. - - #[local]Ltac solve_s_rec := - move => *; eapply rtc_l; eauto; - hauto lq:on ctrs:ERed.R. - - #[local]Ltac solve_s := - repeat (induction 1; last by solve_s_rec); apply rtc_refl. - - Lemma AbsCong n (a b : PTm (S n)) : - rtc ERed.R a b -> - rtc ERed.R (PAbs a) (PAbs b). - Proof. solve_s. Qed. - - Lemma AppCong n (a0 a1 b0 b1 : PTm n) : - rtc ERed.R a0 a1 -> - rtc ERed.R b0 b1 -> - rtc ERed.R (PApp a0 b0) (PApp a1 b1). - Proof. solve_s. Qed. - - Lemma PairCong n (a0 a1 b0 b1 : PTm n) : - rtc ERed.R a0 a1 -> - rtc ERed.R b0 b1 -> - rtc ERed.R (PPair a0 b0) (PPair a1 b1). - Proof. solve_s. Qed. - - Lemma ProjCong n p (a0 a1 : PTm n) : - rtc ERed.R a0 a1 -> - rtc ERed.R (PProj p a0) (PProj p a1). - Proof. solve_s. Qed. -End EReds. +End RPar. Module EPar. - Inductive R {n} : PTm n -> PTm n -> Prop := + Inductive R {n} : Tm n -> Tm n -> Prop := (****************** Eta ***********************) - | AppEta a0 a1 : - R a0 a1 -> - R a0 (PAbs (PApp (ren_PTm shift a1) (VarPTm var_zero))) - | PairEta a0 a1 : - R a0 a1 -> - R a0 (PPair (PProj PL a1) (PProj PR a1)) + | AppEta a : + R a (Abs (App (ren_Tm shift a) (VarTm var_zero))) + | PairEta a : + R a (Pair (Proj1 a) (Proj2 a)) (*************** Congruence ********************) - | Var i : R (VarPTm i) (VarPTm i) + | Var i : R (VarTm i) (VarTm i) | AbsCong a0 a1 : R a0 a1 -> - R (PAbs a0) (PAbs a1) + R (Abs a0) (Abs a1) | AppCong a0 a1 b0 b1 : R a0 a1 -> R b0 b1 -> - R (PApp a0 b0) (PApp a1 b1) + R (App a0 b0) (App a1 b1) | PairCong a0 a1 b0 b1 : R a0 a1 -> R b0 b1 -> - R (PPair a0 b0) (PPair a1 b1) - | ProjCong p a0 a1 : + R (Pair a0 b0) (Pair a1 b1) + | Proj1Cong a0 a1 : R a0 a1 -> - R (PProj p a0) (PProj p a1) - | ConstCong k : - R (PConst k) (PConst k) - | UnivCong i : - R (PUniv i) (PUniv i) - | BotCong : - R PBot PBot. + R (Proj1 a0) (Proj1 a1) + | Proj2Cong a0 a1 : + R a0 a1 -> + R (Proj2 a0) (Proj2 a1). - Lemma refl n (a : PTm n) : EPar.R a a. + Lemma refl n (a : Tm n) : EPar.R a a. Proof. induction a; hauto lq:on ctrs:EPar.R. Qed. - Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) : - R a b -> R (ren_PTm ξ a) (ren_PTm ξ b). + Lemma AppEta' n (a b : Tm n) : + b = (Abs (App (ren_Tm shift a) (VarTm var_zero))) -> + R a b. + Proof. move => ->. by apply AppEta. Qed. + + Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) : + R a b -> R (ren_Tm ξ a) (ren_Tm ξ b). Proof. move => h. move : m ξ. elim : n a b /h. - - move => n a0 a1 ha iha m ξ /=. - move /(_ _ ξ) /AppEta : iha. - by asimpl. - + move => n a m ξ /=. + apply AppEta'. by asimpl. all : qauto ctrs:R. Qed. - Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop. + Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop. - Lemma AppEta' n (a0 a1 b : PTm n) : - b = (PAbs (PApp (ren_PTm shift a1) (VarPTm var_zero))) -> - R a0 a1 -> - R a0 b. - Proof. move => ->; apply AppEta. Qed. - - Lemma morphing n m (a b : PTm n) (ρ0 ρ1 : fin n -> PTm m) : - R a b -> + Lemma morph n m (a : Tm n) (ρ0 ρ1 : fin n -> Tm m) : (forall i, R (ρ0 i) (ρ1 i)) -> - R (subst_PTm ρ0 a) (subst_PTm ρ1 b). + R (subst_Tm ρ0 a) (subst_Tm ρ1 a). Proof. - move => h. move : m ρ0 ρ1. elim : n a b / h => n. - - move => a0 a1 ha iha m ρ0 ρ1 hρ /=. - apply : AppEta'; eauto. by asimpl. - - hauto lq:on ctrs:R. - - hauto lq:on ctrs:R. - - hauto l:on ctrs:R use:renaming inv:nat. - - hauto q:on ctrs:R. - - hauto q:on ctrs:R. - - hauto q:on ctrs:R. - - hauto l:on ctrs:R use:renaming inv:nat. - - hauto lq:on ctrs:R. - - hauto lq:on ctrs:R. - Qed. - - Lemma substing n a0 a1 (b0 b1 : PTm n) : - R a0 a1 -> - R b0 b1 -> - R (subst_PTm (scons b0 VarPTm) a0) (subst_PTm (scons b1 VarPTm) a1). - Proof. - move => h0 h1. apply morphing => //. - hauto lq:on ctrs:R inv:nat. - Qed. + move : m ρ0 ρ1. End EPar. -Module OExp. - Inductive R {n} : PTm n -> PTm n -> Prop := - (****************** Eta ***********************) - | AppEta a : - R a (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero))) - | PairEta a : - R a (PPair (PProj PL a) (PProj PR a)). - - Lemma merge n (t a b : PTm n) : - rtc R a b -> - EPar.R t a -> - EPar.R t b. - Proof. - move => h. move : t. elim : a b /h. - - eauto using EPar.refl. - - hauto q:on ctrs:EPar.R inv:R. - Qed. - - Lemma commutativity n (a b c : PTm n) : - EPar.R a b -> R a c -> exists d, R b d /\ EPar.R c d. - Proof. - move => h. - inversion 1; subst. - - hauto q:on ctrs:EPar.R, R use:EPar.renaming, EPar.refl. - - hauto lq:on ctrs:EPar.R, R. - Qed. - - Lemma commutativity0 n (a b c : PTm n) : - EPar.R a b -> rtc R a c -> exists d, rtc R b d /\ EPar.R c d. - Proof. - move => + h. move : b. - elim : a c / h. - - sfirstorder. - - hauto lq:on rew:off ctrs:rtc use:commutativity. - Qed. - -End OExp. - - Local Ltac com_helper := split; [hauto lq:on ctrs:RPar.R use: RPar.refl, RPar.renaming |hauto lq:on ctrs:EPar.R use:EPar.refl, EPar.renaming]. -Module RPars. +Lemma RPars_AbsCong n (a b : Tm (S n)) : + rtc RPar.R a b -> + rtc RPar.R (Abs a) (Abs b). +Proof. induction 1; hauto l:on ctrs:RPar.R, rtc. Qed. - #[local]Ltac solve_s_rec := - move => *; eapply rtc_l; eauto; - hauto lq:on ctrs:RPar.R use:RPar.refl. +Lemma RPars_AppCong n (a0 a1 b : Tm n) : + rtc RPar.R a0 a1 -> + rtc RPar.R (App a0 b) (App a1 b). +Proof. + move => h. move : b. + elim : a0 a1 /h. + - qauto ctrs:RPar.R, rtc. + - move => x y z h0 h1 ih b. + apply rtc_l with (y := App y b) => //. + hauto lq:on ctrs:RPar.R use:RPar.refl. +Qed. - #[local]Ltac solve_s := - repeat (induction 1; last by solve_s_rec); apply rtc_refl. +Lemma RPars_PairCong n (a0 a1 b0 b1 : Tm n) : + rtc RPar.R a0 a1 -> + rtc RPar.R b0 b1 -> + rtc RPar.R (Pair a0 b0) (Pair a1 b1). +Proof. + move => h. move : b0 b1. + elim : a0 a1 /h. + - move => x b0 b1 h. + elim : b0 b1 /h. + by auto using rtc_refl. + move => x0 y z h0 h1 h2. + apply : rtc_l; eauto. + by eauto using RPar.refl, RPar.PairCong. + - move => x y z h0 h1 ih b0 b1 h. + apply : rtc_l; eauto. + by eauto using RPar.refl, RPar.PairCong. +Qed. - Lemma AbsCong n (a b : PTm (S n)) : - rtc RPar.R a b -> - rtc RPar.R (PAbs a) (PAbs b). - Proof. solve_s. Qed. - Lemma AppCong n (a0 a1 b0 b1 : PTm n) : - rtc RPar.R a0 a1 -> - rtc RPar.R b0 b1 -> - rtc RPar.R (PApp a0 b0) (PApp a1 b1). - Proof. solve_s. Qed. +Lemma RPars_renaming n (a0 a1 : Tm n) m (ξ : fin n -> fin m) : + rtc RPar.R a0 a1 -> + rtc RPar.R (ren_Tm ξ a0) (ren_Tm ξ a1). +Proof. + induction 1. + - apply rtc_refl. + - eauto using RPar.renaming, rtc_l. +Qed. - Lemma PairCong n (a0 a1 b0 b1 : PTm n) : - rtc RPar.R a0 a1 -> - rtc RPar.R b0 b1 -> - rtc RPar.R (PPair a0 b0) (PPair a1 b1). - Proof. solve_s. Qed. +Lemma RPars_Abs_inv n (a : Tm (S n)) b : + rtc RPar.R (Abs a) b -> exists a', b = Abs a' /\ rtc RPar.R a a'. +Proof. + move E : (Abs a) => b0 h. move : a E. + elim : b0 b / h. + - hauto lq:on ctrs:rtc. + - hauto lq:on ctrs:rtc inv:RPar.R, rtc. +Qed. - Lemma ProjCong n p (a0 a1 : PTm n) : - rtc RPar.R a0 a1 -> - rtc RPar.R (PProj p a0) (PProj p a1). - Proof. solve_s. Qed. - - Lemma renaming n (a0 a1 : PTm n) m (ξ : fin n -> fin m) : - rtc RPar.R a0 a1 -> - rtc RPar.R (ren_PTm ξ a0) (ren_PTm ξ a1). - Proof. - induction 1. - - apply rtc_refl. - - eauto using RPar.renaming, rtc_l. - Qed. - - Lemma weakening n (a0 a1 : PTm n) : - rtc RPar.R a0 a1 -> - rtc RPar.R (ren_PTm shift a0) (ren_PTm shift a1). - Proof. apply renaming. Qed. - - Lemma Abs_inv n (a : PTm (S n)) b : - rtc RPar.R (PAbs a) b -> exists a', b = PAbs a' /\ rtc RPar.R a a'. - Proof. - move E : (PAbs a) => b0 h. move : a E. - elim : b0 b / h. - - hauto lq:on ctrs:rtc. - - hauto lq:on ctrs:rtc inv:RPar.R, rtc. - Qed. - - Lemma morphing n m (a b : PTm n) (ρ : fin n -> PTm m) : - rtc RPar.R a b -> - rtc RPar.R (subst_PTm ρ a) (subst_PTm ρ b). - Proof. induction 1; qauto l:on ctrs:rtc use:RPar.substing. Qed. - - Lemma substing n (a b : PTm (S n)) c : - rtc RPar.R a b -> - rtc RPar.R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons c VarPTm) b). - Proof. hauto lq:on use:morphing inv:nat. Qed. - - Lemma antirenaming n m (a : PTm n) (b : PTm m) (ρ : fin n -> PTm m) : - (forall i, var_or_const (ρ i)) -> - rtc RPar.R (subst_PTm ρ a) b -> exists b0, rtc RPar.R a b0 /\ subst_PTm ρ b0 = b. - Proof. - move E :(subst_PTm ρ a) => u hρ h. - move : a E. - elim : u b /h. - - sfirstorder. - - move => a b c h0 h1 ih1 a0 ?. subst. - move /RPar.antirenaming : h0. - move /(_ hρ). - move => [b0 [h2 ?]]. subst. - hauto lq:on rew:off ctrs:rtc. - Qed. - -End RPars. - -Module RPars'. - - #[local]Ltac solve_s_rec := - move => *; eapply rtc_l; eauto; - hauto lq:on ctrs:RPar'.R use:RPar'.refl. - - #[local]Ltac solve_s := - repeat (induction 1; last by solve_s_rec); apply rtc_refl. - - Lemma AbsCong n (a b : PTm (S n)) : - rtc RPar'.R a b -> - rtc RPar'.R (PAbs a) (PAbs b). - Proof. solve_s. Qed. - - Lemma AppCong n (a0 a1 b0 b1 : PTm n) : - rtc RPar'.R a0 a1 -> - rtc RPar'.R b0 b1 -> - rtc RPar'.R (PApp a0 b0) (PApp a1 b1). - Proof. solve_s. Qed. - - Lemma PairCong n (a0 a1 b0 b1 : PTm n) : - rtc RPar'.R a0 a1 -> - rtc RPar'.R b0 b1 -> - rtc RPar'.R (PPair a0 b0) (PPair a1 b1). - Proof. solve_s. Qed. - - Lemma ProjCong n p (a0 a1 : PTm n) : - rtc RPar'.R a0 a1 -> - rtc RPar'.R (PProj p a0) (PProj p a1). - Proof. solve_s. Qed. - - Lemma renaming n (a0 a1 : PTm n) m (ξ : fin n -> fin m) : - rtc RPar'.R a0 a1 -> - rtc RPar'.R (ren_PTm ξ a0) (ren_PTm ξ a1). - Proof. - induction 1. - - apply rtc_refl. - - eauto using RPar'.renaming, rtc_l. - Qed. - - Lemma weakening n (a0 a1 : PTm n) : - rtc RPar'.R a0 a1 -> - rtc RPar'.R (ren_PTm shift a0) (ren_PTm shift a1). - Proof. apply renaming. Qed. - - Lemma Abs_inv n (a : PTm (S n)) b : - rtc RPar'.R (PAbs a) b -> exists a', b = PAbs a' /\ rtc RPar'.R a a'. - Proof. - move E : (PAbs a) => b0 h. move : a E. - elim : b0 b / h. - - hauto lq:on ctrs:rtc. - - hauto lq:on ctrs:rtc inv:RPar'.R, rtc. - Qed. - - Lemma morphing n m (a b : PTm n) (ρ : fin n -> PTm m) : - rtc RPar'.R a b -> - rtc RPar'.R (subst_PTm ρ a) (subst_PTm ρ b). - Proof. induction 1; qauto l:on ctrs:rtc use:RPar'.substing. Qed. - - Lemma substing n (a b : PTm (S n)) c : - rtc RPar'.R a b -> - rtc RPar'.R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons c VarPTm) b). - Proof. hauto lq:on use:morphing inv:nat. Qed. - - Lemma antirenaming n m (a : PTm n) (b : PTm m) (ρ : fin n -> PTm m) : - (forall i, var_or_const (ρ i)) -> - rtc RPar'.R (subst_PTm ρ a) b -> exists b0, rtc RPar'.R a b0 /\ subst_PTm ρ b0 = b. - Proof. - move E :(subst_PTm ρ a) => u hρ h. - move : a E. - elim : u b /h. - - sfirstorder. - - move => a b c h0 h1 ih1 a0 ?. subst. - move /RPar'.antirenaming : h0. - move /(_ hρ). - move => [b0 [h2 ?]]. subst. - hauto lq:on rew:off ctrs:rtc. - Qed. - -End RPars'. - -Lemma Abs_EPar n a (b : PTm n) : - EPar.R (PAbs a) b -> - (exists d, EPar.R a d /\ - rtc RPar.R (PApp (ren_PTm shift b) (VarPTm var_zero)) d) /\ +Lemma Abs_EPar n a (b : Tm n) : + EPar.R (Abs a) b -> + (forall c m (ξ : fin n -> fin m), exists d, + EPar.R (subst_Tm (scons c VarTm) a) d /\ + rtc RPar.R (App b c) d) /\ (exists d, - EPar.R a d /\ forall p, - rtc RPar.R (PProj p b) (PAbs (PProj p d))). + EPar.R a d /\ + rtc RPar.R (Proj1 b) (Abs (Proj1 d)) /\ + rtc RPar.R (Proj2 b) (Abs (Proj2 d))). Proof. - move E : (PAbs a) => u h. - move : a E. - elim : n u b /h => //=. - - move => n a0 a1 ha iha b ?. subst. - specialize iha with (1 := eq_refl). - move : iha => [[d [ih0 ih1]] _]. - split; exists d. - + split => //. - apply : rtc_l. - apply RPar.AppAbs; eauto => //=. - apply RPar.refl. - by apply RPar.refl. - move :ih1; substify; by asimpl. - + split => // p. - apply : rtc_l. - apply : RPar.ProjAbs. - by apply RPar.refl. - eauto using RPars.ProjCong, RPars.AbsCong. - - move => n ? a1 ha iha a0 ?. subst. specialize iha with (1 := eq_refl). - move : iha => [_ [d [ih0 ih1]]]. - split. - + exists (PPair (PProj PL d) (PProj PR d)). - split; first by apply EPar.PairEta. - apply : rtc_l. - apply RPar.AppPair; eauto using RPar.refl. - suff h : forall p, rtc RPar.R (PApp (PProj p (ren_PTm shift a1)) (VarPTm var_zero)) (PProj p d) by - sfirstorder use:RPars.PairCong. - move => p. move /(_ p) /RPars.weakening in ih1. - apply relations.rtc_transitive with (y := PApp (ren_PTm shift (PAbs (PProj p d))) (VarPTm var_zero)). - by eauto using RPars.AppCong, rtc_refl. - apply relations.rtc_once => /=. - apply : RPar.AppAbs'; eauto using RPar.refl. - by asimpl. - + exists d. repeat split => //. move => p. - apply : rtc_l; eauto. - hauto q:on use:RPar.ProjPair', RPar.refl. - - move => n a0 a1 ha _ ? [*]. subst. - split. - + exists a1. split => //. - apply rtc_once. apply : RPar.AppAbs'; eauto using RPar.refl. by asimpl. - + exists a1. split => // p. - apply rtc_once. apply : RPar.ProjAbs; eauto using RPar.refl. -Qed. +(* move E : (Abs a) => u h. *) +(* move : a E. *) +(* elim : n u b /h => //=. *) +(* - move => n a0 a1 ha iha b ?. subst. *) +(* split. *) +(* + move => c. *) +(* specialize iha with (1 := eq_refl). *) +(* move : iha => [+ _]. move /(_ c) => [d [ih0 ih1]]. *) +(* exists d. *) +(* split => //. *) +(* apply : rtc_l. *) +(* apply RPar.AppAbs; eauto => //=. *) +(* apply RPar.refl. *) +(* by apply RPar.refl. *) +(* by asimpl. *) +(* + specialize iha with (1 := eq_refl). *) +(* move : iha => [_ [d [ih0 [ih1 ih2]]]]. *) +(* exists d. *) +(* repeat split => //. *) +(* * apply : rtc_l. apply : RPar.Proj1Abs. apply RPar.refl. *) +(* apply : RPars_AbsCong. *) +(* apply *) +(* - move => n a0 a1 ha iha a ? c. subst. *) +(* specialize iha with (1 := eq_refl) (c := c). *) +(* move : iha => [d [ih0 ih1]]. *) +(* exists (Pair (Proj1 d) (Proj2 d)). split => //. *) +(* + move { ih1}. *) +(* hauto lq:on ctrs:EPar.R. *) +(* + apply : rtc_l. *) +(* apply RPar.AppPair. *) +(* admit. *) +(* admit. *) +(* apply RPar.refl. *) +(* admit. *) +(* - admit. *) +Admitted. -Lemma Pair_EPar n (a b c : PTm n) : - EPar.R (PPair a b) c -> - (forall p, exists d, rtc RPar.R (PProj p c) d /\ EPar.R (if p is PL then a else b) d) /\ - (exists d0 d1, rtc RPar.R (PApp (ren_PTm shift c) (VarPTm var_zero)) - (PPair (PApp (ren_PTm shift d0) (VarPTm var_zero))(PApp (ren_PTm shift d1) (VarPTm var_zero))) /\ - EPar.R a d0 /\ EPar.R b d1). -Proof. - move E : (PPair a b) => u h. move : a b E. - elim : n u c /h => //=. - - move => n a0 a1 ha iha a b ?. subst. - specialize iha with (1 := eq_refl). - move : iha => [_ [d0 [d1 [ih0 [ih1 ih2]]]]]. - split. - + move => p. - exists (PAbs (PApp (ren_PTm shift (if p is PL then d0 else d1)) (VarPTm var_zero))). - split. - * apply : relations.rtc_transitive. - ** apply RPars.ProjCong. apply RPars.AbsCong. eassumption. - ** apply : rtc_l. apply RPar.ProjAbs; eauto using RPar.refl. apply RPars.AbsCong. - apply : rtc_l. apply RPar.ProjPair; eauto using RPar.refl. - hauto l:on. - * hauto lq:on use:EPar.AppEta'. - + exists d0, d1. - repeat split => //. - apply : rtc_l. apply : RPar.AppAbs'; eauto using RPar.refl => //=. - by asimpl; renamify. - - move => n a0 a1 ha iha a b ?. subst. specialize iha with (1 := eq_refl). - split => [p|]. - + move : iha => [/(_ p) [d [ih0 ih1]] _]. - exists d. split=>//. - apply : rtc_l. apply RPar.ProjPair; eauto using RPar.refl. - set q := (X in rtc RPar.R X d). - by have -> : q = PProj p a1 by hauto lq:on. - + move :iha => [iha _]. - move : (iha PL) => [d0 [ih0 ih0']]. - move : (iha PR) => [d1 [ih1 ih1']] {iha}. - exists d0, d1. - apply RPars.weakening in ih0, ih1. - repeat split => //=. - apply : rtc_l. apply RPar.AppPair; eauto using RPar.refl. - apply RPars.PairCong; apply RPars.AppCong; eauto using rtc_refl. - - move => n a0 a1 b0 b1 ha _ hb _ a b [*]. subst. - split. - + move => p. - exists (if p is PL then a1 else b1). - split. - * apply rtc_once. apply : RPar.ProjPair'; eauto using RPar.refl. - * hauto lq:on rew:off. - + exists a1, b1. - split. apply rtc_once. apply RPar.AppPair; eauto using RPar.refl. - split => //. -Qed. -Lemma commutativity0 n (a b0 b1 : PTm n) : +Lemma commutativity n (a b0 b1 : Tm n) : EPar.R a b0 -> RPar.R a b1 -> exists c, rtc RPar.R b0 c /\ EPar.R b1 c. Proof. move => h. move : b1. @@ -1231,1188 +303,83 @@ Proof. - move => n a b0 ha iha b1 hb. move : iha (hb) => /[apply]. move => [c [ih0 ih1]]. - exists (PAbs (PApp (ren_PTm shift c) (VarPTm var_zero))). + exists (Abs (App (ren_Tm shift c) (VarTm var_zero))). split. - + hauto lq:on ctrs:rtc use:RPars.AbsCong, RPars.AppCong, RPars.renaming. + + sfirstorder use:RPars_AbsCong, RPars_AppCong, RPars_renaming. + hauto lq:on ctrs:EPar.R use:EPar.refl, EPar.renaming. - move => n a b0 hb0 ihb0 b1 /[dup] hb1 {}/ihb0. move => [c [ih0 ih1]]. - exists (PPair (PProj PL c) (PProj PR c)). split. - + apply RPars.PairCong; - by apply RPars.ProjCong. + exists (Pair (Proj1 c) (Proj2 c)). split. + + apply RPars_PairCong; admit. + hauto lq:on ctrs:EPar.R use:EPar.refl, EPar.renaming. - hauto l:on ctrs:rtc inv:RPar.R. - move => n a0 a1 h ih b1. elim /RPar.inv => //= _. move => a2 a3 ? [*]. subst. - hauto lq:on ctrs:rtc, RPar.R, EPar.R use:RPars.AbsCong. + hauto lq:on ctrs:rtc, RPar.R, EPar.R use:RPars_AbsCong. - move => n a0 a1 b0 b1 ha iha hb ihb b2. elim /RPar.inv => //= _. + move => a2 a3 b3 b4 h0 h1 [*]. subst. - move /(_ _ ltac:(by eauto)) : ihb => [b [ihb0 ihb1]]. - have {}/iha : RPar.R (PAbs a2) (PAbs a3) by hauto lq:on ctrs:RPar.R. - move => [c [ih0 /Abs_EPar [[d [ih1 ih2]] _]]]. - exists (subst_PTm (scons b VarPTm) d). - split. - (* By substitution *) - * move /RPars.substing : ih2. - move /(_ b). - asimpl. - eauto using relations.rtc_transitive, RPars.AppCong. - (* By EPar morphing *) - * by apply EPar.substing. + have {}/iha : RPar.R (Abs a2) (Abs a3) by hauto lq:on ctrs:RPar.R. + move => [c [ih0 ih1]]. + + elim /EPar.inv : ha => //= _. + * move => a0 a4 h *. subst. + move /ihb : h1 {ihb}. + move => [c [hb1 hb4]]. + have {}/iha : RPar.R (Abs a2) (Abs a3) by hauto lq:on ctrs:RPar.R. + move => [c0 [hc0 hc1]]. + eexists. + split. + ** apply RPar.AppAbs; eauto. + eauto using RPar.refl. + ** simpl. + + + admit. + move => a2 a3 b3 b4 c0 c1 h0 h1 h2 [*]. subst. - move /(_ _ ltac:(by eauto using RPar.PairCong)) : iha - => [c [ihc0 ihc1]]. - move /(_ _ ltac:(by eauto)) : ihb => [d [ihd0 ihd1]]. - move /Pair_EPar : ihc1 => [_ [d0 [d1 [ih0 [ih1 ih2]]]]]. - move /RPars.substing : ih0. move /(_ d). - asimpl => h. - exists (PPair (PApp d0 d) (PApp d1 d)). - split. - hauto lq:on use:relations.rtc_transitive, RPars.AppCong. - apply EPar.PairCong; by apply EPar.AppCong. - + hauto lq:on ctrs:EPar.R use:RPars.AppCong. - - hauto lq:on ctrs:EPar.R inv:RPar.R use:RPars.PairCong. - - move => n p a b0 h0 ih0 b1. + admit. + + hauto lq:on ctrs:RPar.R, EPar.R. + - hauto lq:on ctrs:RPar.R, EPar.R inv:RPar.R. + - move => n a b0 h0 ih0 b1. elim /RPar.inv => //= _. - + move => ? a0 a1 h [*]. subst. - move /(_ _ ltac:(by eauto using RPar.AbsCong)) : ih0 => [c [ih0 ih1]]. - move /Abs_EPar : ih1 => [_ [d [ih1 ih2]]]. - exists (PAbs (PProj p d)). - qauto l:on ctrs:EPar.R use:RPars.ProjCong, @relations.rtc_transitive. - + move => p0 a0 a1 b2 b3 h1 h2 [*]. subst. - move /(_ _ ltac:(by eauto using RPar.PairCong)) : ih0 => [c [ih0 ih1]]. - move /Pair_EPar : ih1 => [/(_ p)[d [ihd ihd']] _]. - exists d. split => //. - hauto lq:on use:RPars.ProjCong, relations.rtc_transitive. - + hauto lq:on ctrs:EPar.R use:RPars.ProjCong. - - hauto l:on ctrs:EPar.R inv:RPar.R. - - hauto l:on ctrs:EPar.R inv:RPar.R. - - hauto l:on ctrs:EPar.R inv:RPar.R. -Qed. - -Lemma commutativity1 n (a b0 b1 : PTm n) : - EPar.R a b0 -> rtc RPar.R a b1 -> exists c, rtc RPar.R b0 c /\ EPar.R b1 c. -Proof. - move => + h. move : b0. - elim : a b1 / h. - - sfirstorder. - - qauto l:on use:relations.rtc_transitive, commutativity0. -Qed. - -Lemma commutativity n (a b0 b1 : PTm n) : - rtc EPar.R a b0 -> rtc RPar.R a b1 -> exists c, rtc RPar.R b0 c /\ rtc EPar.R b1 c. - move => h. move : b1. elim : a b0 /h. - - sfirstorder. - - move => a0 a1 a2 + ha1 ih b1 +. - move : commutativity1; repeat move/[apply]. - hauto q:on ctrs:rtc. -Qed. - -Lemma Abs_EPar' n a (b : PTm n) : - EPar.R (PAbs a) b -> - (exists d, EPar.R a d /\ - rtc OExp.R (PAbs d) b). -Proof. - move E : (PAbs a) => u h. - move : a E. - elim : n u b /h => //=. - - move => n a0 a1 ha iha a ?. subst. - specialize iha with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - move => n a0 a1 ha iha a ?. subst. - specialize iha with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - hauto l:on ctrs:OExp.R. -Qed. - -Lemma Proj_EPar' n p a (b : PTm n) : - EPar.R (PProj p a) b -> - (exists d, EPar.R a d /\ - rtc OExp.R (PProj p d) b). -Proof. - move E : (PProj p a) => u h. - move : p a E. - elim : n u b /h => //=. - - move => n a0 a1 ha iha a p ?. subst. - specialize iha with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - move => n a0 a1 ha iha a p ?. subst. - specialize iha with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - hauto l:on ctrs:OExp.R. -Qed. - -Lemma App_EPar' n (a b u : PTm n) : - EPar.R (PApp a b) u -> - (exists a0 b0, EPar.R a a0 /\ EPar.R b b0 /\ rtc OExp.R (PApp a0 b0) u). -Proof. - move E : (PApp a b) => t h. - move : a b E. elim : n t u /h => //=. - - move => n a0 a1 ha iha a b ?. subst. - specialize iha with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - move => n a0 a1 ha iha a b ?. subst. - specialize iha with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - hauto l:on ctrs:OExp.R. -Qed. - -Lemma Pair_EPar' n (a b u : PTm n) : - EPar.R (PPair a b) u -> - exists a0 b0, EPar.R a a0 /\ EPar.R b b0 /\ rtc OExp.R (PPair a0 b0) u. -Proof. - move E : (PPair a b) => t h. - move : a b E. elim : n t u /h => //=. - - move => n a0 a1 ha iha a b ?. subst. - specialize iha with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - move => n a0 a1 ha iha a b ?. subst. - specialize iha with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - hauto l:on ctrs:OExp.R. -Qed. - -Lemma Const_EPar' n k (u : PTm n) : - EPar.R (PConst k) u -> - rtc OExp.R (PConst k) u. - move E : (PConst k) => t h. - move : k E. elim : n t u /h => //=. - - move => n a0 a1 h ih k ?. subst. - specialize ih with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - move => n a0 a1 h ih k ?. subst. - specialize ih with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - hauto l:on ctrs:OExp.R. -Qed. - -Lemma Bot_EPar' n (u : PTm n) : - EPar.R (PBot) u -> - rtc OExp.R (PBot) u. - move E : (PBot) => t h. - move : E. elim : n t u /h => //=. - - move => n a0 a1 h ih ?. subst. - specialize ih with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - move => n a0 a1 h ih ?. subst. - specialize ih with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - hauto l:on ctrs:OExp.R. -Qed. - -Lemma Univ_EPar' n i (u : PTm n) : - EPar.R (PUniv i) u -> - rtc OExp.R (PUniv i) u. - move E : (PUniv i) => t h. - move : E. elim : n t u /h => //=. - - move => n a0 a1 h ih ?. subst. - specialize ih with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - move => n a0 a1 h ih ?. subst. - specialize ih with (1 := eq_refl). - hauto lq:on ctrs:OExp.R use:rtc_r. - - hauto l:on ctrs:OExp.R. -Qed. - -Lemma EPar_diamond n (c a1 b1 : PTm n) : - EPar.R c a1 -> - EPar.R c b1 -> - exists d2, EPar.R a1 d2 /\ EPar.R b1 d2. -Proof. - move => h. move : b1. elim : n c a1 / h. - - move => n c a1 ha iha b1 /iha [d2 [hd0 hd1]]. - exists(PAbs (PApp (ren_PTm shift d2) (VarPTm var_zero))). - hauto lq:on ctrs:EPar.R use:EPar.renaming. - - hauto lq:on rew:off ctrs:EPar.R. - - hauto lq:on use:EPar.refl. - - move => n a0 a1 ha iha a2. - move /Abs_EPar' => [d [hd0 hd1]]. - move : iha hd0; repeat move/[apply]. - move => [d2 [h0 h1]]. - have : EPar.R (PAbs d) (PAbs d2) by eauto using EPar.AbsCong. - move : OExp.commutativity0 hd1; repeat move/[apply]. - move => [d1 [hd1 hd2]]. - exists d1. hauto lq:on ctrs:EPar.R use:OExp.merge. - - move => n a0 a1 b0 b1 ha iha hb ihb c. - move /App_EPar' => [a2][b2][/iha [a3 h0]][/ihb [b3 h1]]h2 {iha ihb}. - have : EPar.R (PApp a2 b2)(PApp a3 b3) - by hauto l:on use:EPar.AppCong. - move : OExp.commutativity0 h2; repeat move/[apply]. - move => [d h]. - exists d. hauto lq:on rew:off ctrs:EPar.R use:OExp.merge. - - move => n a0 a1 b0 b1 ha iha hb ihb c. - move /Pair_EPar' => [a2][b2][/iha [a3 h0]][/ihb [b3 h1]]h2 {iha ihb}. - have : EPar.R (PPair a2 b2)(PPair a3 b3) - by hauto l:on use:EPar.PairCong. - move : OExp.commutativity0 h2; repeat move/[apply]. - move => [d h]. - exists d. hauto lq:on rew:off ctrs:EPar.R use:OExp.merge. - - move => n p a0 a1 ha iha b. - move /Proj_EPar' => [d [/iha [d2 h] h1]] {iha}. - have : EPar.R (PProj p d) (PProj p d2) - by hauto l:on use:EPar.ProjCong. - move : OExp.commutativity0 h1; repeat move/[apply]. - move => [d1 h1]. - exists d1. hauto lq:on rew:off ctrs:EPar.R use:OExp.merge. - - qauto use:Const_EPar', EPar.refl. - - qauto use:Univ_EPar', EPar.refl. - - qauto use:Bot_EPar', EPar.refl. -Qed. - -Function tstar {n} (a : PTm n) := - match a with - | VarPTm i => a - | PAbs a => PAbs (tstar a) - | PApp (PAbs a) b => subst_PTm (scons (tstar b) VarPTm) (tstar a) - | PApp (PPair a b) c => - PPair (PApp (tstar a) (tstar c)) (PApp (tstar b) (tstar c)) - | PApp a b => PApp (tstar a) (tstar b) - | PPair a b => PPair (tstar a) (tstar b) - | PProj p (PPair a b) => if p is PL then (tstar a) else (tstar b) - | PProj p (PAbs a) => (PAbs (PProj p (tstar a))) - | PProj p a => PProj p (tstar a) - | PConst k => PConst k - | PUniv i => PUniv i - | PBot => PBot - end. - -Lemma RPar_triangle n (a : PTm n) : forall b, RPar.R a b -> RPar.R b (tstar a). -Proof. - apply tstar_ind => {n a}. - - hauto lq:on inv:RPar.R ctrs:RPar.R. - - hauto lq:on inv:RPar.R ctrs:RPar.R. - - hauto lq:on use:RPar.cong, RPar.refl ctrs:RPar.R inv:RPar.R. - - hauto lq:on rew:off ctrs:RPar.R inv:RPar.R. - - hauto lq:on rew:off inv:RPar.R ctrs:RPar.R. - - hauto lq:on rew:off inv:RPar.R ctrs:RPar.R. - - hauto drew:off inv:RPar.R use:RPar.refl, RPar.ProjPair'. - - hauto drew:off inv:RPar.R use:RPar.refl, RPar.ProjPair'. - - hauto lq:on inv:RPar.R ctrs:RPar.R. - - hauto lq:on inv:RPar.R ctrs:RPar.R. - - hauto lq:on inv:RPar.R ctrs:RPar.R. - - hauto lq:on inv:RPar.R ctrs:RPar.R. - - hauto lq:on inv:RPar.R ctrs:RPar.R. -Qed. - -Function tstar' {n} (a : PTm n) := - match a with - | VarPTm i => a - | PAbs a => PAbs (tstar' a) - | PApp (PAbs a) b => subst_PTm (scons (tstar' b) VarPTm) (tstar' a) - | PApp a b => PApp (tstar' a) (tstar' b) - | PPair a b => PPair (tstar' a) (tstar' b) - | PProj p (PPair a b) => if p is PL then (tstar' a) else (tstar' b) - | PProj p a => PProj p (tstar' a) - | PConst k => PConst k - | PUniv i => PUniv i - | PBot => PBot - end. - -Lemma RPar'_triangle n (a : PTm n) : forall b, RPar'.R a b -> RPar'.R b (tstar' a). -Proof. - apply tstar'_ind => {n a}. - - hauto lq:on inv:RPar'.R ctrs:RPar'.R. - - hauto lq:on inv:RPar'.R ctrs:RPar'.R. - - hauto lq:on use:RPar'.cong, RPar'.refl ctrs:RPar'.R inv:RPar'.R. - - hauto lq:on rew:off ctrs:RPar'.R inv:RPar'.R. - - hauto lq:on rew:off inv:RPar'.R ctrs:RPar'.R. - - hauto drew:off inv:RPar'.R use:RPar'.refl, RPar'.ProjPair'. - - hauto drew:off inv:RPar'.R use:RPar'.refl, RPar'.ProjPair'. - - hauto lq:on inv:RPar'.R ctrs:RPar'.R. - - hauto lq:on inv:RPar'.R ctrs:RPar'.R. - - hauto lq:on inv:RPar'.R ctrs:RPar'.R. - - hauto lq:on inv:RPar'.R ctrs:RPar'.R. -Qed. - -Lemma RPar_diamond n (c a1 b1 : PTm n) : - RPar.R c a1 -> - RPar.R c b1 -> - exists d2, RPar.R a1 d2 /\ RPar.R b1 d2. -Proof. hauto l:on use:RPar_triangle. Qed. - -Lemma RPar'_diamond n (c a1 b1 : PTm n) : - RPar'.R c a1 -> - RPar'.R c b1 -> - exists d2, RPar'.R a1 d2 /\ RPar'.R b1 d2. -Proof. hauto l:on use:RPar'_triangle. Qed. - -Lemma RPar_confluent n (c a1 b1 : PTm n) : - rtc RPar.R c a1 -> - rtc RPar.R c b1 -> - exists d2, rtc RPar.R a1 d2 /\ rtc RPar.R b1 d2. -Proof. - sfirstorder use:relations.diamond_confluent, RPar_diamond. -Qed. - -Lemma EPar_confluent n (c a1 b1 : PTm n) : - rtc EPar.R c a1 -> - rtc EPar.R c b1 -> - exists d2, rtc EPar.R a1 d2 /\ rtc EPar.R b1 d2. -Proof. - sfirstorder use:relations.diamond_confluent, EPar_diamond. -Qed. - -Inductive prov {n} : PTm n -> PTm n -> Prop := -| P_Abs h a : - (forall b, prov h (subst_PTm (scons b VarPTm) a)) -> - prov h (PAbs a) -| P_App h a b : - prov h a -> - prov h (PApp a b) -| P_Pair h a b : - prov h a -> - prov h b -> - prov h (PPair a b) -| P_Proj h p a : - prov h a -> - prov h (PProj p a) -| P_Const k : - prov (PConst k) (PConst k) -| P_Var i : - prov (VarPTm i) (VarPTm i) -| P_Univ i : - prov (PUniv i) (PUniv i) -| P_Bot : - prov PBot PBot. - -Lemma ERed_EPar n (a b : PTm n) : ERed.R a b -> EPar.R a b. -Proof. - induction 1; hauto lq:on ctrs:EPar.R use:EPar.refl. -Qed. - -Lemma EPar_ERed n (a b : PTm n) : EPar.R a b -> rtc ERed.R a b. -Proof. - move => h. elim : n a b /h. - - eauto using rtc_r, ERed.AppEta. - - eauto using rtc_r, ERed.PairEta. - - auto using rtc_refl. - - eauto using EReds.AbsCong. - - eauto using EReds.AppCong. - - eauto using EReds.PairCong. - - eauto using EReds.ProjCong. - - auto using rtc_refl. - - auto using rtc_refl. - - auto using rtc_refl. -Qed. - -Lemma EPar_Par n (a b : PTm n) : EPar.R a b -> Par.R a b. -Proof. - move => h. elim : n a b /h; qauto ctrs:Par.R. -Qed. - -Lemma RPar_Par n (a b : PTm n) : RPar.R a b -> Par.R a b. -Proof. - move => h. elim : n a b /h; hauto lq:on ctrs:Par.R. -Qed. - -Lemma rtc_idem n (R : PTm n -> PTm n -> Prop) (a b : PTm n) : rtc (rtc R) a b -> rtc R a b. -Proof. - induction 1; hauto l:on use:@relations.rtc_transitive, @rtc_r. -Qed. - -Lemma EPars_EReds {n} (a b : PTm n) : rtc EPar.R a b <-> rtc ERed.R a b. -Proof. - sfirstorder use:@relations.rtc_subrel, EPar_ERed, rtc_idem, ERed_EPar. -Qed. - -Lemma prov_rpar n (u : PTm n) a b : prov u a -> RPar.R a b -> prov u b. -Proof. - move => h. - move : b. - elim : u a / h. - (* - qauto l:on ctrs:prov inv:RPar.R use:@rtc_r, RPar_Par. *) - - hauto lq:on ctrs:prov inv:RPar.R use:RPar.substing. - - move => h a b ha iha b0. + + move => a0 a1 h [*]. subst. + admit. + + move => a0 ? a1 h1 [*]. subst. + admit. + + hauto lq:on ctrs:RPar.R, EPar.R. + - move => n a b0 h0 ih0 b1. elim /RPar.inv => //= _. - + move => a0 a1 b1 b2 h0 h1 [*]. subst. - have {}iha : prov h (PAbs a1) by hauto lq:on ctrs:RPar.R. - hauto lq:on inv:prov use:RPar.substing. - + move => a0 a1 b1 b2 c0 c1. - move => h0 h1 h2 [*]. subst. - have {}iha : prov h (PPair a1 b2) by hauto lq:on ctrs:RPar.R. - hauto lq:on inv:prov ctrs:prov. - + hauto lq:on ctrs:prov. - - hauto lq:on ctrs:prov inv:RPar.R. - - move => h p a ha iha b. - elim /RPar.inv => //= _. - + move => p0 a0 a1 h0 [*]. subst. - have {iha} : prov h (PAbs a1) by hauto lq:on ctrs:RPar.R. - hauto lq:on ctrs:prov inv:prov use:RPar.substing. - + move => p0 a0 a1 b0 b1 h0 h1 [*]. subst. - have {iha} : prov h (PPair a1 b1) by hauto lq:on ctrs:RPar.R. - qauto l:on inv:prov. - + hauto lq:on ctrs:prov. - - hauto lq:on ctrs:prov inv:RPar.R. - - hauto l:on ctrs:RPar.R inv:RPar.R. - - hauto l:on ctrs:RPar.R inv:RPar.R. - - hauto l:on ctrs:RPar.R inv:RPar.R. -Qed. + + move => a0 a1 ha [*]. subst. + admit. + + move => a0 a1 b2 h [*]. subst. + admit. + + hauto lq:on ctrs:RPar.R, EPar.R. +Admitted. +Lemma EPar_Par n (a b : Tm n) : EPar.R a b -> Par.R a b. +Proof. induction 1; hauto lq:on ctrs:Par.R. Qed. -Lemma prov_lam n (u : PTm n) a : prov u a <-> prov u (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero))). +Lemma RPar_Par n (a b : Tm n) : RPar.R a b -> Par.R a b. +Proof. induction 1; hauto lq:on ctrs:Par.R. Qed. + +Lemma merge n (t a u : Tm n) : + EPar.R t a -> + RPar.R a u -> + Par.R t u. Proof. - split. - move => h. constructor. move => b. asimpl. by constructor. - inversion 1; subst. - specialize H2 with (b := PBot). - move : H2. asimpl. inversion 1; subst. done. -Qed. + move => h. move : u. + elim:t a/h. + - move => n0 a0 a1 ha iha u hu. + apply iha. + inversion hu; subst. -Lemma prov_pair n (u : PTm n) a : prov u a <-> prov u (PPair (PProj PL a) (PProj PR a)). -Proof. hauto lq:on inv:prov ctrs:prov. Qed. -Lemma prov_ered n (u : PTm n) a b : prov u a -> ERed.R a b -> prov u b. -Proof. - move => h. - move : b. - elim : u a / h. - - move => h a ha iha b. - elim /ERed.inv => // _. - + move => a0 *. subst. - rewrite -prov_lam. - by constructor. - + move => a0 *. subst. - rewrite -prov_pair. - by constructor. - + hauto lq:on ctrs:prov use:ERed.substing. - - hauto lq:on inv:ERed.R, prov ctrs:prov. - - move => h a b ha iha hb ihb b0. - elim /ERed.inv => //_. - + move => a0 *. subst. - rewrite -prov_lam. - by constructor. - + move => a0 *. subst. - rewrite -prov_pair. - by constructor. - + hauto lq:on ctrs:prov. - + hauto lq:on ctrs:prov. - - hauto lq:on inv:ERed.R, prov ctrs:prov. - - hauto lq:on inv:ERed.R, prov ctrs:prov. - - hauto lq:on inv:ERed.R, prov ctrs:prov. - - hauto lq:on inv:ERed.R, prov ctrs:prov. - - hauto lq:on inv:ERed.R, prov ctrs:prov. -Qed. + - hauto lq:on inv:RPar.R. + - move => a0 a1 b0 b1 ha iha hb ihb u. + inversion 1; subst. + + inversion ha. -Lemma prov_ereds n (u : PTm n) a b : prov u a -> rtc ERed.R a b -> prov u b. -Proof. - induction 2; sfirstorder use:prov_ered. -Qed. +best use:EPar_Par, RPar_Par. -Fixpoint extract {n} (a : PTm n) : PTm n := - match a with - | PAbs a => subst_PTm (scons PBot VarPTm) (extract a) - | PApp a b => extract a - | PPair a b => extract a - | PProj p a => extract a - | PConst k => PConst k - | VarPTm i => VarPTm i - | PUniv i => PUniv i - | PBot => PBot - end. - -Lemma ren_extract n m (a : PTm n) (ξ : fin n -> fin m) : - extract (ren_PTm ξ a) = ren_PTm ξ (extract a). -Proof. - move : m ξ. elim : n/a. - - sfirstorder. - - move => n a ih m ξ /=. - rewrite ih. - by asimpl. - - hauto q:on. - - hauto q:on. - - hauto q:on. - - hauto q:on. - - sfirstorder. - - sfirstorder. -Qed. - -Lemma ren_morphing n m (a : PTm n) (ρ : fin n -> PTm m) : - (forall i, ρ i = extract (ρ i)) -> - extract (subst_PTm ρ a) = subst_PTm ρ (extract a). -Proof. - move : m ρ. - elim : n /a => n //=. - move => a ha m ρ hi. - rewrite ha. - - destruct i as [i|] => //. - rewrite ren_extract. - rewrite -hi. - by asimpl. - - by asimpl. -Qed. - -Lemma ren_subst_bot n (a : PTm (S n)) : - extract (subst_PTm (scons PBot VarPTm) a) = subst_PTm (scons PBot VarPTm) (extract a). -Proof. - apply ren_morphing. destruct i as [i|] => //=. -Qed. - -Definition prov_extract_spec {n} u (a : PTm n) := - match u with - | PUniv i => extract a = PUniv i - | VarPTm i => extract a = VarPTm i - | (PConst i) => extract a = (PConst i) - | PBot => extract a = PBot - | _ => True - end. - -Lemma prov_extract n u (a : PTm n) : - prov u a -> prov_extract_spec u a. -Proof. - move => h. - elim : u a /h. - - move => h a ha ih. - case : h ha ih => //=. - + move => i ha ih. - move /(_ PBot) in ih. - rewrite -ih. - by rewrite ren_subst_bot. - + move => p _ /(_ PBot). - by rewrite ren_subst_bot. - + move => i h /(_ PBot). - by rewrite ren_subst_bot => ->. - + move /(_ PBot). - move => h /(_ PBot). - by rewrite ren_subst_bot. - - hauto lq:on. - - hauto lq:on. - - hauto lq:on. - - case => //=. - - sfirstorder. - - sfirstorder. - - sfirstorder. -Qed. - -Definition union {A : Type} (R0 R1 : A -> A -> Prop) a b := - R0 a b \/ R1 a b. - -Module ERPar. - Definition R {n} (a b : PTm n) := union RPar.R EPar.R a b. - Lemma RPar {n} (a b : PTm n) : RPar.R a b -> R a b. - Proof. sfirstorder. Qed. - - Lemma EPar {n} (a b : PTm n) : EPar.R a b -> R a b. - Proof. sfirstorder. Qed. - - Lemma refl {n} ( a : PTm n) : ERPar.R a a. - Proof. - sfirstorder use:RPar.refl, EPar.refl. - Qed. - - Lemma ProjCong n p (a0 a1 : PTm n) : - R a0 a1 -> - rtc R (PProj p a0) (PProj p a1). - Proof. - move => []. - - move => h. - apply rtc_once. - left. - by apply RPar.ProjCong. - - move => h. - apply rtc_once. - right. - by apply EPar.ProjCong. - Qed. - - Lemma AbsCong n (a0 a1 : PTm (S n)) : - R a0 a1 -> - rtc R (PAbs a0) (PAbs a1). - Proof. - move => []. - - move => h. - apply rtc_once. - left. - by apply RPar.AbsCong. - - move => h. - apply rtc_once. - right. - by apply EPar.AbsCong. - Qed. - - Lemma AppCong n (a0 a1 b0 b1 : PTm n) : - R a0 a1 -> - R b0 b1 -> - rtc R (PApp a0 b0) (PApp a1 b1). - Proof. - move => [] + []. - - sfirstorder use:RPar.AppCong, @rtc_once. - - move => h0 h1. - apply : rtc_l. - left. apply RPar.AppCong; eauto; apply RPar.refl. - apply rtc_once. - hauto l:on use:EPar.AppCong, EPar.refl. - - move => h0 h1. - apply : rtc_l. - left. apply RPar.AppCong; eauto; apply RPar.refl. - apply rtc_once. - hauto l:on use:EPar.AppCong, EPar.refl. - - sfirstorder use:EPar.AppCong, @rtc_once. - Qed. - - Lemma PairCong n (a0 a1 b0 b1 : PTm n) : - R a0 a1 -> - R b0 b1 -> - rtc R (PPair a0 b0) (PPair a1 b1). - Proof. - move => [] + []. - - sfirstorder use:RPar.PairCong, @rtc_once. - - move => h0 h1. - apply : rtc_l. - left. apply RPar.PairCong; eauto; apply RPar.refl. - apply rtc_once. - hauto l:on use:EPar.PairCong, EPar.refl. - - move => h0 h1. - apply : rtc_l. - left. apply RPar.PairCong; eauto; apply RPar.refl. - apply rtc_once. - hauto l:on use:EPar.PairCong, EPar.refl. - - sfirstorder use:EPar.PairCong, @rtc_once. - Qed. - - Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) : - R a b -> R (ren_PTm ξ a) (ren_PTm ξ b). - Proof. - sfirstorder use:EPar.renaming, RPar.renaming. - Qed. - -End ERPar. - -Hint Resolve ERPar.AppCong ERPar.refl ERPar.AbsCong ERPar.PairCong ERPar.ProjCong : erpar. - -Module ERPars. - #[local]Ltac solve_s_rec := - move => *; eapply relations.rtc_transitive; eauto; - hauto lq:on db:erpar. - #[local]Ltac solve_s := - repeat (induction 1; last by solve_s_rec); apply rtc_refl. - - Lemma AppCong n (a0 a1 b0 b1 : PTm n) : - rtc ERPar.R a0 a1 -> - rtc ERPar.R b0 b1 -> - rtc ERPar.R (PApp a0 b0) (PApp a1 b1). - Proof. solve_s. Qed. - - Lemma AbsCong n (a0 a1 : PTm (S n)) : - rtc ERPar.R a0 a1 -> - rtc ERPar.R (PAbs a0) (PAbs a1). - Proof. solve_s. Qed. - - Lemma PairCong n (a0 a1 b0 b1 : PTm n) : - rtc ERPar.R a0 a1 -> - rtc ERPar.R b0 b1 -> - rtc ERPar.R (PPair a0 b0) (PPair a1 b1). - Proof. solve_s. Qed. - - Lemma ProjCong n p (a0 a1 : PTm n) : - rtc ERPar.R a0 a1 -> - rtc ERPar.R (PProj p a0) (PProj p a1). - Proof. solve_s. Qed. - - Lemma renaming n (a0 a1 : PTm n) m (ξ : fin n -> fin m) : - rtc ERPar.R a0 a1 -> - rtc ERPar.R (ren_PTm ξ a0) (ren_PTm ξ a1). - Proof. - induction 1. - - apply rtc_refl. - - eauto using ERPar.renaming, rtc_l. - Qed. - -End ERPars. - -Lemma ERPar_Par n (a b : PTm n) : ERPar.R a b -> Par.R a b. -Proof. - sfirstorder use:EPar_Par, RPar_Par. -Qed. - -Lemma Par_ERPar n (a b : PTm n) : Par.R a b -> rtc ERPar.R a b. -Proof. - move => h. elim : n a b /h. - - move => n a0 a1 b0 b1 ha iha hb ihb. - suff ? : rtc ERPar.R (PApp (PAbs a0) b0) (PApp (PAbs a1) b1). - apply : relations.rtc_transitive; eauto. - apply rtc_once. apply ERPar.RPar. - by apply RPar.AppAbs; eauto using RPar.refl. - eauto using ERPars.AppCong,ERPars.AbsCong. - - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc. - apply : rtc_l. apply ERPar.RPar. - apply RPar.AppPair; eauto using RPar.refl. - sfirstorder use:ERPars.AppCong, ERPars.PairCong. - - move => n p a0 a1 ha iha. - apply : rtc_l. apply ERPar.RPar. apply RPar.ProjAbs; eauto using RPar.refl. - sfirstorder use:ERPars.AbsCong, ERPars.ProjCong. - - move => n p a0 a1 b0 b1 ha iha hb ihb. - apply : rtc_l. apply ERPar.RPar. apply RPar.ProjPair; eauto using RPar.refl. - hauto lq:on. - - move => n a0 a1 ha iha. - apply : rtc_l. apply ERPar.EPar. apply EPar.AppEta; eauto using EPar.refl. - hauto lq:on ctrs:rtc - use:ERPars.AppCong, ERPars.AbsCong, ERPars.renaming. - - move => n a0 a1 ha iha. - apply : rtc_l. apply ERPar.EPar. apply EPar.PairEta; eauto using EPar.refl. - sfirstorder use:ERPars.PairCong, ERPars.ProjCong. - - sfirstorder. - - sfirstorder use:ERPars.AbsCong. - - sfirstorder use:ERPars.AppCong. - - sfirstorder use:ERPars.PairCong. - - sfirstorder use:ERPars.ProjCong. - - sfirstorder. - - sfirstorder. - - sfirstorder. -Qed. - -Lemma Pars_ERPar n (a b : PTm n) : rtc Par.R a b -> rtc ERPar.R a b. -Proof. - induction 1; hauto l:on use:Par_ERPar, @relations.rtc_transitive. -Qed. - -Lemma Par_ERPar_iff n (a b : PTm n) : rtc Par.R a b <-> rtc ERPar.R a b. -Proof. - split. - sfirstorder use:Pars_ERPar, @relations.rtc_subrel. - sfirstorder use:ERPar_Par, @relations.rtc_subrel. -Qed. - -Lemma RPar_ERPar n (a b : PTm n) : rtc RPar.R a b -> rtc ERPar.R a b. -Proof. - sfirstorder use:@relations.rtc_subrel. -Qed. - -Lemma EPar_ERPar n (a b : PTm n) : rtc EPar.R a b -> rtc ERPar.R a b. -Proof. - sfirstorder use:@relations.rtc_subrel. -Qed. - -Module Type HindleyRosen. - Parameter A : nat -> Type. - Parameter R0 R1 : forall n, A n -> A n -> Prop. - Axiom diamond_R0 : forall n, relations.diamond (R0 n). - Axiom diamond_R1 : forall n, relations.diamond (R1 n). - Axiom commutativity : forall n, - forall a b c, R0 n a b -> R1 n a c -> exists d, R1 n b d /\ R0 n c d. -End HindleyRosen. - -Module HindleyRosenFacts (M : HindleyRosen). - Import M. - Lemma R0_comm : - forall n a b c, R0 n a b -> rtc (union (R0 n) (R1 n)) a c -> - exists d, rtc (union (R0 n) (R1 n)) b d /\ R0 n c d. - Proof. - move => n a + c + h. - elim : a c /h. - - sfirstorder. - - move => a0 a1 a2 ha ha0 ih b h. - case : ha. - + move : diamond_R0 h; repeat move/[apply]. - hauto lq:on ctrs:rtc. - + move : commutativity h; repeat move/[apply]. - hauto lq:on ctrs:rtc. - Qed. - - Lemma R1_comm : - forall n a b c, R1 n a b -> rtc (union (R0 n) (R1 n)) a c -> - exists d, rtc (union (R0 n) (R1 n)) b d /\ R1 n c d. - Proof. - move => n a + c + h. - elim : a c /h. - - sfirstorder. - - move => a0 a1 a2 ha ha0 ih b h. - case : ha. - + move : commutativity h; repeat move/[apply]. - hauto lq:on ctrs:rtc. - + move : diamond_R1 h; repeat move/[apply]. - hauto lq:on ctrs:rtc. - Qed. - - Lemma U_comm : - forall n a b c, (union (R0 n) (R1 n)) a b -> rtc (union (R0 n) (R1 n)) a c -> - exists d, rtc (union (R0 n) (R1 n)) b d /\ (union (R0 n) (R1 n)) c d. - Proof. - hauto lq:on use:R0_comm, R1_comm. - Qed. - - Lemma U_comms : - forall n a b c, rtc (union (R0 n) (R1 n)) a b -> rtc (union (R0 n) (R1 n)) a c -> - exists d, rtc (union (R0 n) (R1 n)) b d /\ rtc (union (R0 n) (R1 n)) c d. - Proof. - move => n a b + h. - elim : a b /h. - - sfirstorder. - - hecrush ctrs:rtc use:U_comm. - Qed. - -End HindleyRosenFacts. - -Module HindleyRosenER <: HindleyRosen. - Definition A := PTm. - Definition R0 n := rtc (@RPar.R n). - Definition R1 n := rtc (@EPar.R n). - Lemma diamond_R0 : forall n, relations.diamond (R0 n). - sfirstorder use:RPar_confluent. - Qed. - Lemma diamond_R1 : forall n, relations.diamond (R1 n). - sfirstorder use:EPar_confluent. - Qed. - Lemma commutativity : forall n, - forall a b c, R0 n a b -> R1 n a c -> exists d, R1 n b d /\ R0 n c d. - Proof. - hauto l:on use:commutativity. - Qed. -End HindleyRosenER. - -Module ERFacts := HindleyRosenFacts HindleyRosenER. - -Lemma rtc_union n (a b : PTm n) : - rtc (union RPar.R EPar.R) a b <-> - rtc (union (rtc RPar.R) (rtc EPar.R)) a b. -Proof. - split; first by induction 1; hauto lq:on ctrs:rtc. - move => h. - elim :a b /h. - - sfirstorder. - - move => a0 a1 a2. - case. - + move => h0 h1 ih. - apply : relations.rtc_transitive; eauto. - move : h0. - apply relations.rtc_subrel. - sfirstorder. - + move => h0 h1 ih. - apply : relations.rtc_transitive; eauto. - move : h0. - apply relations.rtc_subrel. - sfirstorder. -Qed. - -Lemma prov_erpar n (u : PTm n) a b : prov u a -> ERPar.R a b -> prov u b. -Proof. - move => h []. - - sfirstorder use:prov_rpar. - - move /EPar_ERed. - sfirstorder use:prov_ereds. -Qed. - -Lemma prov_pars n (u : PTm n) a b : prov u a -> rtc Par.R a b -> prov u b. -Proof. - move => h /Pars_ERPar. - move => h0. - move : h. - elim : a b /h0. - - done. - - hauto lq:on use:prov_erpar. -Qed. - -Lemma Par_confluent n (a b c : PTm n) : - rtc Par.R a b -> - rtc Par.R a c -> - exists d, rtc Par.R b d /\ rtc Par.R c d. -Proof. - move : n a b c. - suff : forall (n : nat) (a b c : PTm n), - rtc ERPar.R a b -> - rtc ERPar.R a c -> exists d : PTm n, rtc ERPar.R b d /\ rtc ERPar.R c d. - move => h n a b c h0 h1. - apply Par_ERPar_iff in h0, h1. - move : h h0 h1; repeat move/[apply]. - hauto lq:on use:Par_ERPar_iff. - have h := ERFacts.U_comms. - move => n a b c. - rewrite /HindleyRosenER.R0 /HindleyRosenER.R1 in h. - specialize h with (n := n). - rewrite /HindleyRosenER.A in h. - rewrite /ERPar.R. - have eq : (fun a0 b0 : PTm n => union RPar.R EPar.R a0 b0) = union RPar.R EPar.R by reflexivity. - rewrite !{}eq. - move /rtc_union => + /rtc_union. - move : h; repeat move/[apply]. - hauto lq:on use:rtc_union. -Qed. - -Lemma pars_univ_inv n i (c : PTm n) : - rtc Par.R (PUniv i) c -> - extract c = PUniv i. -Proof. - have : prov (PUniv i) (PUniv i : PTm n) by sfirstorder. - move : prov_pars. repeat move/[apply]. - apply prov_extract. -Qed. - -Lemma pars_const_inv n i (c : PTm n) : - rtc Par.R (PConst i) c -> - extract c = PConst i. -Proof. - have : prov (PConst i) (PConst i : PTm n) by sfirstorder. - move : prov_pars. repeat move/[apply]. - apply prov_extract. -Qed. - -Lemma pars_var_inv n (i : fin n) C : - rtc Par.R (VarPTm i) C -> - extract C = VarPTm i. -Proof. - have : prov (VarPTm i) (VarPTm i) by hauto lq:on ctrs:prov, rtc. - move : prov_pars. repeat move/[apply]. - apply prov_extract. -Qed. - -Lemma pars_univ_inj n i j (C : PTm n) : - rtc Par.R (PUniv i) C -> - rtc Par.R (PUniv j) C -> - i = j. -Proof. - sauto l:on use:pars_univ_inv. -Qed. - -Lemma pars_const_inj n i j (C : PTm n) : - rtc Par.R (PConst i) C -> - rtc Par.R (PConst j) C -> - i = j. -Proof. - sauto l:on use:pars_const_inv. -Qed. - -Definition join {n} (a b : PTm n) := - exists c, rtc Par.R a c /\ rtc Par.R b c. - -Lemma join_transitive n (a b c : PTm n) : - join a b -> join b c -> join a c. -Proof. - rewrite /join. - move => [ab [h0 h1]] [bc [h2 h3]]. - move : Par_confluent h1 h2; repeat move/[apply]. - move => [abc [h4 h5]]. - eauto using relations.rtc_transitive. -Qed. - -Lemma join_symmetric n (a b : PTm n) : - join a b -> join b a. -Proof. sfirstorder unfold:join. Qed. - -Lemma join_refl n (a : PTm n) : join a a. -Proof. hauto lq:on ctrs:rtc unfold:join. Qed. - -Lemma join_univ_inj n i j : - join (PUniv i : PTm n) (PUniv j) -> i = j. -Proof. - sfirstorder use:pars_univ_inj. -Qed. - -Lemma join_const_inj n i j : - join (PConst i : PTm n) (PConst j) -> i = j. -Proof. - sfirstorder use:pars_const_inj. -Qed. - -Lemma join_substing n m (a b : PTm n) (ρ : fin n -> PTm m) : - join a b -> - join (subst_PTm ρ a) (subst_PTm ρ b). -Proof. hauto lq:on unfold:join use:Pars.substing. Qed. - -Fixpoint ne {n} (a : PTm n) := - match a with - | VarPTm i => true - | PApp a b => ne a && nf b - | PAbs a => false - | PUniv _ => false - | PProj _ a => ne a - | PPair _ _ => false - | PConst _ => false - | PBot => true - end -with nf {n} (a : PTm n) := - match a with - | VarPTm i => true - | PApp a b => ne a && nf b - | PAbs a => nf a - | PUniv _ => true - | PProj _ a => ne a - | PPair a b => nf a && nf b - | PConst _ => true - | PBot => true -end. - -Lemma ne_nf n a : @ne n a -> nf a. -Proof. elim : a => //=. Qed. - -Definition wn {n} (a : PTm n) := exists b, rtc RPar'.R a b /\ nf b. -Definition wne {n} (a : PTm n) := exists b, rtc RPar'.R a b /\ ne b. - -(* Weakly neutral implies weakly normal *) -Lemma wne_wn n a : @wne n a -> wn a. -Proof. sfirstorder use:ne_nf. Qed. - -(* Normal implies weakly normal *) -Lemma nf_wn n v : @nf n v -> wn v. -Proof. sfirstorder ctrs:rtc. Qed. - -Lemma nf_refl n (a b : PTm n) (h : RPar'.R a b) : (nf a -> b = a) /\ (ne a -> b = a). -Proof. - elim : a b /h => //=; solve [hauto b:on]. -Qed. - -Lemma ne_nf_ren n m (a : PTm n) (ξ : fin n -> fin m) : - (ne a <-> ne (ren_PTm ξ a)) /\ (nf a <-> nf (ren_PTm ξ a)). -Proof. - move : m ξ. elim : n / a => //=; solve [hauto b:on]. -Qed. - -Lemma wne_app n (a b : PTm n) : - wne a -> wn b -> wne (PApp a b). -Proof. - move => [a0 [? ?]] [b0 [? ?]]. - exists (PApp a0 b0). hauto b:on drew:off use:RPars'.AppCong. -Qed. - -Lemma wn_abs n a (h : wn a) : @wn n (PAbs a). -Proof. - move : h => [v [? ?]]. - exists (PAbs v). - eauto using RPars'.AbsCong. -Qed. - -Lemma wn_pair n (a b : PTm n) : wn a -> wn b -> wn (PPair a b). -Proof. - move => [a0 [? ?]] [b0 [? ?]]. - exists (PPair a0 b0). - hauto lqb:on use:RPars'.PairCong. -Qed. - -Lemma wne_proj n p (a : PTm n) : wne a -> wne (PProj p a). -Proof. - move => [a0 [? ?]]. - exists (PProj p a0). hauto lqb:on use:RPars'.ProjCong. -Qed. - -Create HintDb nfne. -#[export]Hint Resolve nf_wn ne_nf wne_wn nf_refl : nfne. - -Lemma ne_nf_antiren n m (a : PTm n) (ρ : fin n -> PTm m) : - (forall i, var_or_const (ρ i)) -> - (ne (subst_PTm ρ a) -> ne a) /\ (nf (subst_PTm ρ a) -> nf a). -Proof. - move : m ρ. elim : n / a => //; - hauto b:on drew:off use:RPar.var_or_const_up. -Qed. - -Lemma wn_antirenaming n m a (ρ : fin n -> PTm m) : - (forall i, var_or_const (ρ i)) -> - wn (subst_PTm ρ a) -> wn a. -Proof. - rewrite /wn => hρ. - move => [v [rv nfv]]. - move /RPars'.antirenaming : rv. - move /(_ hρ) => [b [hb ?]]. subst. - exists b. split => //=. - move : nfv. - by eapply ne_nf_antiren. -Qed. - -Lemma ext_wn n (a : PTm n) : - wn (PApp a PBot) -> - wn a. -Proof. - move E : (PApp a (PBot)) => a0 [v [hr hv]]. - move : a E. - move : hv. - elim : a0 v / hr. - - hauto q:on inv:PTm ctrs:rtc b:on db: nfne. - - move => a0 a1 a2 hr0 hr1 ih hnfa2. - move /(_ hnfa2) in ih. - move => a. - case : a0 hr0=>// => b0 b1. - elim /RPar'.inv=>// _. - + move => a0 a3 b2 b3 ? ? [? ?] ? [? ?]. subst. - have ? : b3 = (PBot) by hauto lq:on inv:RPar'.R. subst. - suff : wn (PAbs a3) by hauto lq:on ctrs:RPar'.R, rtc unfold:wn. - have : wn (subst_PTm (scons (PBot) VarPTm) a3) by sfirstorder. - move => h. apply wn_abs. - move : h. apply wn_antirenaming. - hauto lq:on rew:off inv:nat. - + hauto q:on inv:RPar'.R ctrs:rtc b:on. -Qed. - -Module Join. - Lemma ProjCong p n (a0 a1 : PTm n) : - join a0 a1 -> - join (PProj p a0) (PProj p a1). - Proof. hauto lq:on use:Pars.ProjCong unfold:join. Qed. - - Lemma PairCong n (a0 a1 b0 b1 : PTm n) : - join a0 a1 -> - join b0 b1 -> - join (PPair a0 b0) (PPair a1 b1). - Proof. hauto lq:on use:Pars.PairCong unfold:join. Qed. - - Lemma AppCong n (a0 a1 b0 b1 : PTm n) : - join a0 a1 -> - join b0 b1 -> - join (PApp a0 b0) (PApp a1 b1). - Proof. hauto lq:on use:Pars.AppCong. Qed. - - Lemma AbsCong n (a b : PTm (S n)) : - join a b -> - join (PAbs a) (PAbs b). - Proof. hauto lq:on use:Pars.AbsCong. Qed. - - Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) : - join a b -> join (ren_PTm ξ a) (ren_PTm ξ b). - Proof. - induction 1; hauto lq:on use:Pars.renaming. - Qed. - - Lemma weakening n (a b : PTm n) : - join a b -> join (ren_PTm shift a) (ren_PTm shift b). - Proof. - apply renaming. - Qed. - - Lemma FromPar n (a b : PTm n) : - Par.R a b -> - join a b. - Proof. - hauto lq:on ctrs:rtc use:rtc_once. - Qed. -End Join. - -Lemma abs_eq n a (b : PTm n) : - join (PAbs a) b <-> join a (PApp (ren_PTm shift b) (VarPTm var_zero)). -Proof. - split. - - move => /Join.weakening h. - have {h} : join (PApp (ren_PTm shift (PAbs a)) (VarPTm var_zero)) (PApp (ren_PTm shift b) (VarPTm var_zero)) - by hauto l:on use:Join.AppCong, join_refl. - simpl. - move => ?. apply : join_transitive; eauto. - apply join_symmetric. apply Join.FromPar. - apply : Par.AppAbs'; eauto using Par.refl. by asimpl. - - move /Join.AbsCong. - move /join_transitive. apply. - apply join_symmetric. apply Join.FromPar. apply Par.AppEta. apply Par.refl. -Qed. - -Lemma pair_eq n (a0 a1 b : PTm n) : - join (PPair a0 a1) b <-> join a0 (PProj PL b) /\ join a1 (PProj PR b). -Proof. - split. - - move => h. - have /Join.ProjCong {}h := h. - have h0 : forall p, join (if p is PL then a0 else a1) (PProj p (PPair a0 a1)) - by hauto lq:on use:join_symmetric, Join.FromPar, Par.ProjPair', Par.refl. - hauto lq:on rew:off use:join_transitive, join_symmetric. - - move => [h0 h1]. - move : h0 h1. - move : Join.PairCong; repeat move/[apply]. - move /join_transitive. apply. apply join_symmetric. - apply Join.FromPar. hauto lq:on ctrs:Par.R use:Par.refl. -Qed. - -Lemma join_pair_inj n (a0 a1 b0 b1 : PTm n) : - join (PPair a0 a1) (PPair b0 b1) <-> join a0 b0 /\ join a1 b1. -Proof. - split; last by hauto lq:on use:Join.PairCong. - move /pair_eq => [h0 h1]. - have : join (PProj PL (PPair b0 b1)) b0 by hauto lq:on use:Join.FromPar, Par.refl, Par.ProjPair'. - have : join (PProj PR (PPair b0 b1)) b1 by hauto lq:on use:Join.FromPar, Par.refl, Par.ProjPair'. - eauto using join_transitive. -Qed. + best ctrs:Par.R inv:EPar.R,RPar.R use:EPar_Par, RPar_Par. diff --git a/theories/logrel.v b/theories/logrel.v deleted file mode 100644 index b9c854d..0000000 --- a/theories/logrel.v +++ /dev/null @@ -1,1036 +0,0 @@ -Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax fp_red compile. -From Hammer Require Import Tactics. -From Equations Require Import Equations. -Require Import ssreflect ssrbool. -Require Import Logic.PropExtensionality (propositional_extensionality). -From stdpp Require Import relations (rtc(..), rtc_subrel). -Import Psatz. - -Definition ProdSpace {n} (PA : Tm n -> Prop) - (PF : Tm n -> (Tm n -> Prop) -> Prop) b : Prop := - forall a PB, PA a -> PF a PB -> PB (App b a). - -Definition SumSpace {n} (PA : Tm n -> Prop) - (PF : Tm n -> (Tm n -> Prop) -> Prop) t : Prop := - wne t \/ exists a b, rtc RPar'.R t (Pair a b) /\ PA a /\ (forall PB, PF a PB -> PB b). - -Definition BindSpace {n} p := if p is TPi then @ProdSpace n else SumSpace. - -Reserved Notation "⟦ A ⟧ i ;; I ↘ S" (at level 70). -Inductive InterpExt {n} i (I : nat -> Tm n -> Prop) : Tm n -> (Tm n -> Prop) -> Prop := -| InterpExt_Ne A : - ne A -> - ⟦ A ⟧ i ;; I ↘ wne -| InterpExt_Bind p A B PA PF : - ⟦ A ⟧ i ;; I ↘ PA -> - (forall a, PA a -> exists PB, PF a PB) -> - (forall a PB, PF a PB -> ⟦ subst_Tm (scons a VarTm) B ⟧ i ;; I ↘ PB) -> - ⟦ TBind p A B ⟧ i ;; I ↘ BindSpace p PA PF - -| InterpExt_Univ j : - j < i -> - ⟦ Univ j ⟧ i ;; I ↘ (I j) - -| InterpExt_Step A A0 PA : - RPar'.R A A0 -> - ⟦ A0 ⟧ i ;; I ↘ PA -> - ⟦ A ⟧ i ;; I ↘ PA -where "⟦ A ⟧ i ;; I ↘ S" := (InterpExt i I A S). - -Lemma InterpExt_Univ' n i I j (PF : Tm n -> Prop) : - PF = I j -> - j < i -> - ⟦ Univ j ⟧ i ;; I ↘ PF. -Proof. hauto lq:on ctrs:InterpExt. Qed. - -Infix " (Tm n -> Prop) -> Prop by wf i lt := - InterpUnivN n i := @InterpExt n i - (fun j A => - match j exists PA, InterpUnivN n j A PA - | right _ => False - end). -Arguments InterpUnivN {n}. - -Lemma InterpExt_lt_impl n i I I' A (PA : Tm n -> Prop) : - (forall j, j < i -> I j = I' j) -> - ⟦ A ⟧ i ;; I ↘ PA -> - ⟦ A ⟧ i ;; I' ↘ PA. -Proof. - move => hI h. - elim : A PA /h. - - hauto q:on ctrs:InterpExt. - - hauto lq:on rew:off ctrs:InterpExt. - - hauto q:on ctrs:InterpExt. - - hauto lq:on ctrs:InterpExt. -Qed. - -Lemma InterpExt_lt_eq n i I I' A (PA : Tm n -> Prop) : - (forall j, j < i -> I j = I' j) -> - ⟦ A ⟧ i ;; I ↘ PA = - ⟦ A ⟧ i ;; I' ↘ PA. -Proof. - move => hI. apply propositional_extensionality. - have : forall j, j < i -> I' j = I j by sfirstorder. - firstorder using InterpExt_lt_impl. -Qed. - -Notation "⟦ A ⟧ i ↘ S" := (InterpUnivN i A S) (at level 70). - -Lemma InterpUnivN_nolt n i : - @InterpUnivN n i = @InterpExt n i (fun j (A : Tm n) => exists PA, ⟦ A ⟧ j ↘ PA). -Proof. - simp InterpUnivN. - extensionality A. extensionality PA. - set I0 := (fun _ => _). - set I1 := (fun _ => _). - apply InterpExt_lt_eq. - hauto q:on. -Qed. - -#[export]Hint Rewrite @InterpUnivN_nolt : InterpUniv. - -Lemma RPar_substone n (a b : Tm (S n)) (c : Tm n): - RPar'.R a b -> RPar'.R (subst_Tm (scons c VarTm) a) (subst_Tm (scons c VarTm) b). - Proof. hauto l:on inv:option use:RPar'.substing, RPar'.refl. Qed. - -Lemma InterpExt_Bind_inv n p i I (A : Tm n) B P - (h : ⟦ TBind p A B ⟧ i ;; I ↘ P) : - exists (PA : Tm n -> Prop) (PF : Tm n -> (Tm n -> Prop) -> Prop), - ⟦ A ⟧ i ;; I ↘ PA /\ - (forall a, PA a -> exists PB, PF a PB) /\ - (forall a PB, PF a PB -> ⟦ subst_Tm (scons a VarTm) B ⟧ i ;; I ↘ PB) /\ - P = BindSpace p PA PF. -Proof. - move E : (TBind p A B) h => T h. - move : A B E. - elim : T P / h => //. - - move => //= *. scongruence. - - hauto l:on. - - move => A A0 PA hA hA0 hPi A1 B ?. subst. - elim /RPar'.inv : hA => //= _ p0 A2 A3 B0 B1 hA1 hB0 [*]. subst. - hauto lq:on ctrs:InterpExt use:RPar_substone. -Qed. - -Lemma InterpExt_Ne_inv n i A I P - (h : ⟦ A : Tm n ⟧ i ;; I ↘ P) : - ne A -> - P = wne. -Proof. - elim : A P / h => //=. - qauto l:on ctrs:prov inv:prov use:nf_refl. -Qed. - -Lemma InterpExt_Univ_inv n i I j P - (h : ⟦ Univ j : Tm n ⟧ i ;; I ↘ P) : - P = I j /\ j < i. -Proof. - move : h. - move E : (Univ j) => T h. move : j E. - elim : T P /h => //. - - move => //= *. scongruence. - - hauto l:on. - - hauto lq:on rew:off inv:RPar'.R. -Qed. - -Lemma InterpExt_Bind_nopf n p i I (A : Tm n) B PA : - ⟦ A ⟧ i ;; I ↘ PA -> - (forall a, PA a -> exists PB, ⟦ subst_Tm (scons a VarTm) B ⟧ i ;; I ↘ PB) -> - ⟦ TBind p A B ⟧ i ;; I ↘ (BindSpace p PA (fun a PB => ⟦ subst_Tm (scons a VarTm) B ⟧ i ;; I ↘ PB)). -Proof. - move => h0 h1. apply InterpExt_Bind =>//. -Qed. - -Lemma InterpUnivN_Fun_nopf n p i (A : Tm n) B PA : - ⟦ A ⟧ i ↘ PA -> - (forall a, PA a -> exists PB, ⟦ subst_Tm (scons a VarTm) B ⟧ i ↘ PB) -> - ⟦ TBind p A B ⟧ i ↘ (BindSpace p PA (fun a PB => ⟦ subst_Tm (scons a VarTm) B ⟧ i ↘ PB)). -Proof. - hauto l:on use:InterpExt_Bind_nopf rew:db:InterpUniv. -Qed. - -Lemma InterpExt_cumulative n i j I (A : Tm n) PA : - i <= j -> - ⟦ A ⟧ i ;; I ↘ PA -> - ⟦ A ⟧ j ;; I ↘ PA. -Proof. - move => h h0. - elim : A PA /h0; - hauto l:on ctrs:InterpExt solve+:(by lia). -Qed. - -Lemma InterpUnivN_cumulative n i (A : Tm n) PA : - ⟦ A ⟧ i ↘ PA -> forall j, i <= j -> - ⟦ A ⟧ j ↘ PA. -Proof. - hauto l:on rew:db:InterpUniv use:InterpExt_cumulative. -Qed. - -Lemma InterpExt_preservation n i I (A : Tm n) B P (h : InterpExt i I A P) : - RPar'.R A B -> - ⟦ B ⟧ i ;; I ↘ P. -Proof. - move : B. - elim : A P / h; auto. - - hauto lq:on use:nf_refl ctrs:InterpExt. - - move => p A B PA PF hPA ihPA hPB hPB' ihPB T hT. - elim /RPar'.inv : hT => //. - move => hPar p0 A0 A1 B0 B1 h0 h1 [? ?] ? ?; subst. - apply InterpExt_Bind; auto => a PB hPB0. - apply : ihPB; eauto. - sfirstorder use:RPar'.cong, RPar'.refl. - - hauto lq:on inv:RPar'.R ctrs:InterpExt. - - move => A B P h0 h1 ih1 C hC. - have [D [h2 h3]] := RPar'_diamond _ _ _ _ h0 hC. - hauto lq:on ctrs:InterpExt. -Qed. - -Lemma InterpUnivN_preservation n i (A : Tm n) B P (h : ⟦ A ⟧ i ↘ P) : - RPar'.R A B -> - ⟦ B ⟧ i ↘ P. -Proof. hauto l:on rew:db:InterpUnivN use: InterpExt_preservation. Qed. - -Lemma InterpExt_back_preservation_star n i I (A : Tm n) B P (h : ⟦ B ⟧ i ;; I ↘ P) : - rtc RPar'.R A B -> - ⟦ A ⟧ i ;; I ↘ P. -Proof. induction 1; hauto l:on ctrs:InterpExt. Qed. - -Lemma InterpExt_preservation_star n i I (A : Tm n) B P (h : ⟦ A ⟧ i ;; I ↘ P) : - rtc RPar'.R A B -> - ⟦ B ⟧ i ;; I ↘ P. -Proof. induction 1; hauto l:on use:InterpExt_preservation. Qed. - -Lemma InterpUnivN_preservation_star n i (A : Tm n) B P (h : ⟦ A ⟧ i ↘ P) : - rtc RPar'.R A B -> - ⟦ B ⟧ i ↘ P. -Proof. hauto l:on rew:db:InterpUnivN use:InterpExt_preservation_star. Qed. - -Lemma InterpUnivN_back_preservation_star n i (A : Tm n) B P (h : ⟦ B ⟧ i ↘ P) : - rtc RPar'.R A B -> - ⟦ A ⟧ i ↘ P. -Proof. hauto l:on rew:db:InterpUnivN use:InterpExt_back_preservation_star. Qed. - -Function hfb {n} (A : Tm n) := - match A with - | TBind _ _ _ => true - | Univ _ => true - | _ => ne A - end. - -Inductive hfb_case {n} : Tm n -> Prop := -| hfb_bind p A B : - hfb_case (TBind p A B) -| hfb_univ i : - hfb_case (Univ i) -| hfb_ne A : - ne A -> - hfb_case A. - -Derive Dependent Inversion hfb_inv with (forall n (a : Tm n), hfb_case a) Sort Prop. - -Lemma ne_hfb {n} (A : Tm n) : ne A -> hfb A. -Proof. case : A => //=. Qed. - -Lemma hfb_caseP {n} (A : Tm n) : hfb A -> hfb_case A. -Proof. hauto lq:on ctrs:hfb_case inv:Tm use:ne_hfb. Qed. - -Lemma InterpExtInv n i I (A : Tm n) PA : - ⟦ A ⟧ i ;; I ↘ PA -> - exists B, hfb B /\ rtc RPar'.R A B /\ ⟦ B ⟧ i ;; I ↘ PA. -Proof. - move => h. elim : A PA /h. - - hauto q:on ctrs:InterpExt, rtc use:ne_hfb. - - move => p A B PA PF hPA _ hPF hPF0 _. - exists (TBind p A B). repeat split => //=. - apply rtc_refl. - hauto l:on ctrs:InterpExt. - - move => j ?. exists (Univ j). - hauto l:on ctrs:InterpExt. - - hauto lq:on ctrs:rtc. -Qed. - -Lemma RPar'_Par n (A B : Tm n) : - RPar'.R A B -> - Par.R A B. -Proof. induction 1; hauto lq:on ctrs:Par.R. Qed. - -Lemma RPar's_Pars n (A B : Tm n) : - rtc RPar'.R A B -> - rtc Par.R A B. -Proof. hauto lq:on use:RPar'_Par, rtc_subrel. Qed. - -Lemma RPar's_join n (A B : Tm n) : - rtc RPar'.R A B -> Join.R A B. -Proof. - rewrite /Join.R => h. - have {}h : rtc RPar'.R (Compile.F A) (Compile.F B) by eauto using compile_rpars. - rewrite /join. eauto using RPar's_Pars, rtc_refl. -Qed. - -Lemma bindspace_iff n p (PA : Tm n -> Prop) PF PF0 b : - (forall (a : Tm n) (PB PB0 : Tm n -> Prop), PF a PB -> PF0 a PB0 -> PB = PB0) -> - (forall a, PA a -> exists PB, PF a PB) -> - (forall a, PA a -> exists PB0, PF0 a PB0) -> - (BindSpace p PA PF b <-> BindSpace p PA PF0 b). -Proof. - rewrite /BindSpace => h hPF hPF0. - case : p => /=. - - rewrite /ProdSpace. - split. - move => h1 a PB ha hPF'. - specialize hPF with (1 := ha). - specialize hPF0 with (1 := ha). - sblast. - move => ? a PB ha. - specialize hPF with (1 := ha). - specialize hPF0 with (1 := ha). - sblast. - - rewrite /SumSpace. - hauto lq:on rew:off. -Qed. - -Lemma ne_prov_inv n (a : Tm n) : - ne a -> (exists i, prov (VarTm i) a) \/ prov Bot a. -Proof. - elim : n /a => //=. - - hauto lq:on ctrs:prov. - - hauto lq:on rew:off ctrs:prov b:on. - - hauto lq:on ctrs:prov. - - move => n k. - have : @prov n Bot Bot by auto using P_Bot. - eauto. -Qed. - -Lemma ne_pars_inv n (a b : Tm n) : - ne a -> rtc Par.R a b -> (exists i, prov (VarTm i) b) \/ prov Bot b. -Proof. - move /ne_prov_inv. - sfirstorder use:prov_pars. -Qed. - -Lemma ne_pars_extract n (a b : Tm n) : - ne a -> rtc Par.R a b -> (exists i, extract b = (VarTm i)) \/ extract b = Bot. -Proof. hauto lq:on rew:off use:ne_pars_inv, prov_extract. Qed. - -Lemma const_pars_extract n k b : - rtc Par.R (Const k : Tm n) b -> extract b = Const k. -Proof. hauto l:on use:pars_const_inv, prov_extract. Qed. - -Lemma compile_ne n (a : Tm n) : - ne a = ne (Compile.F a) /\ nf a = nf (Compile.F a). -Proof. - elim : n / a => //=; sfirstorder b:on. -Qed. - -Lemma join_univ_pi_contra n p (A : Tm n) B i : - Join.R (TBind p A B) (Univ i) -> False. -Proof. - rewrite /Join.R /=. - rewrite !pair_eq. - move => [[h _ ]_ ]. - move : h => [C [h0 h1]]. - have ? : extract C = Const p by hauto l:on use:pars_const_inv. - have h : prov (Univ i : Tm n) (Proj PL (Proj PL (Univ i))) by hauto lq:on ctrs:prov. - have {h} : prov (Univ i) C by eauto using prov_pars. - move /prov_extract=>/=. congruence. -Qed. - -Lemma join_bind_ne_contra n p (A : Tm n) B C : - ne C -> - Join.R (TBind p A B) C -> False. -Proof. - rewrite /Join.R. move => hC /=. - rewrite !pair_eq. - move => [[[D [h0 h1]] _] _]. - have {}hC : ne (Compile.F C) by hauto lq:on use:compile_ne. - have {}hC : ne (Proj PL (Proj PL (Compile.F C))) by scongruence. - have : (exists i, extract D = (VarTm i)) \/ extract D = Bot by eauto using ne_pars_extract. - have : extract D = Const p by eauto using const_pars_extract. - sfirstorder. -Qed. - -Lemma join_univ_ne_contra n i C : - ne C -> - Join.R (Univ i : Tm n) C -> False. -Proof. - move => hC [D [h0 h1]]. - move /pars_univ_inv : h0 => ?. - have : (exists i, extract D = (VarTm i)) \/ exists k, extract D =(Const k) by hauto q:on use:ne_pars_extract, compile_ne. - sfirstorder. -Qed. - -#[export]Hint Resolve join_univ_ne_contra join_bind_ne_contra join_univ_pi_contra Join.symmetric Join.transitive : join. - -Lemma InterpExt_Join n i I (A B : Tm n) PA PB : - ⟦ A ⟧ i ;; I ↘ PA -> - ⟦ B ⟧ i ;; I ↘ PB -> - Join.R A B -> - PA = PB. -Proof. - move => h. move : B PB. elim : A PA /h. - - move => A hA B PB /InterpExtInv. - move => [B0 []]. - move /hfb_caseP. elim/hfb_inv => _. - + move => p A0 B1 ? [/RPar's_join h0 h1] h2. subst. exfalso. - eauto with join. - + move => ? ? [/RPar's_join *]. subst. exfalso. - eauto with join. - + hauto lq:on use:InterpExt_Ne_inv. - - move => p A B PA PF hPA ihPA hTot hRes ihPF U PU /InterpExtInv. - move => [B0 []]. - move /hfb_caseP. - elim /hfb_inv => _. - rename B0 into B00. - + move => p0 A0 B0 ? [hr hPi]. subst. - move /InterpExt_Bind_inv : hPi. - move => [PA0][PF0][hPA0][hTot0][hRes0]?. subst. - move => hjoin. - have{}hr : Join.R U (TBind p0 A0 B0) by auto using RPar's_join. - have hj : Join.R (TBind p A B) (TBind p0 A0 B0) by eauto using Join.transitive. - have {hj} : p0 = p /\ Join.R A A0 /\ Join.R B B0 by hauto l:on use:Join.BindInj. - move => [? [h0 h1]]. subst. - have ? : PA0 = PA by hauto l:on. subst. - rewrite /ProdSpace. - extensionality b. - apply propositional_extensionality. - apply bindspace_iff; eauto. - move => a PB PB0 hPB hPB0. - apply : ihPF; eauto. - rewrite /Join.R. - rewrite -!Compile.substing. - hauto l:on use:join_substing. - + move => j ?. subst. - move => [h0 h1] h. - have ? : Join.R U (Univ j) by eauto using RPar's_join. - have : Join.R (TBind p A B) (Univ j) by eauto using Join.transitive. - move => ?. exfalso. - hauto l: on use: join_univ_pi_contra. - + move => A0 ? ? [/RPar's_join ?]. subst. - move => _ ?. exfalso. eauto with join. - - move => j ? B PB /InterpExtInv. - move => [? []]. move/hfb_caseP. - elim /hfb_inv => //= _. - + move => p A0 B0 _ []. - move /RPar's_join => *. - exfalso. eauto with join. - + move => m _ [/RPar's_join h0 + h1]. - have /join_univ_inj {h0 h1} ? : Join.R (Univ j : Tm n) (Univ m) by eauto using Join.transitive. - subst. - move /InterpExt_Univ_inv. firstorder. - + move => A ? ? [/RPar's_join] *. subst. exfalso. eauto with join. - - move => A A0 PA h. - have /Join.symmetric {}h : Join.R A A0 by hauto lq:on ctrs:rtc use:RPar's_join, relations.rtc_once. - eauto with join. -Qed. - -Lemma InterpUniv_Join n i (A B : Tm n) PA PB : - ⟦ A ⟧ i ↘ PA -> - ⟦ B ⟧ i ↘ PB -> - Join.R A B -> - PA = PB. -Proof. hauto l:on use:InterpExt_Join rew:db:InterpUniv. Qed. - -Lemma InterpUniv_Bind_inv n p i (A : Tm n) B P - (h : ⟦ TBind p A B ⟧ i ↘ P) : - exists (PA : Tm n -> Prop) (PF : Tm n -> (Tm n -> Prop) -> Prop), - ⟦ A ⟧ i ↘ PA /\ - (forall a, PA a -> exists PB, PF a PB) /\ - (forall a PB, PF a PB -> ⟦ subst_Tm (scons a VarTm) B ⟧ i ↘ PB) /\ - P = BindSpace p PA PF. -Proof. hauto l:on use:InterpExt_Bind_inv rew:db:InterpUniv. Qed. - -Lemma InterpUniv_Univ_inv n i j P - (h : ⟦ Univ j ⟧ i ↘ P) : - P = (fun (A : Tm n) => exists PA, ⟦ A ⟧ j ↘ PA) /\ j < i. -Proof. hauto l:on use:InterpExt_Univ_inv rew:db:InterpUniv. Qed. - -Lemma InterpExt_Functional n i I (A B : Tm n) PA PB : - ⟦ A ⟧ i ;; I ↘ PA -> - ⟦ A ⟧ i ;; I ↘ PB -> - PA = PB. -Proof. hauto use:InterpExt_Join, join_refl. Qed. - -Lemma InterpUniv_Functional n i (A : Tm n) PA PB : - ⟦ A ⟧ i ↘ PA -> - ⟦ A ⟧ i ↘ PB -> - PA = PB. -Proof. hauto use:InterpExt_Functional rew:db:InterpUniv. Qed. - -Lemma InterpUniv_Join' n i j (A B : Tm n) PA PB : - ⟦ A ⟧ i ↘ PA -> - ⟦ B ⟧ j ↘ PB -> - Join.R A B -> - PA = PB. -Proof. - have [? ?] : i <= max i j /\ j <= max i j by lia. - move => hPA hPB. - have : ⟦ A ⟧ (max i j) ↘ PA by eauto using InterpUnivN_cumulative. - have : ⟦ B ⟧ (max i j) ↘ PB by eauto using InterpUnivN_cumulative. - eauto using InterpUniv_Join. -Qed. - -Lemma InterpUniv_Functional' n i j A PA PB : - ⟦ A : Tm n ⟧ i ↘ PA -> - ⟦ A ⟧ j ↘ PB -> - PA = PB. -Proof. - hauto l:on use:InterpUniv_Join', join_refl. -Qed. - -Lemma InterpExt_Bind_inv_nopf i n I p A B P (h : ⟦TBind p A B ⟧ i ;; I ↘ P) : - exists (PA : Tm n -> Prop), - ⟦ A ⟧ i ;; I ↘ PA /\ - (forall a, PA a -> exists PB, ⟦ subst_Tm (scons a VarTm) B ⟧ i ;; I ↘ PB) /\ - P = BindSpace p PA (fun a PB => ⟦ subst_Tm (scons a VarTm) B ⟧ i ;; I ↘ PB). -Proof. - move /InterpExt_Bind_inv : h. intros (PA & PF & hPA & hPF & hPF' & ?); subst. - exists PA. repeat split => //. - - sfirstorder. - - extensionality b. - case : p => /=. - + extensionality a. - extensionality PB. - extensionality ha. - apply propositional_extensionality. - split. - * hecrush use:InterpExt_Functional. - * sfirstorder. - + rewrite /SumSpace. apply propositional_extensionality. - split; hauto q:on use:InterpExt_Functional. -Qed. - -Lemma InterpUniv_Bind_inv_nopf n i p A B P (h : ⟦TBind p A B ⟧ i ↘ P) : - exists (PA : Tm n -> Prop), - ⟦ A ⟧ i ↘ PA /\ - (forall a, PA a -> exists PB, ⟦ subst_Tm (scons a VarTm) B ⟧ i ↘ PB) /\ - P = BindSpace p PA (fun a PB => ⟦ subst_Tm (scons a VarTm) B ⟧ i ↘ PB). -Proof. hauto l:on use:InterpExt_Bind_inv_nopf rew:db:InterpUniv. Qed. - -Lemma InterpExt_back_clos n i I (A : Tm n) PA : - (forall j, j < i -> forall a b, (RPar'.R a b) -> I j b -> I j a) -> - ⟦ A ⟧ i ;; I ↘ PA -> - forall a b, (RPar'.R a b) -> - PA b -> PA a. -Proof. - move => hI h. - elim : A PA /h. - - hauto q:on ctrs:rtc unfold:wne. - - move => p A B PA PF hPA ihPA hTot hRes ihPF a b hr. - case : p => //=. - + have : forall b0 b1 a, RPar'.R b0 b1 -> RPar'.R (App b0 a) (App b1 a) - by hauto lq:on ctrs:RPar'.R use:RPar'.refl. - hauto lq:on rew:off unfold:ProdSpace. - + hauto lq:on ctrs:rtc unfold:SumSpace. - - eauto. - - eauto. -Qed. - -Lemma InterpExt_back_clos_star n i I (A : Tm n) PA : - (forall j, j < i -> forall a b, (RPar'.R a b) -> I j b -> I j a) -> - ⟦ A ⟧ i ;; I ↘ PA -> - forall a b, (rtc RPar'.R a b) -> - PA b -> PA a. -Proof. induction 3; hauto l:on use:InterpExt_back_clos. Qed. - -Lemma InterpUniv_back_clos n i (A : Tm n) PA : - ⟦ A ⟧ i ↘ PA -> - forall a b, (RPar'.R a b) -> - PA b -> PA a. -Proof. - simp InterpUniv. - apply InterpExt_back_clos. - hauto lq:on ctrs:rtc use:InterpUnivN_back_preservation_star. -Qed. - -Lemma InterpUniv_back_clos_star n i (A : Tm n) PA : - ⟦ A ⟧ i ↘ PA -> - forall a b, rtc RPar'.R a b -> - PA b -> PA a. -Proof. - move => h a b. - induction 1=> //. - hauto lq:on use:InterpUniv_back_clos. -Qed. - -Lemma pars'_wn {n} a b : - rtc RPar'.R a b -> - @wn n b -> - wn a. -Proof. sfirstorder unfold:wn use:@relations.rtc_transitive. Qed. - -(* P identifies a set of "reducibility candidates" *) -Definition CR {n} (P : Tm n -> Prop) := - (forall a, P a -> wn a) /\ - (forall a, ne a -> P a). - -Lemma adequacy_ext i n I A PA - (hI0 : forall j, j < i -> forall a b, (RPar'.R a b) -> I j b -> I j a) - (hI : forall j, j < i -> CR (I j)) - (h : ⟦ A : Tm n ⟧ i ;; I ↘ PA) : - CR PA /\ wn A. -Proof. - elim : A PA / h. - - hauto unfold:wne use:wne_wn. - - move => p A B PA PF hA hPA hTot hRes ihPF. - rewrite /CR. - have hb : PA Bot by firstorder. - repeat split. - + case : p => /=. - * qauto l:on use:ext_wn unfold:ProdSpace, CR. - * rewrite /SumSpace => a []; first by eauto with nfne. - move => [q0][q1]*. - have : wn q0 /\ wn q1 by hauto q:on. - qauto l:on use:wn_pair, pars'_wn. - + case : p => /=. - * rewrite /ProdSpace. - move => a ha c PB hc hPB. - have hc' : wn c by sfirstorder. - have : wne (App a c) by hauto lq:on use:wne_app ctrs:rtc. - have h : (forall a, ne a -> PB a) by sfirstorder. - suff : (forall a, wne a -> PB a) by hauto l:on. - move => a0 [a1 [h0 h1]]. - eapply InterpExt_back_clos_star with (b := a1); eauto. - * rewrite /SumSpace. - move => a ha. left. - sfirstorder ctrs:rtc. - + have wnA : wn A by firstorder. - apply wn_bind => //. - apply wn_antirenaming with (ρ := scons Bot VarTm);first by hauto q:on inv:option. - hauto lq:on. - - hauto l:on. - - hauto lq:on rew:off ctrs:rtc. -Qed. - -Lemma adequacy i n A PA - (h : ⟦ A : Tm n ⟧ i ↘ PA) : - CR PA /\ wn A. -Proof. - move : i A PA h. - elim /Wf_nat.lt_wf_ind => i ih A PA. - simp InterpUniv. - apply adequacy_ext. - hauto lq:on ctrs:rtc use:InterpUnivN_back_preservation_star. - hauto l:on use:InterpExt_Ne rew:db:InterpUniv. -Qed. - -Lemma adequacy_wne i n A PA a : ⟦ A : Tm n ⟧ i ↘ PA -> wne a -> PA a. -Proof. qauto l:on use:InterpUniv_back_clos_star, adequacy unfold:CR. Qed. - -Lemma adequacy_wn i n A PA (h : ⟦ A : Tm n ⟧ i ↘ PA) a : PA a -> wn a. -Proof. hauto q:on use:adequacy. Qed. - -Definition ρ_ok {n} (Γ : fin n -> Tm n) (ρ : fin n -> Tm 0) := forall i k PA, - ⟦ subst_Tm ρ (Γ i) ⟧ k ↘ PA -> PA (ρ i). - -Definition SemWt {n} Γ (a A : Tm n) := forall ρ, ρ_ok Γ ρ -> exists k PA, ⟦ subst_Tm ρ A ⟧ k ↘ PA /\ PA (subst_Tm ρ a). -Notation "Γ ⊨ a ∈ A" := (SemWt Γ a A) (at level 70). - -(* Semantic context wellformedness *) -Definition SemWff {n} Γ := forall (i : fin n), exists j, Γ ⊨ Γ i ∈ Univ j. -Notation "⊨ Γ" := (SemWff Γ) (at level 70). - -Lemma ρ_ok_bot n (Γ : fin n -> Tm n) : - ρ_ok Γ (fun _ => Bot). -Proof. - rewrite /ρ_ok. - hauto q:on use:adequacy. -Qed. - -Lemma ρ_ok_cons n i (Γ : fin n -> Tm n) ρ a PA A : - ⟦ subst_Tm ρ A ⟧ i ↘ PA -> PA a -> - ρ_ok Γ ρ -> - ρ_ok (funcomp (ren_Tm shift) (scons A Γ)) (scons a ρ). -Proof. - move => h0 h1 h2. - rewrite /ρ_ok. - move => j. - destruct j as [j|]. - - move => m PA0. asimpl => ?. - asimpl. - firstorder. - - move => m PA0. asimpl => h3. - have ? : PA0 = PA by eauto using InterpUniv_Functional'. - by subst. -Qed. - -Definition renaming_ok {n m} (Γ : fin n -> Tm n) (Δ : fin m -> Tm m) (ξ : fin m -> fin n) := - forall (i : fin m), ren_Tm ξ (Δ i) = Γ (ξ i). - -Lemma ρ_ok_renaming n m (Γ : fin n -> Tm n) ρ : - forall (Δ : fin m -> Tm m) ξ, - renaming_ok Γ Δ ξ -> - ρ_ok Γ ρ -> - ρ_ok Δ (funcomp ρ ξ). -Proof. - move => Δ ξ hξ hρ. - rewrite /ρ_ok => i m' PA. - rewrite /renaming_ok in hξ. - rewrite /ρ_ok in hρ. - move => h. - rewrite /funcomp. - apply hρ with (k := m'). - move : h. rewrite -hξ. - by asimpl. -Qed. - -Lemma renaming_SemWt {n} Γ a A : - Γ ⊨ a ∈ A -> - forall {m} Δ (ξ : fin n -> fin m), - renaming_ok Δ Γ ξ -> - Δ ⊨ ren_Tm ξ a ∈ ren_Tm ξ A. -Proof. - rewrite /SemWt => h m Δ ξ hξ ρ hρ. - have /h hρ' : (ρ_ok Γ (funcomp ρ ξ)) by eauto using ρ_ok_renaming. - hauto q:on solve+:(by asimpl). -Qed. - -Lemma weakening_Sem n Γ (a : Tm n) A B i - (h0 : Γ ⊨ B ∈ Univ i) - (h1 : Γ ⊨ a ∈ A) : - funcomp (ren_Tm shift) (scons B Γ) ⊨ ren_Tm shift a ∈ ren_Tm shift A. -Proof. - apply : renaming_SemWt; eauto. - hauto lq:on inv:option unfold:renaming_ok. -Qed. - -Lemma SemWt_Wn n Γ (a : Tm n) A : - Γ ⊨ a ∈ A -> - wn a /\ wn A. -Proof. - move => h. - have {}/h := ρ_ok_bot _ Γ => h. - have h0 : wn (subst_Tm (fun _ : fin n => (Bot : Tm 0)) A) by hauto l:on use:adequacy. - have h1 : wn (subst_Tm (fun _ : fin n => (Bot : Tm 0)) a)by hauto l:on use:adequacy_wn. - move {h}. hauto lq:on use:wn_antirenaming. -Qed. - -Lemma SemWt_Univ n Γ (A : Tm n) i : - Γ ⊨ A ∈ Univ i <-> - forall ρ, ρ_ok Γ ρ -> exists S, ⟦ subst_Tm ρ A ⟧ i ↘ S. -Proof. - rewrite /SemWt. - split. - - hauto lq:on rew:off use:InterpUniv_Univ_inv. - - move => /[swap] ρ /[apply]. - move => [PA hPA]. - exists (S i). eexists. - split. - + simp InterpUniv. apply InterpExt_Univ. lia. - + simpl. eauto. -Qed. - -(* Structural laws for Semantic context wellformedness *) -Lemma SemWff_nil : SemWff null. -Proof. case. Qed. - -Lemma SemWff_cons n Γ (A : Tm n) i : - ⊨ Γ -> - Γ ⊨ A ∈ Univ i -> - (* -------------- *) - ⊨ funcomp (ren_Tm shift) (scons A Γ). -Proof. - move => h h0. - move => j. destruct j as [j|]. - - move /(_ j) : h => [k hk]. - exists k. change (Univ k) with (ren_Tm shift (Univ k : Tm n)). - eauto using weakening_Sem. - - hauto q:on use:weakening_Sem. -Qed. - -(* Semantic typing rules *) -Lemma ST_Var n Γ (i : fin n) : - ⊨ Γ -> - Γ ⊨ VarTm i ∈ Γ i. -Proof. - move /(_ i) => [j /SemWt_Univ h]. - rewrite /SemWt => ρ /[dup] hρ {}/h [S hS]. - exists j, S. - asimpl. firstorder. -Qed. - -Lemma ST_Bind n Γ i j p (A : Tm n) (B : Tm (S n)) : - Γ ⊨ A ∈ Univ i -> - funcomp (ren_Tm shift) (scons A Γ) ⊨ B ∈ Univ j -> - Γ ⊨ TBind p A B ∈ Univ (max i j). -Proof. - move => /SemWt_Univ h0 /SemWt_Univ h1. - apply SemWt_Univ => ρ hρ. - move /h0 : (hρ){h0} => [S hS]. - eexists => /=. - have ? : i <= Nat.max i j by lia. - apply InterpUnivN_Fun_nopf. - - eauto using InterpUnivN_cumulative. - - move => *. asimpl. hauto l:on use:InterpUnivN_cumulative, ρ_ok_cons. -Qed. - -Lemma ST_Conv n Γ (a : Tm n) A B i : - Γ ⊨ a ∈ A -> - Γ ⊨ B ∈ Univ i -> - Join.R A B -> - Γ ⊨ a ∈ B. -Proof. - move => ha /SemWt_Univ h h0. - move => ρ hρ. - have {}h0 : Join.R (subst_Tm ρ A) (subst_Tm ρ B) by eauto using Join.substing. - move /ha : (hρ){ha} => [m [PA [h1 h2]]]. - move /h : (hρ){h} => [S hS]. - have ? : PA = S by eauto using InterpUniv_Join'. subst. - eauto. -Qed. - -Lemma ST_Abs n Γ (a : Tm (S n)) A B i : - Γ ⊨ TBind TPi A B ∈ (Univ i) -> - funcomp (ren_Tm shift) (scons A Γ) ⊨ a ∈ B -> - Γ ⊨ Abs a ∈ TBind TPi A B. -Proof. - rename a into b. - move /SemWt_Univ => + hb ρ hρ. - move /(_ _ hρ) => [PPi hPPi]. - exists i, PPi. split => //. - simpl in hPPi. - move /InterpUniv_Bind_inv_nopf : hPPi. - move => [PA [hPA [hTot ?]]]. subst=>/=. - move => a PB ha. asimpl => hPB. - move : ρ_ok_cons (hPA) (hρ) (ha). repeat move/[apply]. - move /hb. - intros (m & PB0 & hPB0 & hPB0'). - replace PB0 with PB in * by hauto l:on use:InterpUniv_Functional'. - apply : InterpUniv_back_clos; eauto. - apply : RPar'.AppAbs'; eauto using RPar'.refl. - by asimpl. -Qed. - -Lemma ST_App n Γ (b a : Tm n) A B : - Γ ⊨ b ∈ TBind TPi A B -> - Γ ⊨ a ∈ A -> - Γ ⊨ App b a ∈ subst_Tm (scons a VarTm) B. -Proof. - move => hf hb ρ hρ. - move /(_ ρ hρ) : hf; intros (i & PPi & hPi & hf). - move /(_ ρ hρ) : hb; intros (j & PA & hPA & hb). - simpl in hPi. - move /InterpUniv_Bind_inv_nopf : hPi. intros (PA0 & hPA0 & hTot & ?). subst. - have ? : PA0 = PA by eauto using InterpUniv_Functional'. subst. - move : hf (hb). move/[apply]. - move : hTot hb. move/[apply]. - asimpl. hauto lq:on. -Qed. - -Lemma ST_Pair n Γ (a b : Tm n) A B i : - Γ ⊨ TBind TSig A B ∈ (Univ i) -> - Γ ⊨ a ∈ A -> - Γ ⊨ b ∈ subst_Tm (scons a VarTm) B -> - Γ ⊨ Pair a b ∈ TBind TSig A B. -Proof. - move /SemWt_Univ => + ha hb ρ hρ. - move /(_ _ hρ) => [PPi hPPi]. - exists i, PPi. split => //. - simpl in hPPi. - move /InterpUniv_Bind_inv_nopf : hPPi. - move => [PA [hPA [hTot ?]]]. subst=>/=. - rewrite /SumSpace. right. - exists (subst_Tm ρ a), (subst_Tm ρ b). - split. - - hauto l:on use:Pars.substing. - - move /ha : (hρ){ha}. - move => [m][PA0][h0]h1. - move /hb : (hρ){hb}. - move => [k][PB][h2]h3. - have ? : PA0 = PA by eauto using InterpUniv_Functional'. subst. - split => // PB0. - move : h2. asimpl => *. - have ? : PB0 = PB by eauto using InterpUniv_Functional'. by subst. -Qed. - -Lemma ST_Proj1 n Γ (a : Tm n) A B : - Γ ⊨ a ∈ TBind TSig A B -> - Γ ⊨ Proj PL a ∈ A. -Proof. - move => h ρ /[dup]hρ {}/h [m][PA][/= /InterpUniv_Bind_inv_nopf h0]h1. - move : h0 => [S][h2][h3]?. subst. - move : h1 => /=. - rewrite /SumSpace. - case; first by hauto lq:on use:adequacy_wne, wne_proj. - move => [a0 [b0 [h4 [h5 h6]]]]. - exists m, S. split => //=. - have {}h4 : rtc RPar'.R (Proj PL (subst_Tm ρ a)) (Proj PL (Pair a0 b0)) by eauto using RPars'.ProjCong. - have ? : RPar'.R (Proj PL (Pair a0 b0)) a0 by hauto l:on use:RPar'.refl, RPar'.ProjPair'. - have : rtc RPar'.R (Proj PL (subst_Tm ρ a)) a0 by eauto using @relations.rtc_r. - move => h. - apply : InterpUniv_back_clos_star; eauto. -Qed. - -Lemma substing_RPar' n m (A : Tm (S n)) ρ (B : Tm m) C : - RPar'.R B C -> - RPar'.R (subst_Tm (scons B ρ) A) (subst_Tm (scons C ρ) A). -Proof. hauto lq:on inv:option use:RPar'.morphing, RPar'.refl. Qed. - -Lemma substing_RPar's n m (A : Tm (S n)) ρ (B : Tm m) C : - rtc RPar'.R B C -> - rtc RPar'.R (subst_Tm (scons B ρ) A) (subst_Tm (scons C ρ) A). -Proof. induction 1; hauto lq:on ctrs:rtc use:substing_RPar'. Qed. - -Lemma ST_Proj2 n Γ (a : Tm n) A B : - Γ ⊨ a ∈ TBind TSig A B -> - Γ ⊨ Proj PR a ∈ subst_Tm (scons (Proj PL a) VarTm) B. -Proof. - move => h ρ hρ. - move : (hρ) => {}/h [m][PA][/= /InterpUniv_Bind_inv_nopf h0]h1. - move : h0 => [S][h2][h3]?. subst. - move : h1 => /=. - rewrite /SumSpace. - case. - - move => h. - have hp : forall p, wne (Proj p (subst_Tm ρ a)) by auto using wne_proj. - have hp0 := hp PL. have hp1 := hp PR => {hp}. - have : S (Proj PL (subst_Tm ρ a)) by hauto q:on use:adequacy_wne. - move /h3 => [PB]. asimpl. hauto lq:on use:adequacy_wne. - - move => [a0 [b0 [h4 [h5 h6]]]]. - specialize h3 with (1 := h5). - move : h3 => [PB hPB]. - have hr : forall p, rtc RPar'.R (Proj p (subst_Tm ρ a)) (Proj p (Pair a0 b0)) by eauto using RPars'.ProjCong. - have hrl : RPar'.R (Proj PL (Pair a0 b0)) a0 by hauto l:on use:RPar'.ProjPair', RPar'.refl. - have hrr : RPar'.R (Proj PR (Pair a0 b0)) b0 by hauto l:on use:RPar'.ProjPair', RPar'.refl. - exists m, PB. - asimpl. split. - + have h : rtc RPar'.R (Proj PL (subst_Tm ρ a)) a0 by eauto using @relations.rtc_r. - have {}h : rtc RPar'.R (subst_Tm (scons (Proj PL (subst_Tm ρ a)) ρ) B) (subst_Tm (scons a0 ρ) B) by eauto using substing_RPar's. - move : hPB. asimpl. - eauto using InterpUnivN_back_preservation_star. - + hauto lq:on use:@relations.rtc_r, InterpUniv_back_clos_star. -Qed. - -Lemma ne_nf_preservation n (a b : Tm n) : ERed.R b a -> (ne a -> ne b) /\ (nf a -> nf b). -Proof. - move => h. elim : n b a /h => //=. - - move => n a. - split => //=. - hauto lqb:on use:ne_nf_ren db:nfne. - - hauto lqb:on db:nfne. - - hauto lqb:on db:nfne. - - hauto lqb:on db:nfne. - - hauto lqb:on db:nfne. - - hauto lqb:on db:nfne. - - hauto lqb:on db:nfne. - - hauto lqb:on db:nfne. - - hauto lqb:on db:nfne. - - hauto lqb:on db:nfne. -Qed. - -Fixpoint size_tm {n} (a : Tm n) := - match a with - | VarTm _ => 1 - | TBind _ A B => 1 + Nat.add (size_tm A) (size_tm B) - | Abs a => 1 + size_tm a - | App a b => 1 + Nat.add (size_tm a) (size_tm b) - | Proj p a => 1 + size_tm a - | Pair a b => 1 + Nat.add (size_tm a) (size_tm b) - | Bot => 1 - | Const _ => 1 - | Univ _ => 1 - end. - -Lemma size_tm_ren n m (ξ : fin n -> fin m) a : size_tm (ren_Tm ξ a) = size_tm a. -Proof. - move : m ξ. elim : n / a => //=; scongruence. -Qed. - -#[export]Hint Rewrite size_tm_ren : size_tm. - -Lemma size_η_lt n (a b : Tm n) : - ERed.R b a -> - size_tm b < size_tm a. -Proof. - move => h. elim : b a / h => //=; hauto l:on rew:db:size_tm. -Qed. - -Lemma ered_local_confluence n (a b c : Tm n) : - ERed.R b a -> - ERed.R c a -> - exists d, rtc ERed.R d b /\ rtc ERed.R d c. -Proof. - move => h. move : c. - elim : n b a / h => n. - - move => a c. - elim /ERed.inv => //= _. - + move => ? ? [*]. subst. - have : subst_Tm (scons Bot VarTm) (ren_Tm shift c) = (subst_Tm (scons Bot VarTm) (ren_Tm shift a)) - by congruence. - asimpl => ?. subst. - eauto using rtc_refl. - + move => a0 a1 ha ? [*]. subst. - elim /ERed.inv : ha => //= _. - * move => a1 a2 b0 ha ? [*]. subst. - have [a2 [h0 h1]] : exists a2, ERed.R a2 a /\ a1 = ren_Tm shift a2 by admit. subst. - eexists. split; cycle 1. - apply : relations.rtc_r; cycle 1. - apply ERed.AppEta. - apply rtc_refl. - eauto using relations.rtc_once. - * hauto q:on ctrs:rtc, ERed.R inv:ERed.R. - - move => a c ha. - elim /ERed.inv : ha => //= _. - + hauto l:on. - + move => a0 a1 b0 ha ? [*]. subst. - elim /ERed.inv : ha => //= _. - move => p a1 a2 ha ? [*]. subst. - exists a1. split. by apply relations.rtc_once. - apply : rtc_l. apply ERed.PairEta. - apply : rtc_l. apply ERed.PairCong1. eauto using ERed.ProjCong. - apply rtc_refl. - + move => a0 b0 b1 ha ? [*]. subst. - elim /ERed.inv : ha => //= _ p a0 a1 h ? [*]. subst. - exists a0. split; first by apply relations.rtc_once. - apply : rtc_l; first by apply ERed.PairEta. - apply relations.rtc_once. - hauto lq:on ctrs:ERed.R. - - move => a0 a1 ha iha c. - elim /ERed.inv => //= _. - + move => a2 ? [*]. subst. - elim /ERed.inv : ha => //=_. - * move => a1 a2 b0 ha ? [*] {iha}. subst. - have [a0 [h0 h1]] : exists a0, ERed.R a0 c /\ a1 = ren_Tm shift a0 by admit. subst. - exists a0. split; last by apply relations.rtc_once. - apply relations.rtc_once. apply ERed.AppEta. - * hauto q:on inv:ERed.R. - + hauto l:on use:EReds.AbsCong. - - move => a0 a1 b ha iha c. - elim /ERed.inv => //= _. - + hauto lq:on ctrs:rtc use:EReds.AppCong. - + hauto lq:on use:@relations.rtc_once ctrs:ERed.R. - - move => a b0 b1 hb ihb c. - elim /ERed.inv => //=_. - + move => a0 a1 a2 ha ? [*]. subst. - move {ihb}. exists (App a0 b0). - hauto lq:on use:@relations.rtc_once ctrs:ERed.R. - + hauto lq:on ctrs:rtc use:EReds.AppCong. - - move => a0 a1 b ha iha c. - elim /ERed.inv => //= _. - + move => ? ?[*]. subst. - elim /ERed.inv : ha => //= _ p a1 a2 h ? [*]. subst. - exists a1. split; last by apply relations.rtc_once. - apply : rtc_l. apply ERed.PairEta. - apply relations.rtc_once. hauto lq:on ctrs:ERed.R. - + hauto lq:on ctrs:rtc use:EReds.PairCong. - + hauto lq:on ctrs:ERed.R use:@relations.rtc_once. - - move => a b0 b1 hb hc c. elim /ERed.inv => //= _. - + move => ? ? [*]. subst. - elim /ERed.inv : hb => //= _ p a0 a1 ha ? [*]. subst. - move {hc}. - exists a0. split; last by apply relations.rtc_once. - apply : rtc_l; first by apply ERed.PairEta. - hauto lq:on ctrs:ERed.R use:@relations.rtc_once. - + hauto lq:on ctrs:ERed.R use:@relations.rtc_once. - + hauto lq:on ctrs:rtc use:EReds.PairCong. - - qauto l:on inv:ERed.R use:EReds.ProjCong. - - move => p A0 A1 B hA ihA. - move => c. elim/ERed.inv => //=. - + hauto lq:on ctrs:rtc use:EReds.BindCong. - + hauto lq:on ctrs:ERed.R use:@relations.rtc_once. - - move => p A B0 B1 hB ihB c. - elim /ERed.inv => //=. - + hauto lq:on ctrs:ERed.R use:@relations.rtc_once. - + hauto lq:on ctrs:rtc use:EReds.BindCong. -Admitted. diff --git a/theories/typing.v b/theories/typing.v deleted file mode 100644 index d41facd..0000000 --- a/theories/typing.v +++ /dev/null @@ -1,251 +0,0 @@ -Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax. - -Reserved Notation "Γ ⊢ a ∈ A" (at level 70). -Reserved Notation "Γ ⊢ a ≡ b ∈ A" (at level 70). -Reserved Notation "Γ ⊢ A ≲ B" (at level 70). -Reserved Notation "⊢ Γ" (at level 70). -Inductive Wt : list PTm -> PTm -> PTm -> Prop := -| T_Var i Γ A : - ⊢ Γ -> - lookup i Γ A -> - Γ ⊢ VarPTm i ∈ A - -| T_Bind Γ i p (A : PTm) (B : PTm) : - Γ ⊢ A ∈ PUniv i -> - cons A Γ ⊢ B ∈ PUniv i -> - Γ ⊢ PBind p A B ∈ PUniv i - -| T_Abs Γ (a : PTm) A B i : - Γ ⊢ PBind PPi A B ∈ (PUniv i) -> - (cons A Γ) ⊢ a ∈ B -> - Γ ⊢ PAbs a ∈ PBind PPi A B - -| T_App Γ (b a : PTm) A B : - Γ ⊢ b ∈ PBind PPi A B -> - Γ ⊢ a ∈ A -> - Γ ⊢ PApp b a ∈ subst_PTm (scons a VarPTm) B - -| T_Pair Γ (a b : PTm) A B i : - Γ ⊢ PBind PSig A B ∈ (PUniv i) -> - Γ ⊢ a ∈ A -> - Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B -> - Γ ⊢ PPair a b ∈ PBind PSig A B - -| T_Proj1 Γ (a : PTm) A B : - Γ ⊢ a ∈ PBind PSig A B -> - Γ ⊢ PProj PL a ∈ A - -| T_Proj2 Γ (a : PTm) A B : - Γ ⊢ a ∈ PBind PSig A B -> - Γ ⊢ PProj PR a ∈ subst_PTm (scons (PProj PL a) VarPTm) B - -| T_Univ Γ i : - ⊢ Γ -> - Γ ⊢ PUniv i ∈ PUniv (S i) - -| T_Nat Γ i : - ⊢ Γ -> - Γ ⊢ PNat ∈ PUniv i - -| T_Zero Γ : - ⊢ Γ -> - Γ ⊢ PZero ∈ PNat - -| T_Suc Γ (a : PTm) : - Γ ⊢ a ∈ PNat -> - Γ ⊢ PSuc a ∈ PNat - -| T_Ind Γ P (a : PTm) b c i : - cons PNat Γ ⊢ P ∈ PUniv i -> - Γ ⊢ a ∈ PNat -> - Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P -> - (cons P (cons PNat Γ)) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P) -> - Γ ⊢ PInd P a b c ∈ subst_PTm (scons a VarPTm) P - -| T_Conv Γ (a : PTm) A B : - Γ ⊢ a ∈ A -> - Γ ⊢ A ≲ B -> - Γ ⊢ a ∈ B - -with Eq : list PTm -> PTm -> PTm -> PTm -> Prop := -(* Structural *) -| E_Refl Γ (a : PTm ) A : - Γ ⊢ a ∈ A -> - Γ ⊢ a ≡ a ∈ A - -| E_Symmetric Γ (a b : PTm) A : - Γ ⊢ a ≡ b ∈ A -> - Γ ⊢ b ≡ a ∈ A - -| E_Transitive Γ (a b c : PTm) A : - Γ ⊢ a ≡ b ∈ A -> - Γ ⊢ b ≡ c ∈ A -> - Γ ⊢ a ≡ c ∈ A - -(* Congruence *) -| E_Bind Γ i p (A0 A1 : PTm) B0 B1 : - Γ ⊢ A0 ∈ PUniv i -> - Γ ⊢ A0 ≡ A1 ∈ PUniv i -> - (cons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i -> - Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i - -| E_Abs Γ (a b : PTm) A B i : - Γ ⊢ PBind PPi A B ∈ (PUniv i) -> - (cons A Γ) ⊢ a ≡ b ∈ B -> - Γ ⊢ PAbs a ≡ PAbs b ∈ PBind PPi A B - -| E_App Γ i (b0 b1 a0 a1 : PTm) A B : - Γ ⊢ PBind PPi A B ∈ (PUniv i) -> - Γ ⊢ b0 ≡ b1 ∈ PBind PPi A B -> - Γ ⊢ a0 ≡ a1 ∈ A -> - Γ ⊢ PApp b0 a0 ≡ PApp b1 a1 ∈ subst_PTm (scons a0 VarPTm) B - -| E_Pair Γ (a0 a1 b0 b1 : PTm) A B i : - Γ ⊢ PBind PSig A B ∈ (PUniv i) -> - Γ ⊢ a0 ≡ a1 ∈ A -> - Γ ⊢ b0 ≡ b1 ∈ subst_PTm (scons a0 VarPTm) B -> - Γ ⊢ PPair a0 b0 ≡ PPair a1 b1 ∈ PBind PSig A B - -| E_Proj1 Γ (a b : PTm) A B : - Γ ⊢ a ≡ b ∈ PBind PSig A B -> - Γ ⊢ PProj PL a ≡ PProj PL b ∈ A - -| E_Proj2 Γ i (a b : PTm) A B : - Γ ⊢ PBind PSig A B ∈ (PUniv i) -> - Γ ⊢ a ≡ b ∈ PBind PSig A B -> - Γ ⊢ PProj PR a ≡ PProj PR b ∈ subst_PTm (scons (PProj PL a) VarPTm) B - -| E_IndCong Γ P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 i : - (cons PNat Γ) ⊢ P0 ∈ PUniv i -> - (cons PNat Γ) ⊢ P0 ≡ P1 ∈ PUniv i -> - Γ ⊢ a0 ≡ a1 ∈ PNat -> - Γ ⊢ b0 ≡ b1 ∈ subst_PTm (scons PZero VarPTm) P0 -> - (cons P0 ((cons PNat Γ))) ⊢ c0 ≡ c1 ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P0) -> - Γ ⊢ PInd P0 a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P0 - -| E_SucCong Γ (a b : PTm) : - Γ ⊢ a ≡ b ∈ PNat -> - Γ ⊢ PSuc a ≡ PSuc b ∈ PNat - -| E_Conv Γ (a b : PTm) A B : - Γ ⊢ a ≡ b ∈ A -> - Γ ⊢ A ≲ B -> - Γ ⊢ a ≡ b ∈ B - -(* Beta *) -| E_AppAbs Γ (a : PTm) b A B i: - Γ ⊢ PBind PPi A B ∈ PUniv i -> - Γ ⊢ b ∈ A -> - (cons A Γ) ⊢ a ∈ B -> - Γ ⊢ PApp (PAbs a) b ≡ subst_PTm (scons b VarPTm) a ∈ subst_PTm (scons b VarPTm ) B - -| E_ProjPair1 Γ (a b : PTm) A B i : - Γ ⊢ PBind PSig A B ∈ (PUniv i) -> - Γ ⊢ a ∈ A -> - Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B -> - Γ ⊢ PProj PL (PPair a b) ≡ a ∈ A - -| E_ProjPair2 Γ (a b : PTm) A B i : - Γ ⊢ PBind PSig A B ∈ (PUniv i) -> - Γ ⊢ a ∈ A -> - Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B -> - Γ ⊢ PProj PR (PPair a b) ≡ b ∈ subst_PTm (scons a VarPTm) B - -| E_IndZero Γ P i (b : PTm) c : - (cons PNat Γ) ⊢ P ∈ PUniv i -> - Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P -> - (cons P (cons PNat Γ)) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P) -> - Γ ⊢ PInd P PZero b c ≡ b ∈ subst_PTm (scons PZero VarPTm) P - -| E_IndSuc Γ P (a : PTm) b c i : - (cons PNat Γ) ⊢ P ∈ PUniv i -> - Γ ⊢ a ∈ PNat -> - Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P -> - (cons P (cons PNat Γ)) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P) -> - Γ ⊢ PInd P (PSuc a) b c ≡ (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c) ∈ subst_PTm (scons (PSuc a) VarPTm) P - -(* Eta *) -| E_AppEta Γ (b : PTm) A B i : - Γ ⊢ PBind PPi A B ∈ (PUniv i) -> - Γ ⊢ b ∈ PBind PPi A B -> - Γ ⊢ PAbs (PApp (ren_PTm shift b) (VarPTm var_zero)) ≡ b ∈ PBind PPi A B - -| E_PairEta Γ (a : PTm ) A B i : - Γ ⊢ PBind PSig A B ∈ (PUniv i) -> - Γ ⊢ a ∈ PBind PSig A B -> - Γ ⊢ a ≡ PPair (PProj PL a) (PProj PR a) ∈ PBind PSig A B - -with LEq : list PTm -> PTm -> PTm -> Prop := -(* Structural *) -| Su_Transitive Γ (A B C : PTm) : - Γ ⊢ A ≲ B -> - Γ ⊢ B ≲ C -> - Γ ⊢ A ≲ C - -(* Congruence *) -| Su_Univ Γ i j : - ⊢ Γ -> - i <= j -> - Γ ⊢ PUniv i ≲ PUniv j - -| Su_Pi Γ (A0 A1 : PTm) B0 B1 i : - Γ ⊢ A0 ∈ PUniv i -> - Γ ⊢ A1 ≲ A0 -> - (cons A0 Γ) ⊢ B0 ≲ B1 -> - Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 - -| Su_Sig Γ (A0 A1 : PTm) B0 B1 i : - Γ ⊢ A1 ∈ PUniv i -> - Γ ⊢ A0 ≲ A1 -> - (cons A1 Γ) ⊢ B0 ≲ B1 -> - Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 - -(* Injecting from equalities *) -| Su_Eq Γ (A : PTm) B i : - Γ ⊢ A ≡ B ∈ PUniv i -> - Γ ⊢ A ≲ B - -(* Projection axioms *) -| Su_Pi_Proj1 Γ (A0 A1 : PTm) B0 B1 : - Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 -> - Γ ⊢ A1 ≲ A0 - -| Su_Sig_Proj1 Γ (A0 A1 : PTm) B0 B1 : - Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 -> - Γ ⊢ A0 ≲ A1 - -| Su_Pi_Proj2 Γ (a0 a1 A0 A1 : PTm ) B0 B1 : - Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 -> - Γ ⊢ a0 ≡ a1 ∈ A1 -> - Γ ⊢ subst_PTm (scons a0 VarPTm) B0 ≲ subst_PTm (scons a1 VarPTm) B1 - -| Su_Sig_Proj2 Γ (a0 a1 A0 A1 : PTm) B0 B1 : - Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 -> - Γ ⊢ a0 ≡ a1 ∈ A0 -> - Γ ⊢ subst_PTm (scons a0 VarPTm) B0 ≲ subst_PTm (scons a1 VarPTm) B1 - -with Wff : list PTm -> Prop := -| Wff_Nil : - ⊢ nil -| Wff_Cons Γ (A : PTm) i : - ⊢ Γ -> - Γ ⊢ A ∈ PUniv i -> - (* -------------------------------- *) - ⊢ (cons A Γ) - -where -"Γ ⊢ a ∈ A" := (Wt Γ a A) and "⊢ Γ" := (Wff Γ) and "Γ ⊢ a ≡ b ∈ A" := (Eq Γ a b A) and "Γ ⊢ A ≲ B" := (LEq Γ A B). - -Scheme wf_ind := Induction for Wff Sort Prop - with wt_ind := Induction for Wt Sort Prop - with eq_ind := Induction for Eq Sort Prop - with le_ind := Induction for LEq Sort Prop. - -Combined Scheme wt_mutual from wf_ind, wt_ind, eq_ind, le_ind. - -(* Lemma lem : *) -(* (forall n (Γ : fin n -> PTm n), ⊢ Γ -> ...) /\ *) -(* (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> ...) /\ *) -(* (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> ...) /\ *) -(* (forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> ...). *) -(* Proof. apply wt_mutual. ... *)