X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fbigops.ma;h=03db53f7d5ff0f9863abaeba9cfe10f5bd681a33;hb=7ad16d18416a08382d62747fce4a0ac18ee557e0;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..03db53f7d 100644 --- a/matita/matita/lib/arithmetics/bigops.ma +++ b/matita/matita/lib/arithmetics/bigops.ma @@ -67,7 +67,7 @@ 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" @@ -142,13 +142,29 @@ a ≤ b → b ≤ c → \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) 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 [|//] 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 // + 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]_{xbigop_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.