X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fbigops.ma;h=97bda394da6d144854b122000456d94146ce0a22;hb=5535cd4e08fd8d1e7e6e067eac1bb6c1bf8fcbbf;hp=3211814db099d89326c46bbc9dfb6d5cde551f21;hpb=53452958508001e7af3090695b619fe92135fb9e;p=helm.git diff --git a/matita/matita/lib/arithmetics/bigops.ma b/matita/matita/lib/arithmetics/bigops.ma index 3211814db..97bda394d 100644 --- a/matita/matita/lib/arithmetics/bigops.ma +++ b/matita/matita/lib/arithmetics/bigops.ma @@ -67,10 +67,10 @@ for @{'bigop ($b-$a) $op $nil (λ${ident j}.((λ${ident j}.$p) (${ident j}+$a))) notation "\big [ op , nil ]_{ ident j ∈ [a,b[ } f" with precedence 80 -for @{'bigop ($b-$a) $op $nil (λ${ident j}.((λ${ident j}.$true) (${ident j}+$a))) +for @{'bigop ($b-$a) $op $nil (λ${ident j}.((λ${ident j}.true) (${ident j}+$a))) (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}. -(* notation "\big [ op , nil ]_{( term 50) a ≤ ident j < b | p } f" +(* notation "\big [ op , nil ]_{( term 55) a ≤ ident j < b | p } f" with precedence 80 for @{\big[$op,$nil]_{${ident j} < ($b-$a) | ((λ${ident j}.$p) (${ident j}+$a))}((λ${ident j}.$f)(${ident j}+$a))}. *) @@ -99,7 +99,7 @@ qed. theorem pad_bigop: ∀k,n,p,B,nil,op.∀f:nat→B. n ≤ k → \big[op,nil]_{i < n | p i}(f i) - = \big[op,nil]_{i < k | if_then_else ? (leb n i) false (p i)}(f i). + = \big[op,nil]_{i < k | if leb n i then false else p i}(f i). #k #n #p #B #nil #op #f #lenk (elim lenk) [@same_bigop #i #lti // >(not_le_to_leb_false …) /2/ |#j #leup #Hind >bigop_Sfalse >(le_to_leb_true … leup) // @@ -114,8 +114,8 @@ record Aop (A:Type[0]) (nil:A) : Type[0] ≝ theorem bigop_sum: ∀k1,k2,p1,p2,B.∀nil.∀op:Aop B nil.∀f,g:nat→B. op (\big[op,nil]_{inill @same_bigop #i #lti >(lt_to_leb_false … lti) normalize /2/ @@ -141,14 +141,30 @@ a ≤ b → b ≤ c → op (\big[op,nil]_{i ∈ [b,c[ |p i}(f i)) \big[op,nil]_{i ∈ [a,b[ |p i}(f i). #a #b # c #p #B #nil #op #f #leab #lebc ->(plus_minus_m_m (c-a) (b-a)) in ⊢ (??%?) /2/ ->minus_minus >(commutative_plus a) (plus_minus_m_m (c-a) (b-a)) in ⊢ (??%?); /2/ +>minus_plus >(commutative_plus a) bigop_sum (cut (∀i. b -a ≤ i → i+a = i-(b-a)+b)) [#i #lei >plus_minus // (bigop_sumI a (S a) (S b)) [|@le_S_S //|//] @eq_f2 + [@same_bigop // |bigop_a [|//] @eq_f2 [|//] bigop_Strue // >Hind >bigop_sum @same_bigop #i #lti @leb_elim // #lei cut (i = n*k2+(i-n*k2)) /2/ - #eqi [|#H] (>eqi in ⊢ (???%)) + #eqi [|#H] >eqi in ⊢ (???%); >div_plus_times /2/ >Hp1 >(mod_plus_times …) /2/ |>bigop_Sfalse // >Hind >(pad_bigop (S n*k2)) // @same_bigop #i #lti @leb_elim // #lei cut (i = n*k2+(i-n*k2)) /2/ - #eqi >eqi in ⊢ (???%) >div_plus_times /2/ + #eqi >eqi in ⊢ (???%); >div_plus_times /2/ ] qed. @@ -169,7 +185,19 @@ record ACop (A:Type[0]) (nil:A) : Type[0] ≝ {aop :> Aop A nil; comm: ∀a,b.aop a b = aop b a }. - + +lemma bigop_op: ∀k,p,B.∀nil.∀op:ACop B nil.∀f,g: nat → B. +op (\big[op,nil]_{ibigop_Strue // >bigop_Strue // >bigop_Strue // + normalize assoc >comm in ⊢ (??(????%?)?); + bigop_Sfalse // >bigop_Sfalse // >bigop_Sfalse // + ] +qed. + lemma bigop_diff: ∀p,B.∀nil.∀op:ACop B nil.∀f:nat → B.∀i,n. i < n → p i = true → \big[op,nil]_{x(not_eq_to_eqb_false … (lt_to_not_eq … Hi)) //] #Hcut cases (true_or_false (p n)) #pn [>bigop_Strue // >bigop_Strue // - >assoc >(comm ?? op (f i) (f n)) Hind // + normalize >assoc >(comm ?? op (f i) (f n)) Hind // |>bigop_Sfalse // >bigop_Sfalse // >Hind // ] |bigop_Strue // @eq_f >bigop_Sfalse @@ -235,7 +263,7 @@ theorem bigop_iso: ∀n1,n2,p1,p2,B.∀nil.∀op:ACop B nil.∀f1,f2. iso B (mk_range B f1 n1 p1) (mk_range B f2 n2 p2) → \big[op,nil]_{ibigop_Sfalse [@(Hind ? (le_O_n ?)) [/2/ | @(transitive_sub … (sub_lt …) sub2) //] @@ -252,7 +280,7 @@ theorem bigop_iso: ∀n1,n2,p1,p2,B.∀nil.∀op:ACop B nil.∀f1,f2. |#j #ltj #p2j (cases (sub2 j ltj (andb_true_r …p2j))) * #ltkj #p1kj #eqj % // % // (cases (le_to_or_lt_eq …(le_S_S_to_le …ltkj))) // - #eqkj @False_ind generalize in match p2j @eqb_elim + #eqkj @False_ind lapply p2j @eqb_elim normalize /2/ ] |>bigop_Sfalse // @(Hind ? len) @@ -265,7 +293,30 @@ theorem bigop_iso: ∀n1,n2,p1,p2,B.∀nil.∀op:ACop B nil.∀f1,f2. ] qed. -(* Sigma e Pi *) +(* distributivity *) + +record Dop (A:Type[0]) (nil:A): Type[0] ≝ + {sum : ACop A nil; + prod: A → A →A; + null: \forall a. prod a nil = nil; + distr: ∀a,b,c:A. prod a (sum b c) = sum (prod a b) (prod a c) + }. + +theorem bigop_distr: ∀n,p,B,nil.∀R:Dop B nil.∀f,a. + let aop ≝ sum B nil R in + let mop ≝ prod B nil R in + mop a \big[aop,nil]_{ibigop_Strue // >bigop_Strue // >(distr B nil R) >Hind // + |>bigop_Sfalse // >bigop_Sfalse // + ] +qed. + +(* Sigma e Pi + + notation "Σ_{ ident i < n | p } f" with precedence 80 for @{'bigop $n plus 0 (λ${ident i}.p) (λ${ident i}. $f)}. @@ -301,7 +352,8 @@ notation "Π_{ ident j ∈ [a,b[ | p } f" with precedence 80 for @{'bigop ($b-$a) times 1 (λ${ident j}.((λ${ident j}.$p) (${ident j}+$a))) (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}. - + +*) (* definition p_ord_times \def @@ -720,44 +772,7 @@ cut (O < p) ] qed. -(*distributive property for iter_p_gen*) -theorem distributive_times_plus_iter_p_gen: \forall A:Type. -\forall plusA: A \to A \to A. \forall basePlusA: A. -\forall timesA: A \to A \to A. -\forall n:nat. \forall k:A. \forall p:nat \to bool. \forall g:nat \to A. -(symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a basePlusA) = a) \to -(symmetric A timesA) \to (distributive A timesA plusA) \to -(\forall a:A. (timesA a basePlusA) = basePlusA) - - \to -((timesA k (iter_p_gen n p A g basePlusA plusA)) = - (iter_p_gen n p A (\lambda i:nat.(timesA k (g i))) basePlusA plusA)). -intros. -elim n -[ simplify. - apply H5 -| cut( (p n1) = true \lor (p n1) = false) - [ elim Hcut - [ rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H7). - rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H7) in \vdash (? ? ? %). - rewrite > (H4). - rewrite > (H3 k (g n1)). - apply eq_f. - assumption - | rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H7). - rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H7) in \vdash (? ? ? %). - assumption - ] - | elim ((p n1)) - [ left. - reflexivity - | right. - reflexivity - ] - ] -] -qed. theorem iter_p_gen_2_eq: \forall A:Type.