From 37905ef451d98f0c857d62876fdbe12f0ee8ccaf Mon Sep 17 00:00:00 2001 From: Andrea Asperti Date: Fri, 7 Mar 2014 13:32:16 +0000 Subject: [PATCH] Minor changes to make the script more robust to strategy changes. --- matita/matita/lib/arithmetics/binomial.ma | 2 +- .../lib/arithmetics/chebyshev/bertrand.ma | 2 +- matita/matita/lib/basics/finset.ma | 14 ++++------- matita/matita/lib/hints_declaration.ma | 23 ++++++++++++++++--- 4 files changed, 27 insertions(+), 14 deletions(-) diff --git a/matita/matita/lib/arithmetics/binomial.ma b/matita/matita/lib/arithmetics/binomial.ma index c3834c8d8..6b4d3ff52 100644 --- a/matita/matita/lib/arithmetics/binomial.ma +++ b/matita/matita/lib/arithmetics/binomial.ma @@ -103,7 +103,7 @@ theorem binomial_law:∀a,b,n. >bigop_Strue in ⊢ (??(??%)?); // associative_plus @eq_f2 [bc_n_n >bc_n_n normalize // - |>bigop_0 >associative_plus >commutative_plus in ⊢ (??(??%)?); + |>(bigop_0 ??? plusA) >associative_plus >commutative_plus in ⊢ (??(??%)?); (bigop_0 n ?? plusA) @eq_f2 [cut (∀a,b. plus a b = plusAC a b) [//] #Hplus >Hplus >(bigop_op n ??? plusAC) @same_bigop // diff --git a/matita/matita/lib/arithmetics/chebyshev/bertrand.ma b/matita/matita/lib/arithmetics/chebyshev/bertrand.ma index 99770436b..8f41262e3 100644 --- a/matita/matita/lib/arithmetics/chebyshev/bertrand.ma +++ b/matita/matita/lib/arithmetics/chebyshev/bertrand.ma @@ -166,7 +166,7 @@ elim (log p (2*n)) [@leb_true_to_le // |>commutative_times in ⊢ (??%); > times_exp @(transitive_le ? (exp n 2)) - [exp_2 in ⊢ (??%); @le_times // + [exp_2 in ⊢ (??%); @le_times [@le18|//] |@(le_exp1 … (lt_O_S ?)) @(le_plus_to_le 3) cut (3+2*n/3*3 = S(2*n/3)*3) [//] #eq1 >eq1 diff --git a/matita/matita/lib/basics/finset.ma b/matita/matita/lib/basics/finset.ma index 33b6b664a..df8d0a87c 100644 --- a/matita/matita/lib/basics/finset.ma +++ b/matita/matita/lib/basics/finset.ma @@ -64,16 +64,12 @@ lemma memb_enumn: ∀m,n,i:DeqNat. ∀h:n ≤ m. ∀k: i < m. n ≤ i → (¬ (memb (Nat_to m) (mk_Sig ?? i k) (enumnaux n m h))) = true. #m #n elim n -n // #n #Hind #i #ltm #k #ltni @sym_eq @noteq_to_eqnot @sym_not_eq % #H cases (orb_true_l … H) - [#H1 @(absurd … (\P H1)) % #Hfalse - cut (∀A,P,a,a1,h,h1.mk_Sig A P a h = mk_Sig A P a1 h1 → a = a1) - [#A #P #a #a1 #h #h1 #H destruct (H) %] #Hcut - lapply (Hcut nat (λi.iHfalse @le_n - |<(notb_notb (memb …)) >Hind normalize /2/ + [whd in ⊢ (??%?→?); #H1 @(absurd … ltni) @le_to_not_lt + >(eqb_true_to_eq … H1) @le_n + |<(notb_notb (memb …)) >Hind normalize /2 by lt_to_le, absurd/ ] qed. - lemma enumn_unique_aux: ∀n,m. ∀h:n ≤ m. uniqueb (Nat_to m) (enumnaux n m h) = true. #n elim n -n // #n #Hind #m #h @true_to_andb_true // @memb_enumn // qed. @@ -209,13 +205,13 @@ unification hint 0 ≔ C1,C2; definition enum_prod ≝ λA,B:DeqSet.λl1.λl2. compose ??? (mk_Prod A B) l1 l2. - + lemma enum_prod_unique: ∀A,B,l1,l2. uniqueb A l1 = true → uniqueb B l2 = true → uniqueb ? (enum_prod A B l1 l2) = true. #A #B #l1 elim l1 // #a #tl #Hind #l2 #H1 #H2 @uniqueb_append - [@unique_map_inj // + [@unique_map_inj [#x #y #Heq @(eq_f … \snd … Heq) | //] |@Hind // @(andb_true_r … H1) |#p #H3 cases (memb_map_to_exists … H3) #b * #Hmemb #eqp "≔ (list0 ( (list1 (ident x) sep , ) opt (: T) ) sep ,) opt (; (lis } }}. +(* it seems unbelivable, but it works! *) +notation > "≔ (list0 ( (list1 (ident x) sep , ) opt (: T) ) sep ,) opt (; (list1 (ident U ≟ term 19 V ) sep ,)) ⊢ term 19 Px ≡≡ term 19 Py" + with precedence 90 + for @{ ${ fold right + @{ ${ default + @{ ${ fold right + @{ 'hint_decl2 $Px $Py } + rec acc1 @{ let ( ${ident U} : ?) ≝ $V in $acc1} } } + @{ 'hint_decl2 $Px $Py } + } + } + rec acc @{ + ${ fold right @{ $acc } rec acc2 + @{ ∀${ident x}:${ default @{ $T } @{ ? } }.$acc2 } } + } + }}. + include "basics/pts.ma". definition hint_declaration_Type0 ≝ λA:Type[0] .λa,b:A.Prop. @@ -65,8 +82,8 @@ definition hint_declaration_CProp0 ≝ λA:CProp[0].λa,b:A.Prop. definition hint_declaration_CProp1 ≝ λA:CProp[1].λa,b:A.Prop. definition hint_declaration_CProp2 ≝ λa,b:CProp[2].Prop. -interpretation "hint_decl_Type2" 'hint_decl a b = (hint_declaration_Type2 a b). -interpretation "hint_decl_CProp2" 'hint_decl a b = (hint_declaration_CProp2 a b). +interpretation "hint_decl_Type2" 'hint_decl2 a b = (hint_declaration_Type2 a b). +interpretation "hint_decl_CProp2" 'hint_decl2 a b = (hint_declaration_CProp2 a b). interpretation "hint_decl_Type1" 'hint_decl a b = (hint_declaration_Type1 ? a b). interpretation "hint_decl_CProp1" 'hint_decl a b = (hint_declaration_CProp1 ? a b). interpretation "hint_decl_CProp0" 'hint_decl a b = (hint_declaration_CProp0 ? a b). @@ -104,4 +121,4 @@ unification hint 0 ≔ R : setoid; (* ---------------------------------------- *) ⊢ setoid ≡ target1 ? MR sol. -*) \ No newline at end of file +*) -- 2.39.2