From: Ferruccio Guidi Date: Sun, 7 Aug 2011 20:44:42 +0000 (+0000) Subject: - cpr is now defined and the cpr_flat propery is proved! (it did not X-Git-Tag: make_still_working~2342 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=baccd5a2f3b79c295b1f9444575bfb351577634e;p=helm.git - cpr is now defined and the cpr_flat propery is proved! (it did not hold in the first version of the calculus because cpr (aka pr2) was badly designed). - relocation properties of tpr closed! - fixed bug in drop (base case was too restrictive) - some refactoring --- diff --git a/matita/matita/lib/lambda-delta/notation.ma b/matita/matita/lib/lambda-delta/notation.ma index 86e31894c..375d06abc 100644 --- a/matita/matita/lib/lambda-delta/notation.ma +++ b/matita/matita/lib/lambda-delta/notation.ma @@ -11,7 +11,7 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -(* *****************************************************************) +(* Syntax *******************************************************************) notation "hvbox( ⋆ )" non associative with precedence 90 @@ -36,7 +36,11 @@ notation "hvbox( 𝕗 { I } break (term 90 T1) . break (term 90 T) )" notation "hvbox( T . break 𝕓 { I } break (term 90 T1) )" non associative with precedence 89 for @{ 'DBind $T $I $T1 }. - +(* +notation "hvbox( | L | )" + non associative with precedence 70 + for @{ 'Length $L }. +*) notation "hvbox( # term 90 x )" non associative with precedence 90 for @{ 'Weight $x }. @@ -45,7 +49,7 @@ notation "hvbox( # [ x , break y ] )" non associative with precedence 90 for @{ 'Weight $x $y }. -(* substitution *************************************************************) +(* Substitution *************************************************************) notation "hvbox( T1 break [ d , break e ] ≈ break T2 )" non associative with precedence 45 @@ -63,7 +67,7 @@ notation "hvbox( L ⊢ break term 90 T1 break [ d , break e ] ≫ break T2 )" non associative with precedence 45 for @{ 'PSubst $L $T1 $d $e $T2 }. -(* reduction ****************************************************************) +(* Reduction ****************************************************************) notation "hvbox( T1 ⇒ break T2 )" non associative with precedence 45 diff --git a/matita/matita/lib/lambda-delta/reduction/cpr.ma b/matita/matita/lib/lambda-delta/reduction/cpr.ma new file mode 100644 index 000000000..80ccfd947 --- /dev/null +++ b/matita/matita/lib/lambda-delta/reduction/cpr.ma @@ -0,0 +1,53 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "lambda-delta/syntax/length.ma". +include "lambda-delta/reduction/tpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +definition cpr: lenv → term → term → Prop ≝ + λL,T1,T2. ∃∃T. T1 ⇒ T & L ⊢ T [0, |L|] ≫ T2. + +interpretation + "context-sensitive parallel reduction (term)" + 'PRed L T1 T2 = (cpr L T1 T2). + +(* Basic properties *********************************************************) + +lemma cpr_pr: ∀T1,T2. T1 ⇒ T2 → ∀L. L ⊢ T1 ⇒ T2. +/2/ qed. + +axiom cpr_pts: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → L ⊢ T1 ⇒ T2. + +lemma cpr_refl: ∀L,T. L ⊢ T ⇒ T. +/2/ qed. + +(* NOTE: new property *) +lemma cpr_flat: ∀I,L,V1,V2,T1,T2. + L ⊢ V1 ⇒ V2 → L ⊢ T1 ⇒ T2 → L ⊢ 𝕗{I} V1. T1 ⇒ 𝕗{I} V2. T2. +#I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/ +qed. + +axiom cpr_delta: ∀L,K,V0,V,i. + ↓[0, i] L ≡ K. 𝕓{Abbr} V0 → ↑[0, i + 1] V0 ≡ V → L ⊢ #i ⇒ V. +(* +#L #K #V0 #V #i #HLK #HV0 +@ex2_1_intro [2: // |1: skip ] +@pts_subst [4,6,7,8: // |1,2,3: skip |5: +*) + +lemma cpr_cast: ∀L,V,T1,T2. + L ⊢ T1 ⇒ T2 → L ⊢ 𝕗{Cast} V. T1 ⇒ T2. +#L #V #T1 #T2 * /3/ +qed. + +(* Basic inversion lemmas ***************************************************) diff --git a/matita/matita/lib/lambda-delta/reduction/lpr.ma b/matita/matita/lib/lambda-delta/reduction/lpr.ma new file mode 100644 index 000000000..92f77215c --- /dev/null +++ b/matita/matita/lib/lambda-delta/reduction/lpr.ma @@ -0,0 +1,41 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "lambda-delta/reduction/tpr.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) + +inductive lpr: lenv → lenv → Prop ≝ +| lpr_sort: lpr (⋆) (⋆) +| lpr_item: ∀K1,K2,I,V1,V2. + lpr K1 K2 → V1 ⇒ V2 → lpr (K1. 𝕓{I} V1) (K2. 𝕓{I} V2) (*𝕓*) +. + +interpretation + "context-free parallel reduction (environment)" + 'PRed L1 L2 = (lpr L1 L2). + +(* Basic inversion lemmas ***************************************************) + +lemma lpr_inv_item1_aux: ∀L1,L2. L1 ⇒ L2 → ∀K1,I,V1. L1 = K1. 𝕓{I} V1 → + ∃∃K2,V2. K1 ⇒ K2 & V1 ⇒ V2 & L2 = K2. 𝕓{I} V2. +#L1 #L2 * -L1 L2 +[ #K1 #I #V1 #H destruct +| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #L #J #W #H destruct - K1 I V1 /2 width=5/ +] +qed. + +lemma lpr_inv_item1: ∀K1,I,V1,L2. K1. 𝕓{I} V1 ⇒ L2 → + ∃∃K2,V2. K1 ⇒ K2 & V1 ⇒ V2 & L2 = K2. 𝕓{I} V2. +/2/ qed. diff --git a/matita/matita/lib/lambda-delta/reduction/lpr_defs.ma b/matita/matita/lib/lambda-delta/reduction/lpr_defs.ma deleted file mode 100644 index 8adec5ca5..000000000 --- a/matita/matita/lib/lambda-delta/reduction/lpr_defs.ma +++ /dev/null @@ -1,41 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "lambda-delta/reduction/tpr_defs.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -inductive lpr: lenv → lenv → Prop ≝ -| lpr_sort: lpr (⋆) (⋆) -| lpr_item: ∀K1,K2,I,V1,V2. - lpr K1 K2 → V1 ⇒ V2 → lpr (K1. 𝕓{I} V1) (K2. 𝕓{I} V2) (*𝕓*) -. - -interpretation - "context-free parallel reduction (environment)" - 'PRed L1 L2 = (lpr L1 L2). - -(* Basic inversion lemmas ***************************************************) - -lemma lpr_inv_item1_aux: ∀L1,L2. L1 ⇒ L2 → ∀K1,I,V1. L1 = K1. 𝕓{I} V1 → - ∃∃K2,V2. K1 ⇒ K2 & V1 ⇒ V2 & L2 = K2. 𝕓{I} V2. -#L1 #L2 * -L1 L2 -[ #K1 #I #V1 #H destruct -| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #L #J #W #H destruct - K1 I V1 /2 width=5/ -] -qed. - -lemma lpr_inv_item1: ∀K1,I,V1,L2. K1. 𝕓{I} V1 ⇒ L2 → - ∃∃K2,V2. K1 ⇒ K2 & V1 ⇒ V2 & L2 = K2. 𝕓{I} V2. -/2/ qed. diff --git a/matita/matita/lib/lambda-delta/reduction/tpr.ma b/matita/matita/lib/lambda-delta/reduction/tpr.ma new file mode 100644 index 000000000..3145ccaa5 --- /dev/null +++ b/matita/matita/lib/lambda-delta/reduction/tpr.ma @@ -0,0 +1,240 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "lambda-delta/substitution/pts.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) + +inductive tpr: term → term → Prop ≝ +| tpr_sort : ∀k. tpr (⋆k) (⋆k) +| tpr_lref : ∀i. tpr (#i) (#i) +| tpr_bind : ∀I,V1,V2,T1,T2. tpr V1 V2 → tpr T1 T2 → + tpr (𝕓{I} V1. T1) (𝕓{I} V2. T2) +| tpr_flat : ∀I,V1,V2,T1,T2. tpr V1 V2 → tpr T1 T2 → + tpr (𝕗{I} V1. T1) (𝕗{I} V2. T2) +| tpr_beta : ∀V1,V2,W,T1,T2. + tpr V1 V2 → tpr T1 T2 → + tpr (𝕚{Appl} V1. 𝕚{Abst} W. T1) (𝕚{Abbr} V2. T2) +| tpr_delta: ∀V1,V2,T1,T2,T. + tpr V1 V2 → tpr T1 T2 → ⋆. 𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T → + tpr (𝕚{Abbr} V1. T1) (𝕚{Abbr} V2. T) +| tpr_theta: ∀V,V1,V2,W1,W2,T1,T2. + tpr V1 V2 → ↑[0,1] V2 ≡ V → tpr W1 W2 → tpr T1 T2 → + tpr (𝕚{Appl} V1. 𝕚{Abbr} W1. T1) (𝕚{Abbr} W2. 𝕚{Appl} V. T2) +| tpr_zeta : ∀V,T,T1,T2. ↑[0,1] T1 ≡ T → tpr T1 T2 → + tpr (𝕚{Abbr} V. T) T2 +| tpr_tau : ∀V,T1,T2. tpr T1 T2 → tpr (𝕚{Cast} V. T1) T2 +. + +interpretation + "context-free parallel reduction (term)" + 'PRed T1 T2 = (tpr T1 T2). + +(* Basic properties *********************************************************) + +lemma tpr_refl: ∀T. T ⇒ T. +#T elim T -T // +#I elim I -I /2/ +qed. + +(* Basic inversion lemmas ***************************************************) + +lemma tpr_inv_sort1_aux: ∀U1,U2. U1 ⇒ U2 → ∀k. U1 = ⋆k → U2 = ⋆k. +#U1 #U2 * -U1 U2 +[ #k0 #k #H destruct -k0 // +| #i #k #H destruct +| #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct +| #V1 #V2 #W #T1 #T2 #_ #_ #k #H destruct +| #V1 #V2 #T1 #T2 #T #_ #_ #_ #k #H destruct +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #k #H destruct +| #V #T #T1 #T2 #_ #_ #k #H destruct +| #V #T1 #T2 #_ #k #H destruct +] +qed. + +lemma tpr_inv_sort1: ∀k,U2. ⋆k ⇒ U2 → U2 = ⋆k. +/2/ qed. + +lemma tpr_inv_lref1_aux: ∀U1,U2. U1 ⇒ U2 → ∀i. U1 = #i → U2 = #i. +#U1 #U2 * -U1 U2 +[ #k #i #H destruct +| #j #i #H destruct -j // +| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct +| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct +| #V1 #V2 #W #T1 #T2 #_ #_ #i #H destruct +| #V1 #V2 #T1 #T2 #T #_ #_ #_ #i #H destruct +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #i #H destruct +| #V #T #T1 #T2 #_ #_ #i #H destruct +| #V #T1 #T2 #_ #i #H destruct +] +qed. + +lemma tpr_inv_lref1: ∀i,U2. #i ⇒ U2 → U2 = #i. +/2/ qed. + +lemma tpr_inv_abbr1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,T1. U1 = 𝕚{Abbr} V1. T1 → + ∨∨ ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abbr} V2. T2 + | ∃∃V2,T2,T. V1 ⇒ V2 & T1 ⇒ T2 & + ⋆. 𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T & + U2 = 𝕚{Abbr} V2. T + | ∃∃T. ↑[0,1] T ≡ T1 & T ⇒ U2. +#U1 #U2 * -U1 U2 +[ #k #V #T #H destruct +| #i #V #T #H destruct +| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /3 width=5/ +| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct +| #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct +| #V1 #V2 #T1 #T2 #T #HV12 #HT12 #HT2 #V0 #T0 #H destruct -V1 T1 /3 width=7/ +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct +| #V #T #T1 #T2 #HT1 #HT12 #V0 #T0 #H destruct -V T /3/ +| #V #T1 #T2 #_ #V0 #T0 #H destruct +] +qed. + +lemma tpr_inv_abbr1: ∀V1,T1,U2. 𝕚{Abbr} V1. T1 ⇒ U2 → + ∨∨ ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abbr} V2. T2 + | ∃∃V2,T2,T. V1 ⇒ V2 & T1 ⇒ T2 & + ⋆. 𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T & + U2 = 𝕚{Abbr} V2. T + | ∃∃T. ↑[0,1] T ≡ T1 & tpr T U2. +/2/ qed. + +lemma tpr_inv_abst1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,T1. U1 = 𝕚{Abst} V1. T1 → + ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abst} V2. T2. +#U1 #U2 * -U1 U2 +[ #k #V #T #H destruct +| #i #V #T #H destruct +| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /2 width=5/ +| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct +| #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct +| #V1 #V2 #T1 #T2 #T #_ #_ #_ #V0 #T0 #H destruct +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct +| #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct +| #V #T1 #T2 #_ #V0 #T0 #H destruct +] +qed. + +lemma tpr_inv_abst1: ∀V1,T1,U2. 𝕚{Abst} V1. T1 ⇒ U2 → + ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abst} V2. T2. +/2/ qed. + +lemma tpr_inv_bind1: ∀V1,T1,U2,I. 𝕓{I} V1. T1 ⇒ U2 → + ∨∨ ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕓{I} V2. T2 + | ∃∃V2,T2,T. V1 ⇒ V2 & T1 ⇒ T2 & + ⋆. 𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T & + U2 = 𝕚{Abbr} V2. T & I = Abbr + | ∃∃T. ↑[0,1] T ≡ T1 & tpr T U2 & I = Abbr. +#V1 #T1 #U2 * #H +[ elim (tpr_inv_abbr1 … H) -H * /3 width=7/ +| /3/ +] +qed. + +lemma tpr_inv_appl1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,U0. U1 = 𝕚{Appl} V1. U0 → + ∨∨ ∃∃V2,T2. V1 ⇒ V2 & U0 ⇒ T2 & + U2 = 𝕚{Appl} V2. T2 + | ∃∃V2,W,T1,T2. V1 ⇒ V2 & T1 ⇒ T2 & + U0 = 𝕚{Abst} W. T1 & + U2 = 𝕓{Abbr} V2. T2 + | ∃∃V2,V,W1,W2,T1,T2. V1 ⇒ V2 & W1 ⇒ W2 & T1 ⇒ T2 & + ↑[0,1] V2 ≡ V & + U0 = 𝕚{Abbr} W1. T1 & + U2 = 𝕚{Abbr} W2. 𝕚{Appl} V. T2. +#U1 #U2 * -U1 U2 +[ #k #V #T #H destruct +| #i #V #T #H destruct +| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct +| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /3 width=5/ +| #V1 #V2 #W #T1 #T2 #HV12 #HT12 #V #T #H destruct -V1 T /3 width=8/ +| #V1 #V2 #T1 #T2 #T #_ #_ #_ #V0 #T0 #H destruct +| #V #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HV2 #HW12 #HT12 #V0 #T0 #H + destruct -V1 T0 /3 width=12/ +| #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct +| #V #T1 #T2 #_ #V0 #T0 #H destruct +] +qed. + +lemma tpr_inv_appl1: ∀V1,U0,U2. 𝕚{Appl} V1. U0 ⇒ U2 → + ∨∨ ∃∃V2,T2. V1 ⇒ V2 & U0 ⇒ T2 & + U2 = 𝕚{Appl} V2. T2 + | ∃∃V2,W,T1,T2. V1 ⇒ V2 & T1 ⇒ T2 & + U0 = 𝕚{Abst} W. T1 & + U2 = 𝕓{Abbr} V2. T2 + | ∃∃V2,V,W1,W2,T1,T2. V1 ⇒ V2 & W1 ⇒ W2 & T1 ⇒ T2 & + ↑[0,1] V2 ≡ V & + U0 = 𝕚{Abbr} W1. T1 & + U2 = 𝕚{Abbr} W2. 𝕚{Appl} V. T2. +/2/ qed. + +lemma tpr_inv_cast1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,T1. U1 = 𝕚{Cast} V1. T1 → + (∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Cast} V2. T2) + ∨ T1 ⇒ U2. +#U1 #U2 * -U1 U2 +[ #k #V #T #H destruct +| #i #V #T #H destruct +| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct +| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /3 width=5/ +| #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct +| #V1 #V2 #T1 #T2 #T #_ #_ #_ #V0 #T0 #H destruct +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct +| #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct +| #V #T1 #T2 #HT12 #V0 #T0 #H destruct -V T1 /2/ +] +qed. + +lemma tpr_inv_cast1: ∀V1,T1,U2. 𝕚{Cast} V1. T1 ⇒ U2 → + (∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Cast} V2. T2) + ∨ T1 ⇒ U2. +/2/ qed. + +lemma tpr_inv_flat1: ∀V1,U0,U2,I. 𝕗{I} V1. U0 ⇒ U2 → + ∨∨ ∃∃V2,T2. V1 ⇒ V2 & U0 ⇒ T2 & + U2 = 𝕗{I} V2. T2 + | ∃∃V2,W,T1,T2. V1 ⇒ V2 & T1 ⇒ T2 & + U0 = 𝕚{Abst} W. T1 & + U2 = 𝕓{Abbr} V2. T2 & I = Appl + | ∃∃V2,V,W1,W2,T1,T2. V1 ⇒ V2 & W1 ⇒ W2 & T1 ⇒ T2 & + ↑[0,1] V2 ≡ V & + U0 = 𝕚{Abbr} W1. T1 & + U2 = 𝕚{Abbr} W2. 𝕚{Appl} V. T2 & + I = Appl + | (U0 ⇒ U2 ∧ I = Cast). +#V1 #U0 #U2 * #H +[ elim (tpr_inv_appl1 … H) -H * /3 width=12/ +| elim (tpr_inv_cast1 … H) -H [1: *] /3 width=5/ +] +qed. + +lemma tpr_inv_lref2_aux: ∀T1,T2. T1 ⇒ T2 → ∀i. T2 = #i → + ∨∨ T1 = #i + | ∃∃V,T,T0. ↑[O,1] T0 ≡ T & T0 ⇒ #i & + T1 = 𝕚{Abbr} V. T + | ∃∃V,T. T ⇒ #i & T1 = 𝕚{Cast} V. T. +#T1 #T2 * -T1 T2 +[ #k #i #H destruct +| #j #i /2/ +| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct +| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct +| #V1 #V2 #W #T1 #T2 #_ #_ #i #H destruct +| #V1 #V2 #T1 #T2 #T #_ #_ #_ #i #H destruct +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #i #H destruct +| #V #T #T1 #T2 #HT1 #HT12 #i #H destruct /3 width=6/ +| #V #T1 #T2 #HT12 #i #H destruct /3/ +] +qed. + +lemma tpr_inv_lref2: ∀T1,i. T1 ⇒ #i → + ∨∨ T1 = #i + | ∃∃V,T,T0. ↑[O,1] T0 ≡ T & T0 ⇒ #i & + T1 = 𝕓{Abbr} V. T + | ∃∃V,T. T ⇒ #i & T1 = 𝕗{Cast} V. T. +/2/ qed. diff --git a/matita/matita/lib/lambda-delta/reduction/tpr_defs.ma b/matita/matita/lib/lambda-delta/reduction/tpr_defs.ma deleted file mode 100644 index cb9609951..000000000 --- a/matita/matita/lib/lambda-delta/reduction/tpr_defs.ma +++ /dev/null @@ -1,240 +0,0 @@ -(* - ||M|| This file is part of HELM, an Hypertextual, Electronic - ||A|| Library of Mathematics, developed at the Computer Science - ||T|| Department of the University of Bologna, Italy. - ||I|| - ||T|| - ||A|| This file is distributed under the terms of the - \ / GNU General Public License Version 2 - \ / - V_______________________________________________________________ *) - -include "lambda-delta/substitution/pts_defs.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) - -inductive tpr: term → term → Prop ≝ -| tpr_sort : ∀k. tpr (⋆k) (⋆k) -| tpr_lref : ∀i. tpr (#i) (#i) -| tpr_bind : ∀I,V1,V2,T1,T2. tpr V1 V2 → tpr T1 T2 → - tpr (𝕓{I} V1. T1) (𝕓{I} V2. T2) -| tpr_flat : ∀I,V1,V2,T1,T2. tpr V1 V2 → tpr T1 T2 → - tpr (𝕗{I} V1. T1) (𝕗{I} V2. T2) -| tpr_beta : ∀V1,V2,W,T1,T2. - tpr V1 V2 → tpr T1 T2 → - tpr (𝕚{Appl} V1. 𝕚{Abst} W. T1) (𝕚{Abbr} V2. T2) -| tpr_delta: ∀V1,V2,T1,T2,T. - tpr V1 V2 → tpr T1 T2 → ⋆. 𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T → - tpr (𝕚{Abbr} V1. T1) (𝕚{Abbr} V2. T) -| tpr_theta: ∀V,V1,V2,W1,W2,T1,T2. - tpr V1 V2 → ↑[0,1] V2 ≡ V → tpr W1 W2 → tpr T1 T2 → - tpr (𝕚{Appl} V1. 𝕚{Abbr} W1. T1) (𝕚{Abbr} W2. 𝕚{Appl} V. T2) -| tpr_zeta : ∀V,T,T1,T2. ↑[0,1] T1 ≡ T → tpr T1 T2 → - tpr (𝕚{Abbr} V. T) T2 -| tpr_tau : ∀V,T1,T2. tpr T1 T2 → tpr (𝕚{Cast} V. T1) T2 -. - -interpretation - "context-free parallel reduction (term)" - 'PRed T1 T2 = (tpr T1 T2). - -(* Basic properties *********************************************************) - -lemma tpr_refl: ∀T. T ⇒ T. -#T elim T -T // -#I elim I -I /2/ -qed. - -(* Basic inversion lemmas ***************************************************) - -lemma tpr_inv_sort1_aux: ∀U1,U2. U1 ⇒ U2 → ∀k. U1 = ⋆k → U2 = ⋆k. -#U1 #U2 * -U1 U2 -[ #k0 #k #H destruct -k0 // -| #i #k #H destruct -| #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct -| #V1 #V2 #W #T1 #T2 #_ #_ #k #H destruct -| #V1 #V2 #T1 #T2 #T #_ #_ #_ #k #H destruct -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #k #H destruct -| #V #T #T1 #T2 #_ #_ #k #H destruct -| #V #T1 #T2 #_ #k #H destruct -] -qed. - -lemma tpr_inv_sort1: ∀k,U2. ⋆k ⇒ U2 → U2 = ⋆k. -/2/ qed. - -lemma tpr_inv_lref1_aux: ∀U1,U2. U1 ⇒ U2 → ∀i. U1 = #i → U2 = #i. -#U1 #U2 * -U1 U2 -[ #k #i #H destruct -| #j #i #H destruct -j // -| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct -| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct -| #V1 #V2 #W #T1 #T2 #_ #_ #i #H destruct -| #V1 #V2 #T1 #T2 #T #_ #_ #_ #i #H destruct -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #i #H destruct -| #V #T #T1 #T2 #_ #_ #i #H destruct -| #V #T1 #T2 #_ #i #H destruct -] -qed. - -lemma tpr_inv_lref1: ∀i,U2. #i ⇒ U2 → U2 = #i. -/2/ qed. - -lemma tpr_inv_abbr1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,T1. U1 = 𝕚{Abbr} V1. T1 → - ∨∨ ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abbr} V2. T2 - | ∃∃V2,T2,T. V1 ⇒ V2 & T1 ⇒ T2 & - ⋆. 𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T & - U2 = 𝕚{Abbr} V2. T - | ∃∃T. ↑[0,1] T ≡ T1 & T ⇒ U2. -#U1 #U2 * -U1 U2 -[ #k #V #T #H destruct -| #i #V #T #H destruct -| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /3 width=5/ -| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct -| #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct -| #V1 #V2 #T1 #T2 #T #HV12 #HT12 #HT2 #V0 #T0 #H destruct -V1 T1 /3 width=7/ -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct -| #V #T #T1 #T2 #HT1 #HT12 #V0 #T0 #H destruct -V T /3/ -| #V #T1 #T2 #_ #V0 #T0 #H destruct -] -qed. - -lemma tpr_inv_abbr1: ∀V1,T1,U2. 𝕚{Abbr} V1. T1 ⇒ U2 → - ∨∨ ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abbr} V2. T2 - | ∃∃V2,T2,T. V1 ⇒ V2 & T1 ⇒ T2 & - ⋆. 𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T & - U2 = 𝕚{Abbr} V2. T - | ∃∃T. ↑[0,1] T ≡ T1 & tpr T U2. -/2/ qed. - -lemma tpr_inv_abst1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,T1. U1 = 𝕚{Abst} V1. T1 → - ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abst} V2. T2. -#U1 #U2 * -U1 U2 -[ #k #V #T #H destruct -| #i #V #T #H destruct -| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /2 width=5/ -| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct -| #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct -| #V1 #V2 #T1 #T2 #T #_ #_ #_ #V0 #T0 #H destruct -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct -| #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct -| #V #T1 #T2 #_ #V0 #T0 #H destruct -] -qed. - -lemma tpr_inv_abst1: ∀V1,T1,U2. 𝕚{Abst} V1. T1 ⇒ U2 → - ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abst} V2. T2. -/2/ qed. - -lemma tpr_inv_bind1: ∀V1,T1,U2,I. 𝕓{I} V1. T1 ⇒ U2 → - ∨∨ ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕓{I} V2. T2 - | ∃∃V2,T2,T. V1 ⇒ V2 & T1 ⇒ T2 & - ⋆. 𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T & - U2 = 𝕚{Abbr} V2. T & I = Abbr - | ∃∃T. ↑[0,1] T ≡ T1 & tpr T U2 & I = Abbr. -#V1 #T1 #U2 * #H -[ elim (tpr_inv_abbr1 … H) -H * /3 width=7/ -| /3/ -] -qed. - -lemma tpr_inv_appl1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,U0. U1 = 𝕚{Appl} V1. U0 → - ∨∨ ∃∃V2,T2. V1 ⇒ V2 & U0 ⇒ T2 & - U2 = 𝕚{Appl} V2. T2 - | ∃∃V2,W,T1,T2. V1 ⇒ V2 & T1 ⇒ T2 & - U0 = 𝕚{Abst} W. T1 & - U2 = 𝕓{Abbr} V2. T2 - | ∃∃V2,V,W1,W2,T1,T2. V1 ⇒ V2 & W1 ⇒ W2 & T1 ⇒ T2 & - ↑[0,1] V2 ≡ V & - U0 = 𝕚{Abbr} W1. T1 & - U2 = 𝕚{Abbr} W2. 𝕚{Appl} V. T2. -#U1 #U2 * -U1 U2 -[ #k #V #T #H destruct -| #i #V #T #H destruct -| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct -| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /3 width=5/ -| #V1 #V2 #W #T1 #T2 #HV12 #HT12 #V #T #H destruct -V1 T /3 width=8/ -| #V1 #V2 #T1 #T2 #T #_ #_ #_ #V0 #T0 #H destruct -| #V #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HV2 #HW12 #HT12 #V0 #T0 #H - destruct -V1 T0 /3 width=12/ -| #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct -| #V #T1 #T2 #_ #V0 #T0 #H destruct -] -qed. - -lemma tpr_inv_appl1: ∀V1,U0,U2. 𝕚{Appl} V1. U0 ⇒ U2 → - ∨∨ ∃∃V2,T2. V1 ⇒ V2 & U0 ⇒ T2 & - U2 = 𝕚{Appl} V2. T2 - | ∃∃V2,W,T1,T2. V1 ⇒ V2 & T1 ⇒ T2 & - U0 = 𝕚{Abst} W. T1 & - U2 = 𝕓{Abbr} V2. T2 - | ∃∃V2,V,W1,W2,T1,T2. V1 ⇒ V2 & W1 ⇒ W2 & T1 ⇒ T2 & - ↑[0,1] V2 ≡ V & - U0 = 𝕚{Abbr} W1. T1 & - U2 = 𝕚{Abbr} W2. 𝕚{Appl} V. T2. -/2/ qed. - -lemma tpr_inv_cast1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,T1. U1 = 𝕚{Cast} V1. T1 → - (∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Cast} V2. T2) - ∨ T1 ⇒ U2. -#U1 #U2 * -U1 U2 -[ #k #V #T #H destruct -| #i #V #T #H destruct -| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct -| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /3 width=5/ -| #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct -| #V1 #V2 #T1 #T2 #T #_ #_ #_ #V0 #T0 #H destruct -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct -| #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct -| #V #T1 #T2 #HT12 #V0 #T0 #H destruct -V T1 /2/ -] -qed. - -lemma tpr_inv_cast1: ∀V1,T1,U2. 𝕚{Cast} V1. T1 ⇒ U2 → - (∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Cast} V2. T2) - ∨ T1 ⇒ U2. -/2/ qed. - -lemma tpr_inv_flat1: ∀V1,U0,U2,I. 𝕗{I} V1. U0 ⇒ U2 → - ∨∨ ∃∃V2,T2. V1 ⇒ V2 & U0 ⇒ T2 & - U2 = 𝕗{I} V2. T2 - | ∃∃V2,W,T1,T2. V1 ⇒ V2 & T1 ⇒ T2 & - U0 = 𝕚{Abst} W. T1 & - U2 = 𝕓{Abbr} V2. T2 & I = Appl - | ∃∃V2,V,W1,W2,T1,T2. V1 ⇒ V2 & W1 ⇒ W2 & T1 ⇒ T2 & - ↑[0,1] V2 ≡ V & - U0 = 𝕚{Abbr} W1. T1 & - U2 = 𝕚{Abbr} W2. 𝕚{Appl} V. T2 & - I = Appl - | (U0 ⇒ U2 ∧ I = Cast). -#V1 #U0 #U2 * #H -[ elim (tpr_inv_appl1 … H) -H * /3 width=12/ -| elim (tpr_inv_cast1 … H) -H [1: *] /3 width=5/ -] -qed. - -lemma tpr_inv_lref2_aux: ∀T1,T2. T1 ⇒ T2 → ∀i. T2 = #i → - ∨∨ T1 = #i - | ∃∃V,T,T0. ↑[O,1] T0 ≡ T & T0 ⇒ #i & - T1 = 𝕚{Abbr} V. T - | ∃∃V,T. T ⇒ #i & T1 = 𝕚{Cast} V. T. -#T1 #T2 * -T1 T2 -[ #k #i #H destruct -| #j #i /2/ -| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct -| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct -| #V1 #V2 #W #T1 #T2 #_ #_ #i #H destruct -| #V1 #V2 #T1 #T2 #T #_ #_ #_ #i #H destruct -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #i #H destruct -| #V #T #T1 #T2 #HT1 #HT12 #i #H destruct /3 width=6/ -| #V #T1 #T2 #HT12 #i #H destruct /3/ -] -qed. - -lemma tpr_inv_lref2: ∀T1,i. T1 ⇒ #i → - ∨∨ T1 = #i - | ∃∃V,T,T0. ↑[O,1] T0 ≡ T & T0 ⇒ #i & - T1 = 𝕓{Abbr} V. T - | ∃∃V,T. T ⇒ #i & T1 = 𝕗{Cast} V. T. -/2/ qed. diff --git a/matita/matita/lib/lambda-delta/reduction/tpr_lift.ma b/matita/matita/lib/lambda-delta/reduction/tpr_lift.ma new file mode 100644 index 000000000..d639f7119 --- /dev/null +++ b/matita/matita/lib/lambda-delta/reduction/tpr_lift.ma @@ -0,0 +1,102 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "lambda-delta/substitution/pts_lift.ma". +include "lambda-delta/reduction/tpr.ma". + +(* Relocation properties ****************************************************) + +lemma tpr_lift: ∀T1,T2. T1 ⇒ T2 → + ∀d,e,U1. ↑[d, e] T1 ≡ U1 → ∀U2. ↑[d, e] T2 ≡ U2 → U1 ⇒ U2. +#T1 #T2 #H elim H -H T1 T2 +[ #k #d #e #U1 #HU1 #U2 #HU2 + lapply (lift_mono … HU1 … HU2) -HU1 #H destruct -U1; + lapply (lift_inv_sort1 … HU2) -HU2 #H destruct -U2 // +| #i #d #e #U1 #HU1 #U2 #HU2 + lapply (lift_mono … HU1 … HU2) -HU1 #H destruct -U1; + lapply (lift_inv_lref1 … HU2) * * #Hid #H destruct -U2 // +| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2 + elim (lift_inv_bind1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct -X1; + elim (lift_inv_bind1 … HX2) -HX2 #W2 #U2 #HVW2 #HTU2 #HX2 destruct -X2 /3/ +| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2 + elim (lift_inv_flat1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct -X1; + elim (lift_inv_flat1 … HX2) -HX2 #W2 #U2 #HVW2 #HTU2 #HX2 destruct -X2 /3/ +| #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2 + elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct -X1; + elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct -X; + elim (lift_inv_bind1 … HX2) -HX2 #V3 #T3 #HV23 #HT23 #HX2 destruct -X2 /3/ +| #V1 #V2 #T1 #T2 #T0 #HV12 #HT12 #HT2 #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2 + elim (lift_inv_bind1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct -X1; + elim (lift_inv_bind1 … HX2) -HX2 #W2 #U0 #HVW2 #HTU0 #HX2 destruct -X2; + elim (lift_total T2 (d + 1) e) #U2 #HTU2 + @tpr_delta + [4: @(pts_lift_le … HT2 … HTU2 HTU0 ?) /2/ |1: skip |2: /2/ |3: /2/ ] (**) (*/3. is too slow *) +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #X1 #HX1 #X2 #HX2 + elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct -X1; + elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct -X; + elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct -X2; + elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct -X; + elim (lift_trans_ge … HV2 … HV3 ?) -HV2 HV3 V // /3/ +| #V #T #T1 #T2 #HT1 #_ #IHT12 #d #e #X #HX #T0 #HT20 + elim (lift_inv_bind1 … HX) -HX #V3 #T3 #_ #HT3 #HX destruct -X; + elim (lift_trans_ge … HT1 … HT3 ?) -HT1 HT3 T // /3 width=6/ +| #V #T1 #T2 #_ #IHT12 #d #e #X #HX #T #HT2 + elim (lift_inv_flat1 … HX) -HX #V0 #T0 #_ #HT0 #HX destruct -X /3/ +] +qed. + +lemma tpr_inv_lift: ∀T1,T2. T1 ⇒ T2 → + ∀d,e,U1. ↑[d, e] U1 ≡ T1 → + ∃∃U2. ↑[d, e] U2 ≡ T2 & U1 ⇒ U2. +#T1 #T2 #H elim H -H T1 T2 +[ #k #d #e #U1 #HU1 + lapply (lift_inv_sort2 … HU1) -HU1 #H destruct -U1 /2/ +| #i #d #e #U1 #HU1 + lapply (lift_inv_lref2 … HU1) -HU1 * * #Hid #H destruct -U1 /3/ +| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X #HX + elim (lift_inv_bind2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct -X; + elim (IHV12 … HV01) -IHV12 HV01; + elim (IHT12 … HT01) -IHT12 HT01 /3 width=5/ +| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X #HX + elim (lift_inv_flat2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct -X; + elim (IHV12 … HV01) -IHV12 HV01; + elim (IHT12 … HT01) -IHT12 HT01 /3 width=5/ +| #V1 #V2 #W1 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X #HX + elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct -X; + elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct -Y; + elim (IHV12 … HV01) -IHV12 HV01; + elim (IHT12 … HT01) -IHT12 HT01 /3 width=5/ +| #V1 #V2 #T1 #T2 #T0 #_ #_ #HT20 #IHV12 #IHT12 #d #e #X #HX + elim (lift_inv_bind2 … HX) -HX #W1 #U1 #HWV1 #HUT1 #HX destruct -X; + elim (IHV12 … HWV1) -IHV12 HWV1 #W2 #HWV2 #HW12 + elim (IHT12 … HUT1) -IHT12 HUT1 #U2 #HUT2 #HU12 + elim (pts_inv_lift1_le … HT20 … HUT2 ?) -HT20 HUT2 // [3: /2 width=5/ |2: skip ] #U0 #HU20 #HUT0 + @ex2_1_intro [2: /2/ |1: skip |3: /2/ ] (**) (* /3 width=5/ is slow *) +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #X #HX + elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct -X; + elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct -Y; + elim (IHV12 … HV01) -IHV12 HV01 #V3 #HV32 #HV03 + elim (IHW12 … HW01) -IHW12 HW01 #W3 #HW32 #HW03 + elim (IHT12 … HT01) -IHT12 HT01 #T3 #HT32 #HT03 + elim (lift_trans_le … HV32 … HV2 ?) -HV32 HV2 V2 // #V2 #HV32 #HV2 + @ex2_1_intro [2: /3/ |1: skip |3: /2/ ] (**) (* /4 width=5/ is slow *) +| #V #T #T1 #T2 #HT1 #_ #IHT12 #d #e #X #HX + elim (lift_inv_bind2 … HX) -HX #V0 #T0 #_ #HT0 #H destruct -X; + elim (lift_div_le … HT1 … HT0 ?) -HT1 HT0 T // #T #HT0 #HT1 + elim (IHT12 … HT1) -IHT12 HT1 /3 width=5/ +| #V #T1 #T2 #_ #IHT12 #d #e #X #HX + elim (lift_inv_flat2 … HX) -HX #V0 #T0 #_ #HT01 #H destruct -X; + elim (IHT12 … HT01) -IHT12 HT01 /3/ +] +qed. diff --git a/matita/matita/lib/lambda-delta/reduction/tpr_main.ma b/matita/matita/lib/lambda-delta/reduction/tpr_main.ma deleted file mode 100644 index 1d5d7d713..000000000 --- a/matita/matita/lib/lambda-delta/reduction/tpr_main.ma +++ /dev/null @@ -1,130 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -(* -include "lambda-delta/substitution/lift_main.ma". -include "lambda-delta/substitution/drop_main.ma". -*) -include "lambda-delta/reduction/tpr_defs.ma". - -axiom tpr_lift: ∀T1,T2. T1 ⇒ T2 → - ∀d,e,U1. ↑[d, e] T1 ≡ U1 → ∀U2. ↑[d, e] T2 ≡ U2 → U1 ⇒ U2. -(* -#L #T1 #T2 #H elim H -H L T1 T2 -[ #L #k #d #e #K #_ #U1 #HU1 #U2 #HU2 - lapply (lift_mono … HU1 … HU2) -HU1 #H destruct -U1; - lapply (lift_inv_sort1 … HU2) -HU2 #H destruct -U2 // -| #L #i #d #e #K #_ #U1 #HU1 #U2 #HU2 - lapply (lift_mono … HU1 … HU2) -HU1 #H destruct -U1; - lapply (lift_inv_lref1 … HU2) * * #Hid #H destruct -U2 // -| #L #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #K #HKL #X1 #HX1 #X2 #HX2 - elim (lift_inv_bind1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct -X1; - elim (lift_inv_bind1 … HX2) -HX2 #W2 #U2 #HVW2 #HTU2 #HX2 destruct -X2 - /5 width=5/ -| #L #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #K #HKL #X1 #HX1 #X2 #HX2 - elim (lift_inv_flat1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct -X1; - elim (lift_inv_flat1 … HX2) -HX2 #W2 #U2 #HVW2 #HTU2 #HX2 destruct -X2 - /3 width=5/ -| #L #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #K #HKL #X1 #HX1 #X2 #HX2 - elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct -X1; - elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct -X; - elim (lift_inv_bind1 … HX2) -HX2 #V3 #T3 #HV23 #HT23 #HX2 destruct -X2 - /5 width=5/ -| #L #K0 #V1 #V2 #V0 #i #HLK0 #HV12 #HV20 #IHV12 #d #e #K #HKL #X #HX #V3 #HV03 - lapply (lift_inv_lref1 … HX) -HX * * #Hid #H destruct -X; - [ -HV12; - elim (lift_trans_ge … HV20 … HV03 ?) -HV20 HV03 V0 // #V0 #HV20 #HV03 - elim (drop_trans_le … HKL … HLK0 ?) -HKL HLK0 L /2/ #L #HKL #HLK0 - elim (drop_inv_skip2 … HLK0 ?) -HLK0 /2/ #K1 #V #HK10 #HV #H destruct -L; - @pr_delta [4,6: // |1,2,3: skip |5: /2 width=5/] (**) (* /2 width=5/ *) - | -IHV12; - lapply (lift_trans_be … HV20 … HV03 ? ?) -HV20 HV03 V0 /2/ #HV23 - lapply (drop_trans_ge … HKL … HLK0 ?) -HKL HLK0 L // -Hid #HKK0 - @pr_delta /width=6/ - ] -| #L #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #K #HKL #X1 #HX1 #X2 #HX2 - elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct -X1; - elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct -X; - elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct -X2; - elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct -X; - lapply (lift_trans_ge … HV2 … HV3 ?) -HV2 HV3 V // * #V #HV2 #HV3 - @pr_theta /3 width=5/ -| #L #V #T #T1 #T2 #HT1 #_ #IHT12 #d #e #K #HKL #X #HX #T0 #HT20 - elim (lift_inv_bind1 … HX) -HX #V3 #T3 #_ #HT3 #HX destruct -X; - lapply (lift_trans_ge … HT1 … HT3 ?) -HT1 HT3 T // * #T #HT1 #HT3 - @pr_zeta [2: // |1:skip | /2 width=5/] (**) (* /2 width=5/ *) -| #L #V #T1 #T2 #_ #IHT12 #d #e #K #HKL #X #HX #T #HT2 - elim (lift_inv_flat1 … HX) -HX #V0 #T0 #_ #HT0 #HX destruct -X; - @pr_tau /2 width=5/ -] -qed. -*) -axiom tpr_inv_lift: ∀T1,T2. T1 ⇒ T2 → - ∀d,e,U1. ↑[d, e] U1 ≡ T1 → - ∃∃U2. ↑[d, e] U2 ≡ T2 & U1 ⇒ U2. -(* -#L #T1 #T2 #H elim H -H L T1 T2 -[ #L #k #d #e #K #_ #U1 #HU1 - lapply (lift_inv_sort2 … HU1) -HU1 #H destruct -U1 /2/ -| #L #i #d #e #K #_ #U1 #HU1 - lapply (lift_inv_lref2 … HU1) -HU1 * * #Hid #H destruct -U1 /3/ -| #L #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_bind2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct -X; - elim (IHV12 … HLK … HV01) -IHV12 #V3 #HV32 #HV03 - elim (IHT12 … HT01) -IHT12 HT01 [2,3: -HV32 HV03 /3/] -HLK HV01 /3 width=5/ -| #L #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct -X; - elim (IHV12 … HLK … HV01) -IHV12 HV01 #V3 #HV32 #HV03 - elim (IHT12 … HLK … HT01) -IHT12 HT01 HLK /3 width=5/ -| #L #V1 #V2 #W1 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct -X; - elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct -Y; - elim (IHV12 … HLK … HV01) -IHV12 HV01 #V3 #HV32 #HV03 - elim (IHT12 … HT01) -IHT12 HT01 - [3: -HV32 HV03 @(drop_skip … HLK) /2/ |2: skip ] (**) (* /3 width=5/ is too slow *) - -HLK HW01 - /3 width=5/ -| #L #K0 #V1 #V2 #V0 #i #HLK0 #HV12 #HV20 #IHV12 #d #e #K #HLK #X #HX - lapply (lift_inv_lref2 … HX) -HX * * #Hid #HX destruct -X; - [ -HV12; - elim (drop_conf_lt … HLK … HLK0 Hid) -HLK HLK0 L #L #V3 #HKL #HK0L #HV31 - elim (IHV12 … HK0L … HV31) -IHV12 HK0L HV31 #V4 #HV42 #HV34 - elim (lift_trans_le … HV42 … HV20 ?) -HV42 HV20 V2 // #V2 #HV42 - >arith5 // -Hid #HV20 - @(ex2_1_intro … V2) /2 width=6/ (**) (* /3 width=8/ is slow *) - | -IHV12; - lapply (drop_conf_ge … HLK … HLK0 Hid) -HLK HLK0 L #HK - elim (lift_free … HV20 d (i - e + 1) ? ? ?) -HV20 /2/ - >arith3 /2/ -Hid /3 width=8/ (**) (* just /3 width=8/ is a bit slow *) - ] -| #L #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct -X; - elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct -Y; - elim (IHV12 ? ? ? HLK ? HV01) -IHV12 HV01 #V3 #HV32 #HV03 - elim (IHW12 ? ? ? HLK ? HW01) -IHW12 #W3 #HW32 #HW03 - elim (IHT12 … HT01) -IHT12 HT01 - [3: -HV2 HV32 HV03 HW32 HW03 @(drop_skip … HLK) /2/ |2: skip ] (**) (* /3/ is too slow *) - -HLK HW01 #T3 #HT32 #HT03 - elim (lift_trans_le … HV32 … HV2 ?) -HV32 HV2 V2 // #V2 #HV32 #HV2 - @(ex2_1_intro … (𝕓{Abbr}W3.𝕗{Appl}V2.T3)) /3/ (**) (* /4/ loops *) -| #L #V #T #T1 #T2 #HT1 #_ #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_bind2 … HX) -HX #V0 #T0 #_ #HT0 #H destruct -X; - elim (lift_div_le … HT1 … HT0 ?) -HT1 HT0 T // #T #HT0 #HT1 - elim (IHT12 … HLK … HT1) -IHT12 HLK HT1 /3 width=5/ -| #L #V #T1 #T2 #_ #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #T0 #_ #HT01 #H destruct -X; - elim (IHT12 … HLK … HT01) -IHT12 HLK HT01 /3/ -] -qed. -*) \ No newline at end of file diff --git a/matita/matita/lib/lambda-delta/reduction/tpr_pts.ma b/matita/matita/lib/lambda-delta/reduction/tpr_pts.ma index ea6b8027b..f8a41f025 100644 --- a/matita/matita/lib/lambda-delta/reduction/tpr_pts.ma +++ b/matita/matita/lib/lambda-delta/reduction/tpr_pts.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "lambda-delta/reduction/lpr_defs.ma". +include "lambda-delta/reduction/lpr.ma". (* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) diff --git a/matita/matita/lib/lambda-delta/reduction/tpr_tpr.ma b/matita/matita/lib/lambda-delta/reduction/tpr_tpr.ma index 53ad09398..b5743ddd1 100644 --- a/matita/matita/lib/lambda-delta/reduction/tpr_tpr.ma +++ b/matita/matita/lib/lambda-delta/reduction/tpr_tpr.ma @@ -11,7 +11,7 @@ include "lambda-delta/substitution/lift_weight.ma". include "lambda-delta/substitution/pts_pts.ma". -include "lambda-delta/reduction/tpr_main.ma". +include "lambda-delta/reduction/tpr_lift.ma". include "lambda-delta/reduction/tpr_pts.ma". (* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) diff --git a/matita/matita/lib/lambda-delta/substitution/drop.ma b/matita/matita/lib/lambda-delta/substitution/drop.ma new file mode 100644 index 000000000..d192c84dc --- /dev/null +++ b/matita/matita/lib/lambda-delta/substitution/drop.ma @@ -0,0 +1,147 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "lambda-delta/substitution/leq.ma". +include "lambda-delta/substitution/lift.ma". + +(* DROPPING *****************************************************************) + +inductive drop: lenv → nat → nat → lenv → Prop ≝ +| drop_sort: ∀d,e. drop (⋆) d e (⋆) +| drop_comp: ∀L1,L2,I,V. drop L1 0 0 L2 → drop (L1. 𝕓{I} V) 0 0 (L2. 𝕓{I} V) +| drop_drop: ∀L1,L2,I,V,e. drop L1 0 e L2 → drop (L1. 𝕓{I} V) 0 (e + 1) L2 +| drop_skip: ∀L1,L2,I,V1,V2,d,e. + drop L1 d e L2 → ↑[d,e] V2 ≡ V1 → + drop (L1. 𝕓{I} V1) (d + 1) e (L2. 𝕓{I} V2) +. + +interpretation "dropping" 'RDrop L1 d e L2 = (drop L1 d e L2). + +(* Basic inversion lemmas ***************************************************) + +lemma drop_inv_refl_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → d = 0 → e = 0 → L1 = L2. +#d #e #L1 #L2 #H elim H -H d e L1 L2 +[ // +| #L1 #L2 #I #V #_ #IHL12 #H1 #H2 + >(IHL12 H1 H2) -IHL12 H1 H2 L1 // +| #L1 #L2 #I #V #e #_ #_ #_ #H + elim (plus_S_eq_O_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #_ #H + elim (plus_S_eq_O_false … H) +] +qed. + +lemma drop_inv_refl: ∀L1,L2. ↓[0, 0] L1 ≡ L2 → L1 = L2. +/2 width=5/ qed. + +lemma drop_inv_sort1_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → L1 = ⋆ → + L2 = ⋆. +#d #e #L1 #L2 * -d e L1 L2 +[ // +| #L1 #L2 #I #V #_ #H destruct +| #L1 #L2 #I #V #e #_ #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct +] +qed. + +lemma drop_inv_sort1: ∀d,e,L2. ↓[d, e] ⋆ ≡ L2 → L2 = ⋆. +/2 width=5/ qed. + +lemma drop_inv_O1_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → d = 0 → + ∀K,I,V. L1 = K. 𝕓{I} V → + (e = 0 ∧ L2 = K. 𝕓{I} V) ∨ + (0 < e ∧ ↓[d, e - 1] K ≡ L2). +#d #e #L1 #L2 * -d e L1 L2 +[ #d #e #_ #K #I #V #H destruct +| #L1 #L2 #I #V #HL12 #H #K #J #W #HX destruct -L1 I V + >(drop_inv_refl … HL12) -HL12 K /3/ +| #L1 #L2 #I #V #e #HL12 #_ #K #J #W #H destruct -L1 I V /3/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H elim (plus_S_eq_O_false … H) +] +qed. + +lemma drop_inv_O1: ∀e,K,I,V,L2. ↓[0, e] K. 𝕓{I} V ≡ L2 → + (e = 0 ∧ L2 = K. 𝕓{I} V) ∨ + (0 < e ∧ ↓[0, e - 1] K ≡ L2). +/2/ qed. + +lemma drop_inv_drop1: ∀e,K,I,V,L2. + ↓[0, e] K. 𝕓{I} V ≡ L2 → 0 < e → ↓[0, e - 1] K ≡ L2. +#e #K #I #V #L2 #H #He +elim (drop_inv_O1 … H) -H * // #H destruct -e; +elim (lt_refl_false … He) +qed. + +lemma drop_inv_skip2_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → 0 < d → + ∀I,K2,V2. L2 = K2. 𝕓{I} V2 → + ∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 & + ↑[d - 1, e] V2 ≡ V1 & + L1 = K1. 𝕓{I} V1. +#d #e #L1 #L2 * -d e L1 L2 +[ #d #e #_ #I #K #V #H destruct +| #L1 #L2 #I #V #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H) +| #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #I #L2 #V2 #H destruct -X Y Z; + /2 width=5/ +] +qed. + +lemma drop_inv_skip2: ∀d,e,I,L1,K2,V2. ↓[d, e] L1 ≡ K2. 𝕓{I} V2 → 0 < d → + ∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 & ↑[d - 1, e] V2 ≡ V1 & + L1 = K1. 𝕓{I} V1. +/2/ qed. + +(* Basic properties *********************************************************) + +lemma drop_refl: ∀L. ↓[0, 0] L ≡ L. +#L elim L -L /2/ +qed. + +lemma drop_drop_lt: ∀L1,L2,I,V,e. + ↓[0, e - 1] L1 ≡ L2 → 0 < e → ↓[0, e] L1. 𝕓{I} V ≡ L2. +#L1 #L2 #I #V #e #HL12 #He >(plus_minus_m_m e 1) /2/ +qed. + +lemma drop_fwd_drop2: ∀L1,I2,K2,V2,e. ↓[O, e] L1 ≡ K2. 𝕓{I2} V2 → + ↓[O, e + 1] L1 ≡ K2. +#L1 elim L1 -L1 +[ #I2 #K2 #V2 #e #H lapply (drop_inv_sort1 … H) -H #H destruct +| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H + elim (drop_inv_O1 … H) -H * #He #H + [ -IHL1; destruct -e K2 I2 V2 /2/ + | @drop_drop >(plus_minus_m_m e 1) /2/ + ] +] +qed. + +lemma drop_leq_drop1: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → + ∀I,K1,V,i. ↓[0, i] L1 ≡ K1. 𝕓{I} V → + d ≤ i → i < d + e → + ∃∃K2. K1 [0, d + e - i - 1] ≈ K2 & + ↓[0, i] L2 ≡ K2. 𝕓{I} V. +#L1 #L2 #d #e #H elim H -H L1 L2 d e +[ #d #e #I #K1 #V #i #H + lapply (drop_inv_sort1 … H) -H #H destruct +| #L1 #L2 #I1 #I2 #V1 #V2 #_ #_ #I #K1 #V #i #_ #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #I #V #e #HL12 #IHL12 #J #K1 #W #i #H #_ #Hie + elim (drop_inv_O1 … H) -H * #Hi #HLK1 + [ -IHL12 Hie; destruct -i K1 J W; + arith_g1 // /3/ + ] +| #L1 #L2 #I1 #I2 #V1 #V2 #d #e #_ #IHL12 #I #K1 #V #i #H #Hdi >plus_plus_comm_23 #Hide + lapply (plus_S_le_to_pos … Hdi) #Hi + lapply (drop_inv_drop1 … H ?) -H // #HLK1 + elim (IHL12 … HLK1 ? ?) -IHL12 HLK1 [2: /2/ |3: /2/ ] -Hdi Hide >arith_g1 // /3/ +] +qed. diff --git a/matita/matita/lib/lambda-delta/substitution/drop_defs.ma b/matita/matita/lib/lambda-delta/substitution/drop_defs.ma deleted file mode 100644 index 51c657615..000000000 --- a/matita/matita/lib/lambda-delta/substitution/drop_defs.ma +++ /dev/null @@ -1,137 +0,0 @@ -(* - ||M|| This file is part of HELM, an Hypertextual, Electronic - ||A|| Library of Mathematics, developed at the Computer Science - ||T|| Department of the University of Bologna, Italy. - ||I|| - ||T|| - ||A|| This file is distributed under the terms of the - \ / GNU General Public License Version 2 - \ / - V_______________________________________________________________ *) - -include "lambda-delta/substitution/leq_defs.ma". -include "lambda-delta/substitution/lift_defs.ma". - -(* DROPPING *****************************************************************) - -inductive drop: lenv → nat → nat → lenv → Prop ≝ -| drop_refl: ∀L. drop L 0 0 L -| drop_drop: ∀L1,L2,I,V,e. drop L1 0 e L2 → drop (L1. 𝕓{I} V) 0 (e + 1) L2 -| drop_skip: ∀L1,L2,I,V1,V2,d,e. - drop L1 d e L2 → ↑[d,e] V2 ≡ V1 → - drop (L1. 𝕓{I} V1) (d + 1) e (L2. 𝕓{I} V2) -. - -interpretation "dropping" 'RDrop L1 d e L2 = (drop L1 d e L2). - -(* Basic inversion lemmas ***************************************************) - -lemma drop_inv_refl_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → d = 0 → e = 0 → L1 = L2. -#d #e #L1 #L2 * -d e L1 L2 -[ // -| #L1 #L2 #I #V #e #_ #_ #H - elim (plus_S_eq_O_false … H) -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H - elim (plus_S_eq_O_false … H) -] -qed. - -lemma drop_inv_refl: ∀L1,L2. ↓[0, 0] L1 ≡ L2 → L1 = L2. -/2 width=5/ qed. - -lemma drop_inv_sort1_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → L1 = ⋆ → - ∧∧ L2 = ⋆ & d = 0 & e = 0. -#d #e #L1 #L2 * -d e L1 L2 -[ /2/ -| #L1 #L2 #I #V #e #_ #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct -] -qed. - -lemma drop_inv_sort1: ∀d,e,L2. ↓[d, e] ⋆ ≡ L2 → - ∧∧ L2 = ⋆ & d = 0 & e = 0. -/2/ qed. - -lemma drop_inv_O1_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → d = 0 → - ∀K,I,V. L1 = K. 𝕓{I} V → - (e = 0 ∧ L2 = K. 𝕓{I} V) ∨ - (0 < e ∧ ↓[d, e - 1] K ≡ L2). -#d #e #L1 #L2 * -d e L1 L2 -[ /3/ -| #L1 #L2 #I #V #e #HL12 #_ #K #J #W #H destruct -L1 I V /3/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H elim (plus_S_eq_O_false … H) -] -qed. - -lemma drop_inv_O1: ∀e,K,I,V,L2. ↓[0, e] K. 𝕓{I} V ≡ L2 → - (e = 0 ∧ L2 = K. 𝕓{I} V) ∨ - (0 < e ∧ ↓[0, e - 1] K ≡ L2). -/2/ qed. - -lemma drop_inv_drop1: ∀e,K,I,V,L2. - ↓[0, e] K. 𝕓{I} V ≡ L2 → 0 < e → ↓[0, e - 1] K ≡ L2. -#e #K #I #V #L2 #H #He -elim (drop_inv_O1 … H) -H * // #H destruct -e; -elim (lt_refl_false … He) -qed. - -lemma drop_inv_skip2_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → 0 < d → - ∀I,K2,V2. L2 = K2. 𝕓{I} V2 → - ∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 & - ↑[d - 1, e] V2 ≡ V1 & - L1 = K1. 𝕓{I} V1. -#d #e #L1 #L2 * -d e L1 L2 -[ #L #H elim (lt_refl_false … H) -| #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H) -| #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #I #L2 #V2 #H destruct -X Y Z; - /2 width=5/ -] -qed. - -lemma drop_inv_skip2: ∀d,e,I,L1,K2,V2. ↓[d, e] L1 ≡ K2. 𝕓{I} V2 → 0 < d → - ∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 & ↑[d - 1, e] V2 ≡ V1 & - L1 = K1. 𝕓{I} V1. -/2/ qed. - -(* Basic properties *********************************************************) - -lemma drop_drop_lt: ∀L1,L2,I,V,e. - ↓[0, e - 1] L1 ≡ L2 → 0 < e → ↓[0, e] L1. 𝕓{I} V ≡ L2. -#L1 #L2 #I #V #e #HL12 #He >(plus_minus_m_m e 1) /2/ -qed. - -lemma drop_fwd_drop2: ∀L1,I2,K2,V2,e. ↓[O, e] L1 ≡ K2. 𝕓{I2} V2 → - ↓[O, e + 1] L1 ≡ K2. -#L1 elim L1 -L1 -[ #I2 #K2 #V2 #e #H elim (drop_inv_sort1 … H) -H #H destruct -| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H - elim (drop_inv_O1 … H) -H * #He #H - [ -IHL1; destruct -e K2 I2 V2 /2/ - | @drop_drop >(plus_minus_m_m e 1) /2/ - ] -] -qed. - -lemma drop_leq_drop1: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → - ∀I,K1,V,i. ↓[0, i] L1 ≡ K1. 𝕓{I} V → - d ≤ i → i < d + e → - ∃∃K2. K1 [0, d + e - i - 1] ≈ K2 & - ↓[0, i] L2 ≡ K2. 𝕓{I} V. -#L1 #L2 #d #e #H elim H -H L1 L2 d e -[ #d #e #I #K1 #V #i #H - elim (drop_inv_sort1 … H) -H #H destruct -| #L1 #L2 #I1 #I2 #V1 #V2 #_ #_ #I #K1 #V #i #_ #_ #H - elim (lt_zero_false … H) -| #L1 #L2 #I #V #e #HL12 #IHL12 #J #K1 #W #i #H #_ #Hie - elim (drop_inv_O1 … H) -H * #Hi #HLK1 - [ -IHL12 Hie; destruct -i K1 J W; - arith_g1 // /3/ - ] -| #L1 #L2 #I1 #I2 #V1 #V2 #d #e #_ #IHL12 #I #K1 #V #i #H #Hdi >plus_plus_comm_23 #Hide - lapply (plus_S_le_to_pos … Hdi) #Hi - lapply (drop_inv_drop1 … H ?) -H // #HLK1 - elim (IHL12 … HLK1 ? ?) -IHL12 HLK1 [2: /2/ |3: /2/ ] -Hdi Hide >arith_g1 // /3/ -] -qed. diff --git a/matita/matita/lib/lambda-delta/substitution/drop_drop.ma b/matita/matita/lib/lambda-delta/substitution/drop_drop.ma new file mode 100644 index 000000000..96fd0db51 --- /dev/null +++ b/matita/matita/lib/lambda-delta/substitution/drop_drop.ma @@ -0,0 +1,107 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "lambda-delta/substitution/drop.ma". + +(* DROPPING *****************************************************************) + +(* Main properties **********************************************************) + +lemma drop_conf_ge: ∀d1,e1,L,L1. ↓[d1, e1] L ≡ L1 → + ∀e2,L2. ↓[0, e2] L ≡ L2 → d1 + e1 ≤ e2 → + ↓[0, e2 - e1] L1 ≡ L2. +#d1 #e1 #L #L1 #H elim H -H d1 e1 L L1 +[ #d #e #e2 #L2 #H + >(drop_inv_sort1 … H) -H L2 // +| #K1 #K2 #I #V #HK12 #_ #e2 #L2 #H #_ minus_minus_comm /3/ (**) (* explicit constructor *) +] +qed. + +lemma drop_conf_lt: ∀d1,e1,L,L1. ↓[d1, e1] L ≡ L1 → + ∀e2,K2,I,V2. ↓[0, e2] L ≡ K2. 𝕓{I} V2 → + e2 < d1 → let d ≝ d1 - e2 - 1 in + ∃∃K1,V1. ↓[0, e2] L1 ≡ K1. 𝕓{I} V1 & + ↓[d, e1] K2 ≡ K1 & ↑[d, e1] V1 ≡ V2. +#d1 #e1 #L #L1 #H elim H -H d1 e1 L L1 +[ #d #e #e2 #K2 #I #V2 #H + lapply (drop_inv_sort1 … H) -H #H destruct +| #L1 #L2 #I #V #_ #_ #e2 #K2 #J #V2 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #I #V #e #_ #_ #e2 #K2 #J #V2 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #K2 #J #V #H #He2d + elim (drop_inv_O1 … H) -H * + [ -IHL12 He2d #H1 #H2 destruct -e2 K2 J V /2 width=5/ + | -HL12 -HV12 #He #HLK + elim (IHL12 … HLK ?) -IHL12 HLK [ (drop_inv_sort1 … H) -H L2 /2/ +| #K1 #K2 #I #V #HK12 #_ #e2 #L2 #HL2 #H + >(drop_inv_refl … HK12) -HK12 K1; + lapply (le_O_to_eq_O … H) -H #H destruct -e2 /2/ +| #L1 #L2 #I #V #e #_ #IHL12 #e2 #L #HL2 #H + lapply (le_O_to_eq_O … H) -H #H destruct -e2; + elim (IHL12 … HL2 ?) -IHL12 HL2 // #L0 #H #HL0 + lapply (drop_inv_refl … H) -H #H destruct -L1 /3 width=5/ +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #L #H #He2d + elim (drop_inv_O1 … H) -H * + [ -He2d IHL12 #H1 #H2 destruct -e2 L /3 width=5/ + | -HL12 HV12 #He2 #HL2 + elim (IHL12 … HL2 ?) -IHL12 HL2 L2 + [ (drop_inv_sort1 … H) -H L2 // +| #K1 #K2 #I #V #HK12 #_ #e2 #L2 #H #_ normalize + >(drop_inv_refl … HK12) -HK12 K1 // +| /3/ +| #L1 #L2 #I #V1 #V2 #d #e #H_ #_ #IHL12 #e2 #L #H #Hde2 + lapply (lt_to_le_to_lt 0 … Hde2) // #He2 + lapply (lt_to_le_to_lt … (e + e2) He2 ?) // #Hee2 + lapply (drop_inv_drop1 … H ?) -H // #HL2 + @drop_drop_lt // >le_plus_minus // @IHL12 /2/ (**) (* explicit constructor *) +] +qed. + +lemma drop_trans_ge_comm: ∀d1,e1,e2,L1,L2,L. + ↓[d1, e1] L1 ≡ L → ↓[0, e2] L ≡ L2 → d1 ≤ e2 → + ↓[0, e2 + e1] L1 ≡ L2. +#e1 #e1 #e2 >commutative_plus /2 width=5/ +qed. + +axiom drop_div: ∀e1,L1,L. ↓[0, e1] L1 ≡ L → ∀e2,L2. ↓[0, e2] L2 ≡ L → + ∃∃L0. ↓[0, e1] L0 ≡ L2 & ↓[e1, e2] L0 ≡ L1. diff --git a/matita/matita/lib/lambda-delta/substitution/drop_main.ma b/matita/matita/lib/lambda-delta/substitution/drop_main.ma deleted file mode 100644 index 575f25401..000000000 --- a/matita/matita/lib/lambda-delta/substitution/drop_main.ma +++ /dev/null @@ -1,96 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "lambda-delta/substitution/drop_defs.ma". - -(* DROPPING *****************************************************************) - -(* Main properties **********************************************************) - -lemma drop_conf_ge: ∀d1,e1,L,L1. ↓[d1, e1] L ≡ L1 → - ∀e2,L2. ↓[0, e2] L ≡ L2 → d1 + e1 ≤ e2 → - ↓[0, e2 - e1] L1 ≡ L2. -#d1 #e1 #L #L1 #H elim H -H d1 e1 L L1 -[ // -| #L #K #I #V #e #_ #IHLK #e2 #L2 #H #He2 - lapply (drop_inv_drop1 … H ?) -H /2/ #HL2 - minus_minus_comm /3/ (**) (* explicit constructor *) -] -qed. - -lemma drop_conf_lt: ∀d1,e1,L,L1. ↓[d1, e1] L ≡ L1 → - ∀e2,K2,I,V2. ↓[0, e2] L ≡ K2. 𝕓{I} V2 → - e2 < d1 → let d ≝ d1 - e2 - 1 in - ∃∃K1,V1. ↓[0, e2] L1 ≡ K1. 𝕓{I} V1 & - ↓[d, e1] K2 ≡ K1 & ↑[d, e1] V1 ≡ V2. -#d1 #e1 #L #L1 #H elim H -H d1 e1 L L1 -[ #L0 #e2 #K2 #I #V2 #_ #H - elim (lt_zero_false … H) -| #L1 #L2 #I #V #e #_ #_ #e2 #K2 #J #V2 #_ #H - elim (lt_zero_false … H) -| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #K2 #J #V #H #He2d - elim (drop_inv_O1 … H) -H * - [ -IHL12 He2d #H1 #H2 destruct -e2 K2 J V /2 width=5/ - | -HL12 -HV12 #He #HLK - elim (IHL12 … HLK ?) -IHL12 HLK [ le_plus_minus // @IHL12 /2/ (**) (* explicit constructor *) -] -qed. - -lemma drop_trans_ge_comm: ∀d1,e1,e2,L1,L2,L. - ↓[d1, e1] L1 ≡ L → ↓[0, e2] L ≡ L2 → d1 ≤ e2 → - ↓[0, e2 + e1] L1 ≡ L2. -#e1 #e1 #e2 >commutative_plus /2 width=5/ -qed. - -axiom drop_div: ∀e1,L1,L. ↓[0, e1] L1 ≡ L → ∀e2,L2. ↓[0, e2] L2 ≡ L → - ∃∃L0. ↓[0, e1] L0 ≡ L2 & ↓[e1, e2] L0 ≡ L1. diff --git a/matita/matita/lib/lambda-delta/substitution/leq.ma b/matita/matita/lib/lambda-delta/substitution/leq.ma new file mode 100644 index 000000000..3c7981566 --- /dev/null +++ b/matita/matita/lib/lambda-delta/substitution/leq.ma @@ -0,0 +1,60 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "lambda-delta/syntax/lenv.ma". + +(* LOCAL ENVIRONMENT EQUALITY ***********************************************) + +inductive leq: lenv → nat → nat → lenv → Prop ≝ +| leq_sort: ∀d,e. leq (⋆) d e (⋆) +| leq_comp: ∀L1,L2,I1,I2,V1,V2. + leq L1 0 0 L2 → leq (L1. 𝕓{I1} V1) 0 0 (L2. 𝕓{I2} V2) +| leq_eq: ∀L1,L2,I,V,e. leq L1 0 e L2 → leq (L1. 𝕓{I} V) 0 (e + 1) (L2.𝕓{I} V) +| leq_skip: ∀L1,L2,I1,I2,V1,V2,d,e. + leq L1 d e L2 → leq (L1. 𝕓{I1} V1) (d + 1) e (L2. 𝕓{I2} V2) +. + +interpretation "local environment equality" 'Eq L1 d e L2 = (leq L1 d e L2). + +(* Basic properties *********************************************************) + +lemma leq_refl: ∀d,e,L. L [d, e] ≈ L. +#d elim d -d +[ #e elim e -e [ #L elim L -L /2/ | #e #IHe #L elim L -L /2/ ] +| #d #IHd #e #L elim L -L /2/ +] +qed. + +lemma leq_sym: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → L2 [d, e] ≈ L1. +#L1 #L2 #d #e #H elim H -H L1 L2 d e /2/ +qed. + +lemma leq_skip_lt: ∀L1,L2,d,e. leq L1 (d - 1) e L2 → 0 < d → + ∀I1,I2,V1,V2. L1. 𝕓{I1} V1 [d, e] ≈ L2. 𝕓{I2} V2. + +#L1 #L2 #d #e #HL12 #Hd >(plus_minus_m_m d 1) /2/ +qed. + +(* Basic inversion lemmas ***************************************************) + +lemma leq_inv_sort1_aux: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → L1 = ⋆ → L2 = ⋆. +#L1 #L2 #d #e #H elim H -H L1 L2 d e +[ // +| #L1 #L2 #I1 #I2 #V1 #V2 #_ #_ #H destruct +| #L1 #L2 #I #V #e #_ #_ #H destruct +| #L1 #L2 #I1 #I2 #V1 #V2 #d #e #_ #_ #H destruct +qed. + +lemma leq_inv_sort1: ∀L2,d,e. ⋆ [d, e] ≈ L2 → L2 = ⋆. +/2 width=5/ qed. + +lemma leq_inv_sort2: ∀L1,d,e. L1 [d, e] ≈ ⋆ → L1 = ⋆. +/3/ qed. diff --git a/matita/matita/lib/lambda-delta/substitution/leq_defs.ma b/matita/matita/lib/lambda-delta/substitution/leq_defs.ma deleted file mode 100644 index 3c7981566..000000000 --- a/matita/matita/lib/lambda-delta/substitution/leq_defs.ma +++ /dev/null @@ -1,60 +0,0 @@ -(* - ||M|| This file is part of HELM, an Hypertextual, Electronic - ||A|| Library of Mathematics, developed at the Computer Science - ||T|| Department of the University of Bologna, Italy. - ||I|| - ||T|| - ||A|| This file is distributed under the terms of the - \ / GNU General Public License Version 2 - \ / - V_______________________________________________________________ *) - -include "lambda-delta/syntax/lenv.ma". - -(* LOCAL ENVIRONMENT EQUALITY ***********************************************) - -inductive leq: lenv → nat → nat → lenv → Prop ≝ -| leq_sort: ∀d,e. leq (⋆) d e (⋆) -| leq_comp: ∀L1,L2,I1,I2,V1,V2. - leq L1 0 0 L2 → leq (L1. 𝕓{I1} V1) 0 0 (L2. 𝕓{I2} V2) -| leq_eq: ∀L1,L2,I,V,e. leq L1 0 e L2 → leq (L1. 𝕓{I} V) 0 (e + 1) (L2.𝕓{I} V) -| leq_skip: ∀L1,L2,I1,I2,V1,V2,d,e. - leq L1 d e L2 → leq (L1. 𝕓{I1} V1) (d + 1) e (L2. 𝕓{I2} V2) -. - -interpretation "local environment equality" 'Eq L1 d e L2 = (leq L1 d e L2). - -(* Basic properties *********************************************************) - -lemma leq_refl: ∀d,e,L. L [d, e] ≈ L. -#d elim d -d -[ #e elim e -e [ #L elim L -L /2/ | #e #IHe #L elim L -L /2/ ] -| #d #IHd #e #L elim L -L /2/ -] -qed. - -lemma leq_sym: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → L2 [d, e] ≈ L1. -#L1 #L2 #d #e #H elim H -H L1 L2 d e /2/ -qed. - -lemma leq_skip_lt: ∀L1,L2,d,e. leq L1 (d - 1) e L2 → 0 < d → - ∀I1,I2,V1,V2. L1. 𝕓{I1} V1 [d, e] ≈ L2. 𝕓{I2} V2. - -#L1 #L2 #d #e #HL12 #Hd >(plus_minus_m_m d 1) /2/ -qed. - -(* Basic inversion lemmas ***************************************************) - -lemma leq_inv_sort1_aux: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → L1 = ⋆ → L2 = ⋆. -#L1 #L2 #d #e #H elim H -H L1 L2 d e -[ // -| #L1 #L2 #I1 #I2 #V1 #V2 #_ #_ #H destruct -| #L1 #L2 #I #V #e #_ #_ #H destruct -| #L1 #L2 #I1 #I2 #V1 #V2 #d #e #_ #_ #H destruct -qed. - -lemma leq_inv_sort1: ∀L2,d,e. ⋆ [d, e] ≈ L2 → L2 = ⋆. -/2 width=5/ qed. - -lemma leq_inv_sort2: ∀L1,d,e. L1 [d, e] ≈ ⋆ → L1 = ⋆. -/3/ qed. diff --git a/matita/matita/lib/lambda-delta/substitution/lift.ma b/matita/matita/lib/lambda-delta/substitution/lift.ma new file mode 100644 index 000000000..a5b15e110 --- /dev/null +++ b/matita/matita/lib/lambda-delta/substitution/lift.ma @@ -0,0 +1,230 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "lambda-delta/syntax/term.ma". + +(* RELOCATION ***************************************************************) + +inductive lift: term → nat → nat → term → Prop ≝ +| lift_sort : ∀k,d,e. lift (⋆k) d e (⋆k) +| lift_lref_lt: ∀i,d,e. i < d → lift (#i) d e (#i) +| lift_lref_ge: ∀i,d,e. d ≤ i → lift (#i) d e (#(i + e)) +| lift_bind : ∀I,V1,V2,T1,T2,d,e. + lift V1 d e V2 → lift T1 (d + 1) e T2 → + lift (𝕓{I} V1. T1) d e (𝕓{I} V2. T2) +| lift_flat : ∀I,V1,V2,T1,T2,d,e. + lift V1 d e V2 → lift T1 d e T2 → + lift (𝕗{I} V1. T1) d e (𝕗{I} V2. T2) +. + +interpretation "relocation" 'RLift T1 d e T2 = (lift T1 d e T2). + +(* Basic properties *********************************************************) + +lemma lift_lref_ge_minus: ∀d,e,i. d + e ≤ i → ↑[d, e] #(i - e) ≡ #i. +#d #e #i #H >(plus_minus_m_m i e) in ⊢ (? ? ? ? %) /3/ +qed. + +lemma lift_refl: ∀T,d. ↑[d, 0] T ≡ T. +#T elim T -T +[ // +| #i #d elim (lt_or_ge i d) /2/ +| #I elim I -I /2/ +] +qed. + +lemma lift_total: ∀T1,d,e. ∃T2. ↑[d,e] T1 ≡ T2. +#T1 elim T1 -T1 +[ /2/ +| #i #d #e elim (lt_or_ge i d) /3/ +| * #I #V1 #T1 #IHV1 #IHT1 #d #e + elim (IHV1 d e) -IHV1 #V2 #HV12 + [ elim (IHT1 (d+1) e) -IHT1 /3/ + | elim (IHT1 d e) -IHT1 /3/ + ] +] +qed. + +lemma lift_split: ∀d1,e2,T1,T2. ↑[d1, e2] T1 ≡ T2 → ∀d2,e1. + d1 ≤ d2 → d2 ≤ d1 + e1 → e1 ≤ e2 → + ∃∃T. ↑[d1, e1] T1 ≡ T & ↑[d2, e2 - e1] T ≡ T2. +#d1 #e2 #T1 #T2 #H elim H -H d1 e2 T1 T2 +[ /3/ +| #i #d1 #e2 #Hid1 #d2 #e1 #Hd12 #_ #_ + lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 /4/ +| #i #d1 #e2 #Hid1 #d2 #e1 #_ #Hd21 #He12 + lapply (transitive_le …(i+e1) Hd21 ?) /2/ -Hd21 #Hd21 + <(arith_d1 i e2 e1) // /3/ +| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12 + elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b + elim (IHT (d2+1) … ? ? He12) /3 width = 5/ +| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12 + elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b + elim (IHT d2 … ? ? He12) /3 width = 5/ +] +qed. + +(* Basic inversion lemmas ***************************************************) + +lemma lift_inv_refl_aux: ∀d,e,T1,T2. ↑[d, e] T1 ≡ T2 → e = 0 → T1 = T2. +#d #e #T1 #T2 #H elim H -H d e T1 T2 /3/ +qed. + +lemma lift_inv_refl: ∀d,T1,T2. ↑[d, 0] T1 ≡ T2 → T1 = T2. +/2/ qed. + +lemma lift_inv_sort1_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → ∀k. T1 = ⋆k → T2 = ⋆k. +#d #e #T1 #T2 * -d e T1 T2 // +[ #i #d #e #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct +] +qed. + +lemma lift_inv_sort1: ∀d,e,T2,k. ↑[d,e] ⋆k ≡ T2 → T2 = ⋆k. +/2 width=5/ qed. + +lemma lift_inv_lref1_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → ∀i. T1 = #i → + (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)). +#d #e #T1 #T2 * -d e T1 T2 +[ #k #d #e #i #H destruct +| #j #d #e #Hj #i #Hi destruct /3/ +| #j #d #e #Hj #i #Hi destruct /3/ +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct +] +qed. + +lemma lift_inv_lref1: ∀d,e,T2,i. ↑[d,e] #i ≡ T2 → + (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)). +/2/ qed. + +lemma lift_inv_lref1_lt: ∀d,e,T2,i. ↑[d,e] #i ≡ T2 → i < d → T2 = #i. +#d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * // +#Hdi #_ #Hid lapply (le_to_lt_to_lt … Hdi Hid) -Hdi Hid #Hdd +elim (lt_refl_false … Hdd) +qed. + +lemma lift_inv_lref1_ge: ∀d,e,T2,i. ↑[d,e] #i ≡ T2 → d ≤ i → T2 = #(i + e). +#d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * // +#Hid #_ #Hdi lapply (le_to_lt_to_lt … Hdi Hid) -Hdi Hid #Hdd +elim (lt_refl_false … Hdd) +qed. + +lemma lift_inv_bind1_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → + ∀I,V1,U1. T1 = 𝕓{I} V1.U1 → + ∃∃V2,U2. ↑[d,e] V1 ≡ V2 & ↑[d+1,e] U1 ≡ U2 & + T2 = 𝕓{I} V2. U2. +#d #e #T1 #T2 * -d e T1 T2 +[ #k #d #e #I #V1 #U1 #H destruct +| #i #d #e #_ #I #V1 #U1 #H destruct +| #i #d #e #_ #I #V1 #U1 #H destruct +| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/ +| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct +] +qed. + +lemma lift_inv_bind1: ∀d,e,T2,I,V1,U1. ↑[d,e] 𝕓{I} V1. U1 ≡ T2 → + ∃∃V2,U2. ↑[d,e] V1 ≡ V2 & ↑[d+1,e] U1 ≡ U2 & + T2 = 𝕓{I} V2. U2. +/2/ qed. + +lemma lift_inv_flat1_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → + ∀I,V1,U1. T1 = 𝕗{I} V1.U1 → + ∃∃V2,U2. ↑[d,e] V1 ≡ V2 & ↑[d,e] U1 ≡ U2 & + T2 = 𝕗{I} V2. U2. +#d #e #T1 #T2 * -d e T1 T2 +[ #k #d #e #I #V1 #U1 #H destruct +| #i #d #e #_ #I #V1 #U1 #H destruct +| #i #d #e #_ #I #V1 #U1 #H destruct +| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct +| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/ +] +qed. + +lemma lift_inv_flat1: ∀d,e,T2,I,V1,U1. ↑[d,e] 𝕗{I} V1. U1 ≡ T2 → + ∃∃V2,U2. ↑[d,e] V1 ≡ V2 & ↑[d,e] U1 ≡ U2 & + T2 = 𝕗{I} V2. U2. +/2/ qed. + +lemma lift_inv_sort2_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → ∀k. T2 = ⋆k → T1 = ⋆k. +#d #e #T1 #T2 * -d e T1 T2 // +[ #i #d #e #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct +] +qed. + +lemma lift_inv_sort2: ∀d,e,T1,k. ↑[d,e] T1 ≡ ⋆k → T1 = ⋆k. +/2 width=5/ qed. + +lemma lift_inv_lref2_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → ∀i. T2 = #i → + (i < d ∧ T1 = #i) ∨ (d + e ≤ i ∧ T1 = #(i - e)). +#d #e #T1 #T2 * -d e T1 T2 +[ #k #d #e #i #H destruct +| #j #d #e #Hj #i #Hi destruct /3/ +| #j #d #e #Hj #i #Hi destruct (plus_minus_m_m i e) in ⊢ (? ? ? ? %) /3/ -qed. - -lemma lift_refl: ∀T,d. ↑[d, 0] T ≡ T. -#T elim T -T -[ // -| #i #d elim (lt_or_ge i d) /2/ -| #I elim I -I /2/ -] -qed. - -lemma lift_total: ∀T1,d,e. ∃T2. ↑[d,e] T1 ≡ T2. -#T1 elim T1 -T1 -[ /2/ -| #i #d #e elim (lt_or_ge i d) /3/ -| * #I #V1 #T1 #IHV1 #IHT1 #d #e - elim (IHV1 d e) -IHV1 #V2 #HV12 - [ elim (IHT1 (d+1) e) -IHT1 /3/ - | elim (IHT1 d e) -IHT1 /3/ - ] -] -qed. - -(* Basic inversion lemmas ***************************************************) - -lemma lift_inv_refl_aux: ∀d,e,T1,T2. ↑[d, e] T1 ≡ T2 → e = 0 → T1 = T2. -#d #e #T1 #T2 #H elim H -H d e T1 T2 /3/ -qed. - -lemma lift_inv_refl: ∀d,T1,T2. ↑[d, 0] T1 ≡ T2 → T1 = T2. -/2/ qed. - -lemma lift_inv_sort1_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → ∀k. T1 = ⋆k → T2 = ⋆k. -#d #e #T1 #T2 * -d e T1 T2 // -[ #i #d #e #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct -] -qed. - -lemma lift_inv_sort1: ∀d,e,T2,k. ↑[d,e] ⋆k ≡ T2 → T2 = ⋆k. -/2 width=5/ qed. - -lemma lift_inv_lref1_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → ∀i. T1 = #i → - (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)). -#d #e #T1 #T2 * -d e T1 T2 -[ #k #d #e #i #H destruct -| #j #d #e #Hj #i #Hi destruct /3/ -| #j #d #e #Hj #i #Hi destruct /3/ -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct -] -qed. - -lemma lift_inv_lref1: ∀d,e,T2,i. ↑[d,e] #i ≡ T2 → - (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)). -/2/ qed. - -lemma lift_inv_lref1_lt: ∀d,e,T2,i. ↑[d,e] #i ≡ T2 → i < d → T2 = #i. -#d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * // -#Hdi #_ #Hid lapply (le_to_lt_to_lt … Hdi Hid) -Hdi Hid #Hdd -elim (lt_refl_false … Hdd) -qed. - -lemma lift_inv_lref1_ge: ∀d,e,T2,i. ↑[d,e] #i ≡ T2 → d ≤ i → T2 = #(i + e). -#d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * // -#Hid #_ #Hdi lapply (le_to_lt_to_lt … Hdi Hid) -Hdi Hid #Hdd -elim (lt_refl_false … Hdd) -qed. - -lemma lift_inv_bind1_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → - ∀I,V1,U1. T1 = 𝕓{I} V1.U1 → - ∃∃V2,U2. ↑[d,e] V1 ≡ V2 & ↑[d+1,e] U1 ≡ U2 & - T2 = 𝕓{I} V2. U2. -#d #e #T1 #T2 * -d e T1 T2 -[ #k #d #e #I #V1 #U1 #H destruct -| #i #d #e #_ #I #V1 #U1 #H destruct -| #i #d #e #_ #I #V1 #U1 #H destruct -| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/ -| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct -] -qed. - -lemma lift_inv_bind1: ∀d,e,T2,I,V1,U1. ↑[d,e] 𝕓{I} V1. U1 ≡ T2 → - ∃∃V2,U2. ↑[d,e] V1 ≡ V2 & ↑[d+1,e] U1 ≡ U2 & - T2 = 𝕓{I} V2. U2. -/2/ qed. - -lemma lift_inv_flat1_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → - ∀I,V1,U1. T1 = 𝕗{I} V1.U1 → - ∃∃V2,U2. ↑[d,e] V1 ≡ V2 & ↑[d,e] U1 ≡ U2 & - T2 = 𝕗{I} V2. U2. -#d #e #T1 #T2 * -d e T1 T2 -[ #k #d #e #I #V1 #U1 #H destruct -| #i #d #e #_ #I #V1 #U1 #H destruct -| #i #d #e #_ #I #V1 #U1 #H destruct -| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct -| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/ -] -qed. - -lemma lift_inv_flat1: ∀d,e,T2,I,V1,U1. ↑[d,e] 𝕗{I} V1. U1 ≡ T2 → - ∃∃V2,U2. ↑[d,e] V1 ≡ V2 & ↑[d,e] U1 ≡ U2 & - T2 = 𝕗{I} V2. U2. -/2/ qed. - -lemma lift_inv_sort2_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → ∀k. T2 = ⋆k → T1 = ⋆k. -#d #e #T1 #T2 * -d e T1 T2 // -[ #i #d #e #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct -] -qed. - -lemma lift_inv_sort2: ∀d,e,T1,k. ↑[d,e] T1 ≡ ⋆k → T1 = ⋆k. -/2 width=5/ qed. - -lemma lift_inv_lref2_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → ∀i. T2 = #i → - (i < d ∧ T1 = #i) ∨ (d + e ≤ i ∧ T1 = #(i - e)). -#d #e #T1 #T2 * -d e T1 T2 -[ #k #d #e #i #H destruct -| #j #d #e #Hj #i #Hi destruct /3/ -| #j #d #e #Hj #i #Hi destruct le_plus_minus_comm [ @lift_lref_ge @(transitive_le … Hd12) -Hd12 /2/ | -Hd12 /2/ ] + | -Hd12 >(plus_minus_m_m i e2) in ⊢ (? ? ? ? %) /3/ + ] + ] +| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12 + lapply (lift_inv_bind2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct -T2; + elim (IHW … HW2 ?) // -IHW HW2 #W0 #HW2 #HW1 + >plus_plus_comm_23 in HU2 #HU2 elim (IHU … HU2 ?) /3 width = 5/ +| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12 + lapply (lift_inv_flat2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct -T2; + elim (IHW … HW2 ?) // -IHW HW2 #W0 #HW2 #HW1 + elim (IHU … HU2 ?) /3 width = 5/ +] +qed. + +lemma lift_mono: ∀d,e,T,U1. ↑[d,e] T ≡ U1 → ∀U2. ↑[d,e] T ≡ U2 → U1 = U2. +#d #e #T #U1 #H elim H -H d e T U1 +[ #k #d #e #X #HX + lapply (lift_inv_sort1 … HX) -HX // +| #i #d #e #Hid #X #HX + lapply (lift_inv_lref1_lt … HX ?) -HX // +| #i #d #e #Hdi #X #HX + lapply (lift_inv_lref1_ge … HX ?) -HX // +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX + elim (lift_inv_bind1 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/ +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX + elim (lift_inv_flat1 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/ +] +qed. + +lemma lift_trans_be: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T → + ∀d2,e2,T2. ↑[d2, e2] T ≡ T2 → + d1 ≤ d2 → d2 ≤ d1 + e1 → ↑[d1, e1 + e2] T1 ≡ T2. +#d1 #e1 #T1 #T #H elim H -H d1 e1 T1 T +[ #k #d1 #e1 #d2 #e2 #T2 #HT2 #_ #_ + >(lift_inv_sort1 … HT2) -HT2 // +| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #HT2 #Hd12 #_ + lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 + lapply (lift_inv_lref1_lt … HT2 Hid2) /2/ +| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #HT2 #_ #Hd21 + lapply (lift_inv_lref1_ge … HT2 ?) -HT2 + [ @(transitive_le … Hd21 ?) -Hd21 /2/ + | -Hd21 /2/ + ] +| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21 + elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X; + lapply (IHV12 … HV20 ? ?) // -IHV12 HV20 #HV10 + lapply (IHT12 … HT20 ? ?) /2/ +| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21 + elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X; + lapply (IHV12 … HV20 ? ?) // -IHV12 HV20 #HV10 + lapply (IHT12 … HT20 ? ?) /2/ +] +qed. + +lemma lift_trans_le: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T → + ∀d2,e2,T2. ↑[d2, e2] T ≡ T2 → d2 ≤ d1 → + ∃∃T0. ↑[d2, e2] T1 ≡ T0 & ↑[d1 + e2, e1] T0 ≡ T2. +#d1 #e1 #T1 #T #H elim H -H d1 e1 T1 T +[ #k #d1 #e1 #d2 #e2 #X #HX #_ + >(lift_inv_sort1 … HX) -HX /2/ +| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #_ + lapply (lt_to_le_to_lt … (d1+e2) Hid1 ?) // #Hie2 + elim (lift_inv_lref1 … HX) -HX * #Hid2 #HX destruct -X /4/ +| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #Hd21 + lapply (transitive_le … Hd21 Hid1) -Hd21 #Hid2 + lapply (lift_inv_lref1_ge … HX ?) -HX /2/ #HX destruct -X; + >plus_plus_comm_23 /4/ +| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21 + elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X; + elim (IHV12 … HV20 ?) -IHV12 HV20 // + elim (IHT12 … HT20 ?) -IHT12 HT20 /3 width=5/ +| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21 + elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X; + elim (IHV12 … HV20 ?) -IHV12 HV20 // + elim (IHT12 … HT20 ?) -IHT12 HT20 /3 width=5/ +] +qed. + +lemma lift_trans_ge: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T → + ∀d2,e2,T2. ↑[d2, e2] T ≡ T2 → d1 + e1 ≤ d2 → + ∃∃T0. ↑[d2 - e1, e2] T1 ≡ T0 & ↑[d1, e1] T0 ≡ T2. +#d1 #e1 #T1 #T #H elim H -H d1 e1 T1 T +[ #k #d1 #e1 #d2 #e2 #X #HX #_ + >(lift_inv_sort1 … HX) -HX /2/ +| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #Hded + lapply (lt_to_le_to_lt … (d1+e1) Hid1 ?) // #Hid1e + lapply (lt_to_le_to_lt … (d2-e1) Hid1 ?) /2/ #Hid2e + lapply (lt_to_le_to_lt … Hid1e Hded) -Hid1e Hded #Hid2 + lapply (lift_inv_lref1_lt … HX ?) -HX // #HX destruct -X /3/ +| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #_ + elim (lift_inv_lref1 … HX) -HX * #Hied #HX destruct -X; + [2: >plus_plus_comm_23] /4/ +| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded + elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X; + elim (IHV12 … HV20 ?) -IHV12 HV20 // + elim (IHT12 … HT20 ?) -IHT12 HT20 /2/ #T + le_plus_minus_comm [ @lift_lref_ge @(transitive_le … Hd12) -Hd12 /2/ | -Hd12 /2/ ] - | -Hd12 >(plus_minus_m_m i e2) in ⊢ (? ? ? ? %) /3/ - ] - ] -| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12 - lapply (lift_inv_bind2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct -T2; - elim (IHW … HW2 ?) // -IHW HW2 #W0 #HW2 #HW1 - >plus_plus_comm_23 in HU2 #HU2 elim (IHU … HU2 ?) /3 width = 5/ -| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12 - lapply (lift_inv_flat2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct -T2; - elim (IHW … HW2 ?) // -IHW HW2 #W0 #HW2 #HW1 - elim (IHU … HU2 ?) /3 width = 5/ -] -qed. - -lemma lift_free: ∀d1,e2,T1,T2. ↑[d1, e2] T1 ≡ T2 → ∀d2,e1. - d1 ≤ d2 → d2 ≤ d1 + e1 → e1 ≤ e2 → - ∃∃T. ↑[d1, e1] T1 ≡ T & ↑[d2, e2 - e1] T ≡ T2. -#d1 #e2 #T1 #T2 #H elim H -H d1 e2 T1 T2 -[ /3/ -| #i #d1 #e2 #Hid1 #d2 #e1 #Hd12 #_ #_ - lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 /4/ -| #i #d1 #e2 #Hid1 #d2 #e1 #_ #Hd21 #He12 - lapply (transitive_le …(i+e1) Hd21 ?) /2/ -Hd21 #Hd21 - <(arith_d1 i e2 e1) // /3/ -| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12 - elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b - elim (IHT (d2+1) … ? ? He12) /3 width = 5/ -| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12 - elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b - elim (IHT d2 … ? ? He12) /3 width = 5/ -] -qed. - -lemma lift_mono: ∀d,e,T,U1. ↑[d,e] T ≡ U1 → ∀U2. ↑[d,e] T ≡ U2 → U1 = U2. -#d #e #T #U1 #H elim H -H d e T U1 -[ #k #d #e #X #HX - lapply (lift_inv_sort1 … HX) -HX // -| #i #d #e #Hid #X #HX - lapply (lift_inv_lref1_lt … HX ?) -HX // -| #i #d #e #Hdi #X #HX - lapply (lift_inv_lref1_ge … HX ?) -HX // -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_bind1 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/ -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_flat1 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/ -] -qed. - -lemma lift_trans_be: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T → - ∀d2,e2,T2. ↑[d2, e2] T ≡ T2 → - d1 ≤ d2 → d2 ≤ d1 + e1 → ↑[d1, e1 + e2] T1 ≡ T2. -#d1 #e1 #T1 #T #H elim H -H d1 e1 T1 T -[ #k #d1 #e1 #d2 #e2 #T2 #HT2 #_ #_ - >(lift_inv_sort1 … HT2) -HT2 // -| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #HT2 #Hd12 #_ - lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 - lapply (lift_inv_lref1_lt … HT2 Hid2) /2/ -| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #HT2 #_ #Hd21 - lapply (lift_inv_lref1_ge … HT2 ?) -HT2 - [ @(transitive_le … Hd21 ?) -Hd21 /2/ - | -Hd21 /2/ - ] -| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21 - elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X; - lapply (IHV12 … HV20 ? ?) // -IHV12 HV20 #HV10 - lapply (IHT12 … HT20 ? ?) /2/ -| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21 - elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X; - lapply (IHV12 … HV20 ? ?) // -IHV12 HV20 #HV10 - lapply (IHT12 … HT20 ? ?) /2/ -] -qed. - -lemma lift_trans_le: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T → - ∀d2,e2,T2. ↑[d2, e2] T ≡ T2 → d2 ≤ d1 → - ∃∃T0. ↑[d2, e2] T1 ≡ T0 & ↑[d1 + e2, e1] T0 ≡ T2. -#d1 #e1 #T1 #T #H elim H -H d1 e1 T1 T -[ #k #d1 #e1 #d2 #e2 #X #HX #_ - >(lift_inv_sort1 … HX) -HX /2/ -| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #_ - lapply (lt_to_le_to_lt … (d1+e2) Hid1 ?) // #Hie2 - elim (lift_inv_lref1 … HX) -HX * #Hid2 #HX destruct -X /4/ -| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #Hd21 - lapply (transitive_le … Hd21 Hid1) -Hd21 #Hid2 - lapply (lift_inv_lref1_ge … HX ?) -HX /2/ #HX destruct -X; - >plus_plus_comm_23 /4/ -| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21 - elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X; - elim (IHV12 … HV20 ?) -IHV12 HV20 // - elim (IHT12 … HT20 ?) -IHT12 HT20 /3 width=5/ -| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21 - elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X; - elim (IHV12 … HV20 ?) -IHV12 HV20 // - elim (IHT12 … HT20 ?) -IHT12 HT20 /3 width=5/ -] -qed. - -lemma lift_trans_ge: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T → - ∀d2,e2,T2. ↑[d2, e2] T ≡ T2 → d1 + e1 ≤ d2 → - ∃∃T0. ↑[d2 - e1, e2] T1 ≡ T0 & ↑[d1, e1] T0 ≡ T2. -#d1 #e1 #T1 #T #H elim H -H d1 e1 T1 T -[ #k #d1 #e1 #d2 #e2 #X #HX #_ - >(lift_inv_sort1 … HX) -HX /2/ -| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #Hded - lapply (lt_to_le_to_lt … (d1+e1) Hid1 ?) // #Hid1e - lapply (lt_to_le_to_lt … (d2-e1) Hid1 ?) /2/ #Hid2e - lapply (lt_to_le_to_lt … Hid1e Hded) -Hid1e Hded #Hid2 - lapply (lift_inv_lref1_lt … HX ?) -HX // #HX destruct -X /3/ -| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #_ - elim (lift_inv_lref1 … HX) -HX * #Hied #HX destruct -X; - [2: >plus_plus_comm_23] /4/ -| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded - elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X; - elim (IHV12 … HV20 ?) -IHV12 HV20 // - elim (IHT12 … HT20 ?) -IHT12 HT20 /2/ #T - arith_e2 // #V0 #HV10 #HV02 + elim (lift_split … HV12 d (i - e + 1) ? ? ?) -HV12; [2,3,4: normalize /2/ ] -Hdei >arith_e2 // #V0 #HV10 #HV02 @ex2_1_intro [2: @pts_subst [4: /2/ |6,7,8: // |1,2,3: skip |5: @arith5 // ] |1: skip diff --git a/matita/matita/lib/lambda-delta/syntax/length.ma b/matita/matita/lib/lambda-delta/syntax/length.ma new file mode 100644 index 000000000..91e1bd78c --- /dev/null +++ b/matita/matita/lib/lambda-delta/syntax/length.ma @@ -0,0 +1,22 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "lambda-delta/syntax/lenv.ma". + +(* LENGTH *******************************************************************) + +(* the length of a local environment *) +let rec length L ≝ match L with +[ LSort ⇒ 0 +| LPair L _ _ ⇒ length L + 1 +]. + +interpretation "length (local environment)" 'card L = (length L). diff --git a/matita/matita/lib/lambda-delta/xoa.conf.xml b/matita/matita/lib/lambda-delta/xoa.conf.xml index 6eebcc64e..3330813ae 100644 --- a/matita/matita/lib/lambda-delta/xoa.conf.xml +++ b/matita/matita/lib/lambda-delta/xoa.conf.xml @@ -11,7 +11,7 @@
lib/lambda-delta - xoa_defs + xoa xoa_notation basics/pts.ma 2 1 diff --git a/matita/matita/lib/lambda-delta/xoa.ma b/matita/matita/lib/lambda-delta/xoa.ma new file mode 100644 index 000000000..582b2d25f --- /dev/null +++ b/matita/matita/lib/lambda-delta/xoa.ma @@ -0,0 +1,135 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* This file was generated by xoa.native: do not edit *********************) + +include "basics/pts.ma". + +(* multiple existental quantifier (2, 1) *) + +inductive ex2_1 (A0:Type[0]) (P0,P1:A0→Prop) : Prop ≝ + | ex2_1_intro: ∀x0. P0 x0 → P1 x0 → ex2_1 ? ? ? +. + +interpretation "multiple existental quantifier (2, 1)" 'Ex P0 P1 = (ex2_1 ? P0 P1). + +(* multiple existental quantifier (2, 2) *) + +inductive ex2_2 (A0,A1:Type[0]) (P0,P1:A0→A1→Prop) : Prop ≝ + | ex2_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → ex2_2 ? ? ? ? +. + +interpretation "multiple existental quantifier (2, 2)" 'Ex P0 P1 = (ex2_2 ? ? P0 P1). + +(* multiple existental quantifier (3, 1) *) + +inductive ex3_1 (A0:Type[0]) (P0,P1,P2:A0→Prop) : Prop ≝ + | ex3_1_intro: ∀x0. P0 x0 → P1 x0 → P2 x0 → ex3_1 ? ? ? ? +. + +interpretation "multiple existental quantifier (3, 1)" 'Ex P0 P1 P2 = (ex3_1 ? P0 P1 P2). + +(* multiple existental quantifier (3, 2) *) + +inductive ex3_2 (A0,A1:Type[0]) (P0,P1,P2:A0→A1→Prop) : Prop ≝ + | ex3_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → P2 x0 x1 → ex3_2 ? ? ? ? ? +. + +interpretation "multiple existental quantifier (3, 2)" 'Ex P0 P1 P2 = (ex3_2 ? ? P0 P1 P2). + +(* multiple existental quantifier (3, 3) *) + +inductive ex3_3 (A0,A1,A2:Type[0]) (P0,P1,P2:A0→A1→A2→Prop) : Prop ≝ + | ex3_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → ex3_3 ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (3, 3)" 'Ex P0 P1 P2 = (ex3_3 ? ? ? P0 P1 P2). + +(* multiple existental quantifier (4, 3) *) + +inductive ex4_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3:A0→A1→A2→Prop) : Prop ≝ + | ex4_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → P3 x0 x1 x2 → ex4_3 ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (4, 3)" 'Ex P0 P1 P2 P3 = (ex4_3 ? ? ? P0 P1 P2 P3). + +(* multiple existental quantifier (4, 4) *) + +inductive ex4_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3:A0→A1→A2→A3→Prop) : Prop ≝ + | ex4_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → ex4_4 ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (4, 4)" 'Ex P0 P1 P2 P3 = (ex4_4 ? ? ? ? P0 P1 P2 P3). + +(* multiple existental quantifier (5, 3) *) + +inductive ex5_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3,P4:A0→A1→A2→Prop) : Prop ≝ + | ex5_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → P3 x0 x1 x2 → P4 x0 x1 x2 → ex5_3 ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (5, 3)" 'Ex P0 P1 P2 P3 P4 = (ex5_3 ? ? ? P0 P1 P2 P3 P4). + +(* multiple existental quantifier (5, 4) *) + +inductive ex5_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3,P4:A0→A1→A2→A3→Prop) : Prop ≝ + | ex5_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → P4 x0 x1 x2 x3 → ex5_4 ? ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (5, 4)" 'Ex P0 P1 P2 P3 P4 = (ex5_4 ? ? ? ? P0 P1 P2 P3 P4). + +(* multiple existental quantifier (6, 6) *) + +inductive ex6_6 (A0,A1,A2,A3,A4,A5:Type[0]) (P0,P1,P2,P3,P4,P5:A0→A1→A2→A3→A4→A5→Prop) : Prop ≝ + | ex6_6_intro: ∀x0,x1,x2,x3,x4,x5. P0 x0 x1 x2 x3 x4 x5 → P1 x0 x1 x2 x3 x4 x5 → P2 x0 x1 x2 x3 x4 x5 → P3 x0 x1 x2 x3 x4 x5 → P4 x0 x1 x2 x3 x4 x5 → P5 x0 x1 x2 x3 x4 x5 → ex6_6 ? ? ? ? ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (6, 6)" 'Ex P0 P1 P2 P3 P4 P5 = (ex6_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5). + +(* multiple existental quantifier (7, 6) *) + +inductive ex7_6 (A0,A1,A2,A3,A4,A5:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→A3→A4→A5→Prop) : Prop ≝ + | ex7_6_intro: ∀x0,x1,x2,x3,x4,x5. P0 x0 x1 x2 x3 x4 x5 → P1 x0 x1 x2 x3 x4 x5 → P2 x0 x1 x2 x3 x4 x5 → P3 x0 x1 x2 x3 x4 x5 → P4 x0 x1 x2 x3 x4 x5 → P5 x0 x1 x2 x3 x4 x5 → P6 x0 x1 x2 x3 x4 x5 → ex7_6 ? ? ? ? ? ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (7, 6)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6). + +(* multiple disjunction connective (3) *) + +inductive or3 (P0,P1,P2:Prop) : Prop ≝ + | or3_intro0: P0 → or3 ? ? ? + | or3_intro1: P1 → or3 ? ? ? + | or3_intro2: P2 → or3 ? ? ? +. + +interpretation "multiple disjunction connective (3)" 'Or P0 P1 P2 = (or3 P0 P1 P2). + +(* multiple disjunction connective (4) *) + +inductive or4 (P0,P1,P2,P3:Prop) : Prop ≝ + | or4_intro0: P0 → or4 ? ? ? ? + | or4_intro1: P1 → or4 ? ? ? ? + | or4_intro2: P2 → or4 ? ? ? ? + | or4_intro3: P3 → or4 ? ? ? ? +. + +interpretation "multiple disjunction connective (4)" 'Or P0 P1 P2 P3 = (or4 P0 P1 P2 P3). + +(* multiple conjunction connective (3) *) + +inductive and3 (P0,P1,P2:Prop) : Prop ≝ + | and3_intro: P0 → P1 → P2 → and3 ? ? ? +. + +interpretation "multiple conjunction connective (3)" 'And P0 P1 P2 = (and3 P0 P1 P2). + diff --git a/matita/matita/lib/lambda-delta/xoa_defs.ma b/matita/matita/lib/lambda-delta/xoa_defs.ma deleted file mode 100644 index 582b2d25f..000000000 --- a/matita/matita/lib/lambda-delta/xoa_defs.ma +++ /dev/null @@ -1,135 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -(* This file was generated by xoa.native: do not edit *********************) - -include "basics/pts.ma". - -(* multiple existental quantifier (2, 1) *) - -inductive ex2_1 (A0:Type[0]) (P0,P1:A0→Prop) : Prop ≝ - | ex2_1_intro: ∀x0. P0 x0 → P1 x0 → ex2_1 ? ? ? -. - -interpretation "multiple existental quantifier (2, 1)" 'Ex P0 P1 = (ex2_1 ? P0 P1). - -(* multiple existental quantifier (2, 2) *) - -inductive ex2_2 (A0,A1:Type[0]) (P0,P1:A0→A1→Prop) : Prop ≝ - | ex2_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → ex2_2 ? ? ? ? -. - -interpretation "multiple existental quantifier (2, 2)" 'Ex P0 P1 = (ex2_2 ? ? P0 P1). - -(* multiple existental quantifier (3, 1) *) - -inductive ex3_1 (A0:Type[0]) (P0,P1,P2:A0→Prop) : Prop ≝ - | ex3_1_intro: ∀x0. P0 x0 → P1 x0 → P2 x0 → ex3_1 ? ? ? ? -. - -interpretation "multiple existental quantifier (3, 1)" 'Ex P0 P1 P2 = (ex3_1 ? P0 P1 P2). - -(* multiple existental quantifier (3, 2) *) - -inductive ex3_2 (A0,A1:Type[0]) (P0,P1,P2:A0→A1→Prop) : Prop ≝ - | ex3_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → P2 x0 x1 → ex3_2 ? ? ? ? ? -. - -interpretation "multiple existental quantifier (3, 2)" 'Ex P0 P1 P2 = (ex3_2 ? ? P0 P1 P2). - -(* multiple existental quantifier (3, 3) *) - -inductive ex3_3 (A0,A1,A2:Type[0]) (P0,P1,P2:A0→A1→A2→Prop) : Prop ≝ - | ex3_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → ex3_3 ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (3, 3)" 'Ex P0 P1 P2 = (ex3_3 ? ? ? P0 P1 P2). - -(* multiple existental quantifier (4, 3) *) - -inductive ex4_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3:A0→A1→A2→Prop) : Prop ≝ - | ex4_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → P3 x0 x1 x2 → ex4_3 ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (4, 3)" 'Ex P0 P1 P2 P3 = (ex4_3 ? ? ? P0 P1 P2 P3). - -(* multiple existental quantifier (4, 4) *) - -inductive ex4_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3:A0→A1→A2→A3→Prop) : Prop ≝ - | ex4_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → ex4_4 ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (4, 4)" 'Ex P0 P1 P2 P3 = (ex4_4 ? ? ? ? P0 P1 P2 P3). - -(* multiple existental quantifier (5, 3) *) - -inductive ex5_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3,P4:A0→A1→A2→Prop) : Prop ≝ - | ex5_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → P3 x0 x1 x2 → P4 x0 x1 x2 → ex5_3 ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (5, 3)" 'Ex P0 P1 P2 P3 P4 = (ex5_3 ? ? ? P0 P1 P2 P3 P4). - -(* multiple existental quantifier (5, 4) *) - -inductive ex5_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3,P4:A0→A1→A2→A3→Prop) : Prop ≝ - | ex5_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → P4 x0 x1 x2 x3 → ex5_4 ? ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (5, 4)" 'Ex P0 P1 P2 P3 P4 = (ex5_4 ? ? ? ? P0 P1 P2 P3 P4). - -(* multiple existental quantifier (6, 6) *) - -inductive ex6_6 (A0,A1,A2,A3,A4,A5:Type[0]) (P0,P1,P2,P3,P4,P5:A0→A1→A2→A3→A4→A5→Prop) : Prop ≝ - | ex6_6_intro: ∀x0,x1,x2,x3,x4,x5. P0 x0 x1 x2 x3 x4 x5 → P1 x0 x1 x2 x3 x4 x5 → P2 x0 x1 x2 x3 x4 x5 → P3 x0 x1 x2 x3 x4 x5 → P4 x0 x1 x2 x3 x4 x5 → P5 x0 x1 x2 x3 x4 x5 → ex6_6 ? ? ? ? ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (6, 6)" 'Ex P0 P1 P2 P3 P4 P5 = (ex6_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5). - -(* multiple existental quantifier (7, 6) *) - -inductive ex7_6 (A0,A1,A2,A3,A4,A5:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→A3→A4→A5→Prop) : Prop ≝ - | ex7_6_intro: ∀x0,x1,x2,x3,x4,x5. P0 x0 x1 x2 x3 x4 x5 → P1 x0 x1 x2 x3 x4 x5 → P2 x0 x1 x2 x3 x4 x5 → P3 x0 x1 x2 x3 x4 x5 → P4 x0 x1 x2 x3 x4 x5 → P5 x0 x1 x2 x3 x4 x5 → P6 x0 x1 x2 x3 x4 x5 → ex7_6 ? ? ? ? ? ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (7, 6)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6). - -(* multiple disjunction connective (3) *) - -inductive or3 (P0,P1,P2:Prop) : Prop ≝ - | or3_intro0: P0 → or3 ? ? ? - | or3_intro1: P1 → or3 ? ? ? - | or3_intro2: P2 → or3 ? ? ? -. - -interpretation "multiple disjunction connective (3)" 'Or P0 P1 P2 = (or3 P0 P1 P2). - -(* multiple disjunction connective (4) *) - -inductive or4 (P0,P1,P2,P3:Prop) : Prop ≝ - | or4_intro0: P0 → or4 ? ? ? ? - | or4_intro1: P1 → or4 ? ? ? ? - | or4_intro2: P2 → or4 ? ? ? ? - | or4_intro3: P3 → or4 ? ? ? ? -. - -interpretation "multiple disjunction connective (4)" 'Or P0 P1 P2 P3 = (or4 P0 P1 P2 P3). - -(* multiple conjunction connective (3) *) - -inductive and3 (P0,P1,P2:Prop) : Prop ≝ - | and3_intro: P0 → P1 → P2 → and3 ? ? ? -. - -interpretation "multiple conjunction connective (3)" 'And P0 P1 P2 = (and3 P0 P1 P2). - diff --git a/matita/matita/lib/lambda-delta/xoa_props.ma b/matita/matita/lib/lambda-delta/xoa_props.ma index 967656716..f07bc89be 100644 --- a/matita/matita/lib/lambda-delta/xoa_props.ma +++ b/matita/matita/lib/lambda-delta/xoa_props.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "lambda-delta/xoa_notation.ma". -include "lambda-delta/xoa_defs.ma". +include "lambda-delta/xoa.ma". lemma ex2_1_comm: ∀A0. ∀P0,P1:A0→Prop. (∃∃x0. P0 x0 & P1 x0) → ∃∃x0. P1 x0 & P0 x0. #A0 #P0 #P1 * /2/