From: Ferruccio Guidi Date: Fri, 3 Jun 2011 16:52:50 +0000 (+0000) Subject: - we introduce extended existentials (generated) X-Git-Tag: make_still_working~2463 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=861d99cbe515be1a8e6ca204c2cafa40ccdec8a3;p=helm.git - we introduce extended existentials (generated) - we exploit the tactic "*" to reduce the volume of decompositions --- diff --git a/matita/matita/lib/lambda-delta/ground.ma b/matita/matita/lib/lambda-delta/ground.ma index 72a1eaf29..651c2dd99 100644 --- a/matita/matita/lib/lambda-delta/ground.ma +++ b/matita/matita/lib/lambda-delta/ground.ma @@ -10,6 +10,8 @@ V_______________________________________________________________ *) include "basics/list.ma". +include "lambda-delta/xoa_defs.ma". +include "lambda-delta/xoa_notation.ma". include "lambda-delta/notation.ma". (* ARITHMETICAL PROPERTIES **************************************************) diff --git a/matita/matita/lib/lambda-delta/substitution/lift.ma b/matita/matita/lib/lambda-delta/substitution/lift.ma index 181642504..906788968 100644 --- a/matita/matita/lib/lambda-delta/substitution/lift.ma +++ b/matita/matita/lib/lambda-delta/substitution/lift.ma @@ -60,18 +60,20 @@ qed. lemma lift_inv_con22_aux: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → ∀I,V2,U2. T2 = ♭I V2.U2 → - ∃V1,U1. ↑[d,e] V1 ≡ V2 ∧ ↑[d+1,e] U1 ≡ U2 ∧ - T1 = ♭I V1.U1. + ∃∃V1,U1. ↑[d,e] V1 ≡ V2 & ↑[d+1,e] U1 ≡ U2 & + T1 = ♭I V1.U1. #d #e #T1 #T2 #H elim H -H d e T1 T2 [ #k #d #e #I #V2 #U2 #H destruct | #i #d #e #_ #I #V2 #U2 #H destruct | #i #d #e #_ #I #V2 #U2 #H destruct - | #J #W1 #W2 #T1 #T2 #d #e #HW #HT #_ #_ #I #V2 #U2 #H destruct /5/ + | #J #W1 #W2 #T1 #T2 #d #e #HW #HT #_ #_ #I #V2 #U2 #H destruct + /2 width = 5/ + ] qed. lemma lift_inv_con22: ∀d,e,T1,I,V2,U2. ↑[d,e] T1 ≡ ♭I V2. U2 → - ∃V1,U1. ↑[d,e] V1 ≡ V2 ∧ ↑[d+1,e] U1 ≡ U2 ∧ - T1 = ♭I V1. U1. + ∃∃V1,U1. ↑[d,e] V1 ≡ V2 & ↑[d+1,e] U1 ≡ U2 & + T1 = ♭I V1. U1. #d #e #T1 #I #V2 #U2 #H lapply (lift_inv_con22_aux … H) /2/ qed. @@ -80,57 +82,44 @@ qed. theorem lift_trans_rev: ∀d1,e1,T1,T. ↑[d1,e1] T1 ≡ T → ∀d2,e2,T2. ↑[d2 + e1, e2] T2 ≡ T → d1 ≤ d2 → - ∃T0. ↑[d1, e1] T0 ≡ T2 ∧ ↑[d2, e2] T0 ≡ T1. + ∃∃T0. ↑[d1, e1] T0 ≡ T2 & ↑[d2, e2] T0 ≡ T1. #d1 #e1 #T1 #T #H elim H -H d1 e1 T1 T [ #k #d1 #e1 #d2 #e2 #T2 #Hk #Hd12 lapply (lift_inv_sort2 … Hk) -Hk #Hk destruct -T2 /3/ | #i #d1 #e1 #Hid1 #d2 #e2 #T2 #Hi #Hd12 - lapply (lift_inv_lref2 … Hi) -Hi #H - elim H -H #H elim H -H #Hid2 #Hi (**) (* decompose*) - destruct -T2 - [ -Hid2 /5/ (**) (* /4/ *) + lapply (lift_inv_lref2 … Hi) -Hi * * #Hid2 #H destruct -T2 + [ -Hid2 /4/ | elim (lt_false d1 ?) @(le_to_lt_to_lt … Hd12) -Hd12 @(le_to_lt_to_lt … Hid1) -Hid1 /2/ ] | #i #d1 #e1 #Hid1 #d2 #e2 #T2 #Hi #Hd12 - lapply (lift_inv_lref2 … Hi) -Hi #H - elim H -H #H elim H -H #Hid2 #Hi (**) (* decompose*) - destruct -T2 - [ -Hd12; lapply (lt_plus_to_lt_l … Hid2) -Hid2 #Hid2 /4/ + lapply (lift_inv_lref2 … Hi) -Hi * * #Hid2 #H destruct -T2 + [ -Hd12; lapply (lt_plus_to_lt_l … Hid2) -Hid2 #Hid2 /3/ | -Hid1; lapply (arith1 … Hid2) -Hid2 #Hid2 - @(ex_intro … #(i - e2)) @conj + @(ex2_1_intro … #(i - e2)) [ >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_con22 … H) -H #H - elim H -H #W2 #H elim H -H #U2 #H elim H -H #H elim H -H #HW2 #HU2 #H (**) (* decompose*) - destruct -T2; - elim (IHW … HW2 ?) // -IHW HW2 #W0 #H - elim H -H #HW2 #HW1 (**) (* decompose*) - >plus_plus_comm_23 in HU2 #HU2 elim (IHU … HU2 ?) /2/ -IHU HU2 Hd12 #U0 #H - elim H -H #HU2 #HU1 (**) (* decompose*) - @(ex_intro … ♭I W0. U0) /3/ (**) (* /4/ *) + lapply (lift_inv_con22 … 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/ ] qed. theorem 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. + ∃∃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 - <(plus_plus_minus_m_m e1 e2 i) // - @(ex_intro … #(i+e1)) /3/ (**) (* /4/ *) + <(plus_plus_minus_m_m e1 e2 i) /3/ | #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12 - elim (IHV … Hd12 Hd21 He12) -IHV #V0 #H - elim H -H #HV0a #HV0b (**) (* decompose *) - elim (IHT (d2+1) … ? ? He12) /2/ -IHT Hd12 Hd21 He12 #T0 #H - elim H -H #HT0a #HT0b (**) (* decompose *) - @(ex_intro … ♭I V0.T0) /3/ (**) (* /4/ *) + elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b + elim (IHT (d2+1) … ? ? He12) /3 width = 5/ ] qed. diff --git a/matita/matita/lib/lambda-delta/substitution/thin.ma b/matita/matita/lib/lambda-delta/substitution/thin.ma index 00e52947d..ab1c5f271 100644 --- a/matita/matita/lib/lambda-delta/substitution/thin.ma +++ b/matita/matita/lib/lambda-delta/substitution/thin.ma @@ -32,4 +32,4 @@ axiom thin_conf_ge: ∀d1,e1,L,L1. ↓[d1,e1] L ≡ L1 → axiom thin_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. + ∃∃K1,V1. ↓[0,e2] L1 ≡ K1. ♭I V1 & ↓[d,e1] K2 ≡ K1 & ↑[d,e1] V1 ≡ V2. diff --git a/matita/matita/lib/lambda-delta/xoa_defs.ma b/matita/matita/lib/lambda-delta/xoa_defs.ma new file mode 100644 index 000000000..e36373e3c --- /dev/null +++ b/matita/matita/lib/lambda-delta/xoa_defs.ma @@ -0,0 +1,30 @@ +(**************************************************************************) +(* ___ *) +(* ||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". + +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). + +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). + diff --git a/matita/matita/lib/lambda-delta/xoa_notation.ma b/matita/matita/lib/lambda-delta/xoa_notation.ma new file mode 100644 index 000000000..1af6c7fab --- /dev/null +++ b/matita/matita/lib/lambda-delta/xoa_notation.ma @@ -0,0 +1,24 @@ +(**************************************************************************) +(* ___ *) +(* ||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 *********************) + +notation "hvbox(∃∃ ident x0 . term 19 P0 & term 19 P1)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.$P0) (λ${ident x0}.$P1) }. + +notation "hvbox(∃∃ ident x0 , ident x1 . term 19 P0 & term 19 P1 & term 19 P2)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0},${ident x1}.$P0) (λ${ident x0},${ident x1}.$P1) (λ${ident x0},${ident x1}.$P2) }. +