From 67c5cf7ae14c745a94defbe645c5406ccbcf514d Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Thu, 19 Jun 2014 13:33:06 +0000 Subject: [PATCH] - some pending conjectures closed in basic_2 and ground_2 - bug fix in decidability of lift --- .../lambdadelta/basic_2/grammar/item.ma | 28 ++++++- .../lambdadelta/basic_2/grammar/lenv.ma | 12 ++- .../lambdadelta/basic_2/grammar/term.ma | 13 ++- .../lambdadelta/basic_2/substitution/lift.ma | 80 +++++++++---------- .../lambdadelta/ground_2/etc/lib/arith.etc | 14 ++++ .../lambdadelta/ground_2/lib/arith.ma | 15 ++-- .../contribs/lambdadelta/ground_2/lib/bool.ma | 8 ++ 7 files changed, 115 insertions(+), 55 deletions(-) diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/item.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/item.ma index df330a7df..99260ba79 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/item.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/item.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "ground_2/lib/bool.ma". include "ground_2/lib/arith.ma". (* ITEMS ********************************************************************) @@ -43,16 +44,35 @@ inductive item2: Type[0] ≝ (* Basic properties *********************************************************) -axiom eq_item0_dec: ∀I1,I2:item0. Decidable (I1 = I2). +lemma eq_item0_dec: ∀I1,I2:item0. Decidable (I1 = I2). +* #i1 * #i2 [2,3,4,6,7,8: @or_intror #H destruct ] +elim (eq_nat_dec i1 i2) /2 width=1 by or_introl/ +#Hni12 @or_intror #H destruct /2 width=1 by/ +qed-. (* Basic_1: was: bind_dec *) -axiom eq_bind2_dec: ∀I1,I2:bind2. Decidable (I1 = I2). +lemma eq_bind2_dec: ∀I1,I2:bind2. Decidable (I1 = I2). +* * /2 width=1 by or_introl/ +@or_intror #H destruct +qed-. (* Basic_1: was: flat_dec *) -axiom eq_flat2_dec: ∀I1,I2:flat2. Decidable (I1 = I2). +lemma eq_flat2_dec: ∀I1,I2:flat2. Decidable (I1 = I2). +* * /2 width=1 by or_introl/ +@or_intror #H destruct +qed-. (* Basic_1: was: kind_dec *) -axiom eq_item2_dec: ∀I1,I2:item2. Decidable (I1 = I2). +lemma eq_item2_dec: ∀I1,I2:item2. Decidable (I1 = I2). +* [ #a1 ] #I1 * [1,3: #a2 ] #I2 +[2,3: @or_intror #H destruct +| elim (eq_bool_dec a1 a2) #Ha + [ elim (eq_bind2_dec I1 I2) /2 width=1 by or_introl/ #HI ] + @or_intror #H destruct /2 width=1 by/ +| elim (eq_flat2_dec I1 I2) /2 width=1 by or_introl/ #HI + @or_intror #H destruct /2 width=1 by/ +] +qed-. (* Basic_1: removed theorems 21: s_S s_plus s_plus_sym s_minus minus_s_s s_le s_lt s_inj s_inc diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma index c34238632..49d6cf5b8 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma @@ -40,7 +40,17 @@ interpretation "abstraction (local environment)" (* Basic properties *********************************************************) -axiom eq_lenv_dec: ∀L1,L2:lenv. Decidable (L1 = L2). +lemma eq_lenv_dec: ∀L1,L2:lenv. Decidable (L1 = L2). +#L1 elim L1 -L1 [| #L1 #I1 #V1 #IHL1 ] * [2,4: #L2 #I2 #V2 ] +[1,4: @or_intror #H destruct +| elim (eq_bind2_dec I1 I2) #HI + [ elim (eq_term_dec V1 V2) #HV + [ elim (IHL1 L2) -IHL1 /2 width=1 by or_introl/ #HL ] + ] + @or_intror #H destruct /2 width=1 by/ +| /2 width=1 by or_introl/ +] +qed-. (* Basic inversion lemmas ***************************************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/term.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/term.ma index 77725446a..fc039863e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/term.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/term.ma @@ -93,7 +93,18 @@ interpretation "native type annotation (term)" (* Basic properties *********************************************************) (* Basic_1: was: term_dec *) -axiom eq_term_dec: ∀T1,T2:term. Decidable (T1 = T2). +lemma eq_term_dec: ∀T1,T2:term. Decidable (T1 = T2). +#T1 elim T1 -T1 #I1 [| #V1 #T1 #IHV1 #IHT1 ] * #I2 [2,4: #V2 #T2 ] +[1,4: @or_intror #H destruct +| elim (eq_item2_dec I1 I2) #HI + [ elim (IHV1 V2) -IHV1 #HV + [ elim (IHT1 T2) -IHT1 /2 width=1 by or_introl/ #HT ] + ] + @or_intror #H destruct /2 width=1 by/ +| elim (eq_item0_dec I1 I2) /2 width=1 by or_introl/ #HI + @or_intror #H destruct /2 width=1 by/ +] +qed-. (* Basic inversion lemmas ***************************************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lift.ma index f90b349b3..f88486292 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/lift.ma @@ -39,7 +39,7 @@ interpretation "relocation" 'RLift d e T1 T2 = (lift d e T1 T2). (* Basic inversion lemmas ***************************************************) fact lift_inv_O2_aux: ∀d,e,T1,T2. ⇧[d, e] T1 ≡ T2 → e = 0 → T1 = T2. -#d #e #T1 #T2 #H elim H -d -e -T1 -T2 // /3 width=1/ +#d #e #T1 #T2 #H elim H -d -e -T1 -T2 /3 width=1 by eq_f2/ qed-. lemma lift_inv_O2: ∀d,T1,T2. ⇧[d, 0] T1 ≡ T2 → T1 = T2. @@ -60,8 +60,8 @@ fact 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 width=1/ -| #j #d #e #Hj #i #Hi destruct /3 width=1/ +| #j #d #e #Hj #i #Hi destruct /3 width=1 by or_introl, conj/ +| #j #d #e #Hj #i #Hi destruct /3 width=1 by or_intror, conj/ | #p #d #e #i #H destruct | #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct | #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct @@ -104,7 +104,7 @@ fact lift_inv_bind1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → | #i #d #e #_ #a #I #V1 #U1 #H destruct | #i #d #e #_ #a #I #V1 #U1 #H destruct | #p #d #e #a #I #V1 #U1 #H destruct -| #b #J #W1 #W2 #T1 #T2 #d #e #HW #HT #a #I #V1 #U1 #H destruct /2 width=5/ +| #b #J #W1 #W2 #T1 #T2 #d #e #HW #HT #a #I #V1 #U1 #H destruct /2 width=5 by ex3_2_intro/ | #J #W1 #W2 #T1 #T2 #d #e #_ #HT #a #I #V1 #U1 #H destruct ] qed-. @@ -124,7 +124,7 @@ fact lift_inv_flat1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → | #i #d #e #_ #I #V1 #U1 #H destruct | #p #d #e #I #V1 #U1 #H destruct | #a #J #W1 #W2 #T1 #T2 #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 /2 width=5 by ex3_2_intro/ ] qed-. @@ -149,8 +149,8 @@ fact 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 width=1/ -| #j #d #e #Hj #i #Hi destruct (plus_minus_m_m i e) in ⊢ (? ? ? ? %); /2 width=2/ /3 width=2/ +#d #e #i #H >(plus_minus_m_m i e) in ⊢ (? ? ? ? %); /3 width=2 by lift_lref_ge, le_plus_to_minus_r, le_plus_b/ qed. lemma lift_lref_ge_minus_eq: ∀d,e,i,j. d + e ≤ i → j = i - e → ⇧[d, e] #j ≡ #i. @@ -314,18 +314,18 @@ lemma lift_lref_ge_minus_eq: ∀d,e,i,j. d + e ≤ i → j = i - e → ⇧[d, e] (* Basic_1: was: lift_r *) lemma lift_refl: ∀T,d. ⇧[d, 0] T ≡ T. #T elim T -T -[ * #i // #d elim (lt_or_ge i d) /2 width=1/ -| * /2 width=1/ +[ * #i // #d elim (lt_or_ge i d) /2 width=1 by lift_lref_lt, lift_lref_ge/ +| * /2 width=1 by lift_bind, lift_flat/ ] qed. lemma lift_total: ∀T1,d,e. ∃T2. ⇧[d,e] T1 ≡ T2. #T1 elim T1 -T1 -[ * #i /2 width=2/ #d #e elim (lt_or_ge i d) /3 width=2/ +[ * #i /2 width=2/ #d #e elim (lt_or_ge i d) /3 width=2 by lift_lref_lt, lift_lref_ge, ex_intro/ | * [ #a ] #I #V1 #T1 #IHV1 #IHT1 #d #e elim (IHV1 d e) -IHV1 #V2 #HV12 - [ elim (IHT1 (d+1) e) -IHT1 /3 width=2/ - | elim (IHT1 d e) -IHT1 /3 width=2/ + [ elim (IHT1 (d+1) e) -IHT1 /3 width=2 by lift_bind, ex_intro/ + | elim (IHT1 d e) -IHT1 /3 width=2 by lift_flat, ex_intro/ ] ] qed. @@ -337,51 +337,49 @@ lemma lift_split: ∀d1,e2,T1,T2. ⇧[d1, e2] T1 ≡ T2 → #d1 #e2 #T1 #T2 #H elim H -d1 -e2 -T1 -T2 [ /3 width=3/ | #i #d1 #e2 #Hid1 #d2 #e1 #Hd12 #_ #_ - lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 /4 width=3/ + lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 /4 width=3 by lift_lref_lt, ex2_intro/ | #i #d1 #e2 #Hid1 #d2 #e1 #_ #Hd21 #He12 - lapply (transitive_le … (i+e1) Hd21 ?) /2 width=1/ -Hd21 #Hd21 - >(plus_minus_m_m e2 e1 ?) // /3 width=3/ + lapply (transitive_le … (i+e1) Hd21 ?) /2 width=1 by monotonic_le_plus_l/ -Hd21 #Hd21 + >(plus_minus_m_m e2 e1 ?) /3 width=3 by lift_lref_ge, ex2_intro/ | /3 width=3/ | #a #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) /2 width=1/ /3 width=5/ + elim (IHT (d2+1) … ? ? He12) /3 width=5 by lift_bind, le_S_S, ex2_intro/ | #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/ + elim (IHT d2 … ? ? He12) /3 width=5 by lift_flat, ex2_intro/ ] qed. (* Basic_1: was only: dnf_dec2 dnf_dec *) lemma is_lift_dec: ∀T2,d,e. Decidable (∃T1. ⇧[d,e] T1 ≡ T2). #T1 elim T1 -T1 -[ * [1,3: /3 width=2/ ] #i #d #e - elim (lt_dec i d) #Hid - [ /4 width=2/ - | lapply (false_lt_to_le … Hid) -Hid #Hid - elim (lt_dec i (d + e)) #Hide - [ @or_intror * #T1 #H - elim (lift_inv_lref2_be … H Hid Hide) - | lapply (false_lt_to_le … Hide) -Hide /4 width=2/ +[ * [1,3: /3 width=2 by lift_sort, lift_gref, ex_intro, or_introl/ ] #i #d #e + elim (lt_or_ge i d) #Hdi + [ /4 width=3 by lift_lref_lt, ex_intro, or_introl/ + | elim (lt_or_ge i (d + e)) #Hide + [ @or_intror * #T1 #H elim (lift_inv_lref2_be … H Hdi Hide) + | -Hdi /4 width=2 by lift_lref_ge_minus, ex_intro, or_introl/ ] ] | * [ #a ] #I #V2 #T2 #IHV2 #IHT2 #d #e [ elim (IHV2 d e) -IHV2 [ * #V1 #HV12 elim (IHT2 (d+1) e) -IHT2 - [ * #T1 #HT12 @or_introl /3 width=2/ + [ * #T1 #HT12 @or_introl /3 width=2 by lift_bind, ex_intro/ | -V1 #HT2 @or_intror * #X #H - elim (lift_inv_bind2 … H) -H /3 width=2/ + elim (lift_inv_bind2 … H) -H /3 width=2 by ex_intro/ ] | -IHT2 #HV2 @or_intror * #X #H - elim (lift_inv_bind2 … H) -H /3 width=2/ + elim (lift_inv_bind2 … H) -H /3 width=2 by ex_intro/ ] | elim (IHV2 d e) -IHV2 [ * #V1 #HV12 elim (IHT2 d e) -IHT2 [ * #T1 #HT12 /4 width=2/ | -V1 #HT2 @or_intror * #X #H - elim (lift_inv_flat2 … H) -H /3 width=2/ + elim (lift_inv_flat2 … H) -H /3 width=2 by ex_intro/ ] | -IHT2 #HV2 @or_intror * #X #H - elim (lift_inv_flat2 … H) -H /3 width=2/ + elim (lift_inv_flat2 … H) -H /3 width=2 by ex_intro/ ] ] ] diff --git a/matita/matita/contribs/lambdadelta/ground_2/etc/lib/arith.etc b/matita/matita/contribs/lambdadelta/ground_2/etc/lib/arith.etc index 656a44255..1af2859d7 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/etc/lib/arith.etc +++ b/matita/matita/contribs/lambdadelta/ground_2/etc/lib/arith.etc @@ -16,3 +16,17 @@ qed-. lemma arith1: ∀x,y,z,w. z < y → x + (y + w) - z = x + y - z + w. #x #y #z #w #H