]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/arithmetics/bigops.ma
247af0b929590f54866f61663accbf5f97fd8965
[helm.git] / matita / matita / lib / arithmetics / bigops.ma
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic        
3     ||A||  Library of Mathematics, developed at the Computer Science     
4     ||T||  Department of the University of Bologna, Italy.                     
5     ||I||                                                                 
6     ||T||  
7     ||A||  This file is distributed under the terms of the 
8     \   /  GNU General Public License Version 2        
9      \ /      
10       V_______________________________________________________________ *)
11
12 include "basics/types.ma".
13 include "arithmetics/div_and_mod.ma".
14
15 definition sameF_upto: nat → ∀A.relation(nat→A)  ≝
16 λk.λA.λf,g.∀i. i < k → f i = g i.
17      
18 definition sameF_p: nat → (nat → bool) →∀A.relation(nat→A)  ≝
19 λk,p,A,f,g.∀i. i < k → p i = true → f i = g i.
20
21 lemma sameF_upto_le: ∀A,f,g,n,m. 
22  n ≤m → sameF_upto m A f g → sameF_upto n A f g.
23 #A #f #g #n #m #lenm #samef #i #ltin @samef /2/
24 qed.
25
26 lemma sameF_p_le: ∀A,p,f,g,n,m. 
27  n ≤m → sameF_p m p A f g → sameF_p n p A f g.
28 #A #p #f #g #n #m #lenm #samef #i #ltin #pi @samef /2/
29 qed.
30
31 (*
32 definition sumF ≝ λA.λf,g:nat → A.λn,i.
33 if_then_else ? (leb n i) (g (i-n)) (f i). 
34
35 lemma sumF_unfold: ∀A,f,g,n,i. 
36 sumF A f g n i = if_then_else ? (leb n i) (g (i-n)) (f i). 
37 // qed. *)
38
39 definition prodF ≝
40  λA,B.λf:nat→A.λg:nat→B.λm,x.〈 f(div x m), g(mod x m) 〉.
41
42 (* bigop *)
43 let rec bigop (n:nat) (p:nat → bool) (B:Type[0])
44    (nil: B) (op: B → B → B)  (f: nat → B) ≝
45   match n with
46    [ O ⇒ nil
47    | S k ⇒ 
48       match p k with
49       [true ⇒ op (f k) (bigop k p B nil op f)
50       |false ⇒ bigop k p B nil op f]
51    ].
52    
53 notation "\big  [ op , nil ]_{ ident i < n | p } f"
54   with precedence 80
55 for @{'bigop $n $op $nil (λ${ident i}. $p) (λ${ident i}. $f)}.
56
57 notation "\big [ op , nil ]_{ ident i < n } f"
58   with precedence 80
59 for @{'bigop $n $op $nil (λ${ident i}.true) (λ${ident i}. $f)}.
60
61 interpretation "bigop" 'bigop n op nil p f = (bigop n p ? nil op f).
62
63 notation "\big  [ op , nil ]_{ ident j ∈ [a,b[ | p } f"
64   with precedence 80
65 for @{'bigop ($b-$a) $op $nil (λ${ident j}.((λ${ident j}.$p) (${ident j}+$a)))
66   (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
67   
68 notation "\big  [ op , nil ]_{ ident j ∈ [a,b[ } f"
69   with precedence 80
70 for @{'bigop ($b-$a) $op $nil (λ${ident j}.((λ${ident j}.true) (${ident j}+$a)))
71   (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.  
72  
73 (* notation "\big  [ op , nil ]_{( term 55) a ≤ ident j < b | p } f"
74   with precedence 80
75 for @{\big[$op,$nil]_{${ident j} < ($b-$a) | ((λ${ident j}.$p) (${ident j}+$a))}((λ${ident j}.$f)(${ident j}+$a))}.
76 *)
77  
78 interpretation "bigop" 'bigop n op nil p f = (bigop n p ? nil op f).
79    
80 lemma bigop_Strue: ∀k,p,B,nil,op.∀f:nat→B. p k = true →
81   \big[op,nil]_{i < S k | p i}(f i) =
82     op (f k) (\big[op,nil]_{i < k | p i}(f i)).
83 #k #p #B #nil #op #f #H normalize >H // qed.
84
85 lemma bigop_Sfalse: ∀k,p,B,nil,op.∀f:nat→B. p k = false →
86   \big[op,nil]_{ i < S k | p i}(f i) =
87     \big[op,nil]_{i < k | p i}(f i).
88 #k #p #B #nil #op #f #H normalize >H // qed. 
89  
90 lemma same_bigop : ∀k,p1,p2,B,nil,op.∀f,g:nat→B. 
91   sameF_upto k bool p1 p2 → sameF_p k p1 B f g →
92   \big[op,nil]_{i < k | p1 i}(f i) = 
93     \big[op,nil]_{i < k | p2 i}(g i).
94 #k #p1 #p2 #B #nil #op #f #g (elim k) // 
95 #n #Hind #samep #samef normalize >Hind /2/
96 <(samep … (le_n …)) cases(true_or_false (p1 n)) #H1 >H1 
97 normalize // <(samef … (le_n …) H1) // 
98 qed.
99
100 theorem pad_bigop: ∀k,n,p,B,nil,op.∀f:nat→B. n ≤ k → 
101 \big[op,nil]_{i < n | p i}(f i)
102   = \big[op,nil]_{i < k | if leb n i then false else p i}(f i).
103 #k #n #p #B #nil #op #f #lenk (elim lenk) 
104   [@same_bigop #i #lti // >(not_le_to_leb_false …) /2/
105   |#j #leup #Hind >bigop_Sfalse >(le_to_leb_true … leup) // 
106   ] qed.
107
108 record Aop (A:Type[0]) (nil:A) : Type[0] ≝
109   {op :2> A → A → A; 
110    nill:∀a. op nil a = a; 
111    nilr:∀a. op a nil = a;
112    assoc: ∀a,b,c.op a (op b c) = op (op a b) c
113   }.
114
115 theorem bigop_sum: ∀k1,k2,p1,p2,B.∀nil.∀op:Aop B nil.∀f,g:nat→B.
116 op (\big[op,nil]_{i<k1|p1 i}(f i)) \big[op,nil]_{i<k2|p2 i}(g i) =
117       \big[op,nil]_{i<k1+k2|if leb k2 i then p1 (i-k2) else p2 i}
118         (if leb k2 i then f (i-k2) else g i).
119 #k1 #k2 #p1 #p2 #B #nil #op #f #g (elim k1)
120   [normalize >nill @same_bigop #i #lti 
121    >(lt_to_leb_false … lti) normalize /2/
122   |#i #Hind normalize <minus_plus_m_m (cases (p1 i)) 
123    >(le_to_leb_true … (le_plus_n …)) normalize <Hind //
124    <assoc //
125   ]
126 qed.
127
128 lemma plus_minus1: ∀a,b,c. c ≤ b → a + (b -c) = a + b -c.
129 #a #b #c #lecb @sym_eq @plus_to_minus >(commutative_plus c)
130 >associative_plus <plus_minus_m_m //
131 qed.
132
133 theorem bigop_I: ∀n,p,B.∀nil.∀op:Aop B nil.∀f:nat→B.
134 \big[op,nil]_{i∈[0,n[ |p i}(f i) = \big[op,nil]_{i < n|p i}(f i). 
135 #n #p #B #nil #op #f <minus_n_O @same_bigop //
136 qed.
137           
138 theorem bigop_sumI: ∀a,b,c,p,B.∀nil.∀op:Aop B nil.∀f:nat→B.
139 a ≤ b → b ≤ c →
140 \big[op,nil]_{i∈[a,c[ |p i}(f i) = 
141   op (\big[op,nil]_{i ∈ [b,c[ |p i}(f i)) 
142       \big[op,nil]_{i ∈ [a,b[ |p i}(f i).
143 #a #b # c #p #B #nil #op #f #leab #lebc 
144 >(plus_minus_m_m (c-a) (b-a)) in ⊢ (??%?); /2/
145 >minus_plus >(commutative_plus a) <plus_minus_m_m //
146 >bigop_sum (cut (∀i. b -a ≤ i → i+a = i-(b-a)+b))
147   [#i #lei >plus_minus // <plus_minus1 
148      [@eq_f @sym_eq @plus_to_minus /2/ | /2/]] 
149 #H @same_bigop #i #ltic @leb_elim normalize // #lei <H //
150 qed.   
151
152 theorem bigop_a: ∀a,b,B.∀nil.∀op:Aop B nil.∀f:nat→B. a ≤ b →
153 \big[op,nil]_{i∈[a,S b[ }(f i) = 
154   op (\big[op,nil]_{i ∈ [a,b[ }(f (S i))) (f a).
155 #a #b #B #nil #op #f #leab 
156 >(bigop_sumI a (S a) (S b)) [|@le_S_S //|//] @eq_f2 
157   [@same_bigop // |<minus_Sn_n normalize @nilr]
158 qed.
159   
160 theorem bigop_0: ∀n,B.∀nil.∀op:Aop B nil.∀f:nat→B.
161 \big[op,nil]_{i < S n}(f i) = 
162   op (\big[op,nil]_{i < n}(f (S i))) (f 0).
163 #n #B #nil #op #f 
164 <bigop_I >bigop_a [|//] @eq_f2 [|//] <minus_n_O 
165 @same_bigop //
166 qed.    
167
168 theorem bigop_prod: ∀k1,k2,p1,p2,B.∀nil.∀op:Aop B nil.∀f: nat →nat → B.
169 \big[op,nil]_{x<k1|p1 x}(\big[op,nil]_{i<k2|p2 x i}(f x i)) =
170   \big[op,nil]_{i<k1*k2|andb (p1 (div i k2)) (p2 (div i k2) (i \mod k2))}
171      (f (div i k2) (i \mod k2)).
172 #k1 #k2 #p1 #p2 #B #nil #op #f (elim k1) //
173 #n #Hind cases(true_or_false (p1 n)) #Hp1
174   [>bigop_Strue // >Hind >bigop_sum @same_bigop
175    #i #lti @leb_elim // #lei cut (i = n*k2+(i-n*k2)) /2/
176    #eqi [|#H] >eqi in ⊢ (???%);
177      >div_plus_times /2/ >Hp1 >(mod_plus_times …) /2/
178   |>bigop_Sfalse // >Hind >(pad_bigop (S n*k2)) // @same_bigop
179    #i #lti @leb_elim // #lei cut (i = n*k2+(i-n*k2)) /2/
180    #eqi >eqi in ⊢ (???%); >div_plus_times /2/ 
181   ]
182 qed.
183
184 record ACop (A:Type[0]) (nil:A) : Type[0] ≝
185   {aop :> Aop A nil; 
186    comm: ∀a,b.aop a b = aop b a
187   }.
188  
189 lemma bigop_op: ∀k,p,B.∀nil.∀op:ACop B nil.∀f,g: nat → B.
190 op (\big[op,nil]_{i<k|p i}(f i)) (\big[op,nil]_{i<k|p i}(g i)) =
191   \big[op,nil]_{i<k|p i}(op (f i) (g i)).
192 #k #p #B #nil #op #f #g (elim k) [normalize @nill]
193 -k #k #Hind (cases (true_or_false (p k))) #H
194   [>bigop_Strue // >bigop_Strue // >bigop_Strue //
195    normalize <assoc <assoc in ⊢ (???%); @eq_f >assoc >comm in ⊢ (??(????%?)?);
196    <assoc @eq_f @Hind
197   |>bigop_Sfalse // >bigop_Sfalse // >bigop_Sfalse //
198   ]
199 qed.
200
201 lemma bigop_diff: ∀p,B.∀nil.∀op:ACop B nil.∀f:nat → B.∀i,n.
202   i < n → p i = true →
203   \big[op,nil]_{x<n|p x}(f x)=
204     op (f i) (\big[op,nil]_{x<n|andb(notb(eqb i x))(p x)}(f x)).
205 #p #B #nil #op #f #i #n (elim n) 
206   [#ltO @False_ind /2/
207   |#n #Hind #lein #pi cases (le_to_or_lt_eq … (le_S_S_to_le …lein)) #Hi
208     [cut (andb(notb(eqb i n))(p n) = (p n))
209       [>(not_eq_to_eqb_false … (lt_to_not_eq … Hi)) //] #Hcut
210      cases (true_or_false (p n)) #pn 
211       [>bigop_Strue // >bigop_Strue //
212        normalize >assoc >(comm ?? op (f i) (f n)) <assoc >Hind //
213       |>bigop_Sfalse // >bigop_Sfalse // >Hind //  
214       ]
215     |<Hi >bigop_Strue // @eq_f >bigop_Sfalse  
216        [@same_bigop // #k #ltki >not_eq_to_eqb_false /2/
217        |>eq_to_eqb_true // 
218        ]
219      ]
220    ]
221 qed.
222
223 (* range *)
224 record range (A:Type[0]): Type[0] ≝
225   {enum:nat→A; upto:nat; filter:nat→bool}.
226   
227 definition sub_hk: (nat→nat)→(nat→nat)→∀A:Type[0].relation (range A) ≝
228 λh,k,A,I,J.∀i.i<(upto A I) → (filter A I i)=true → 
229   (h i < upto A J
230   ∧ filter A J (h i) = true
231   ∧ k (h i) = i).
232
233 definition iso: ∀A:Type[0].relation (range A) ≝
234   λA,I,J.∃h,k. 
235     (∀i. i < (upto A I) → (filter A I i) = true → 
236        enum A I i = enum A J (h i)) ∧
237     sub_hk h k A I J ∧ sub_hk k h A J I.
238   
239 lemma sub_hkO: ∀h,k,A,I,J. upto A I = 0 → sub_hk h k A I J.
240 #h #k #A #I #J #up0 #i #lti >up0 @False_ind /2/
241 qed.
242
243 lemma sub0_to_false: ∀h,k,A,I,J. upto A I = 0 → sub_hk h k A J I → 
244   ∀i. i < upto A J → filter A J i = false.
245 #h #k #A #I #J #up0 #sub #i #lti cases(true_or_false (filter A J i)) //
246 #ptrue (cases (sub i lti ptrue)) * #hi @False_ind /2/ 
247 qed.
248
249 lemma sub_lt: ∀A,e,p,n,m. n ≤ m → 
250   sub_hk (λx.x) (λx.x) A (mk_range A e n p) (mk_range A e m p).
251 #A #e #f #n #m #lenm #i #lti #fi % // % /2/
252 qed. 
253
254 theorem transitive_sub: ∀h1,k1,h2,k2,A,I,J,K. 
255   sub_hk h1 k1 A I J → sub_hk h2 k2 A J K → 
256     sub_hk (λx.h2(h1 x)) (λx.k1(k2 x)) A I K.
257 #h1 #k1 #h2 #k2 #A #I #J #K #sub1 #sub2 #i #lti #fi 
258 cases(sub1 i lti fi) * #lth1i #fh1i #ei 
259 cases(sub2 (h1 i) lth1i fh1i) * #H1 #H2 #H3 % // % // 
260 qed. 
261
262 theorem bigop_iso: ∀n1,n2,p1,p2,B.∀nil.∀op:ACop B nil.∀f1,f2.
263   iso B (mk_range B f1 n1 p1) (mk_range B f2 n2 p2) →
264   \big[op,nil]_{i<n1|p1 i}(f1 i) = \big[op,nil]_{i<n2|p2 i}(f2 i).
265 #n1 #n2 #p1 #p2 #B #nil #op #f1 #f2 * #h * #k * * #same
266 @(le_gen ? n1) #i lapply p2 (elim i) 
267   [(elim n2) // #m #Hind #p2 #_ #sub1 #sub2
268    >bigop_Sfalse 
269     [@(Hind ? (le_O_n ?)) [/2/ | @(transitive_sub … (sub_lt …) sub2) //]
270     |@(sub0_to_false … sub2) //
271     ]
272   |#n #Hind #p2 #ltn #sub1 #sub2 (cut (n ≤n1)) [/2/] #len
273    cases(true_or_false (p1 n)) #p1n
274     [>bigop_Strue // (cases (sub1 n (le_n …) p1n)) * #hn #p2hn #eqn
275      >(bigop_diff … (h n) n2) // >same // 
276      @eq_f @(Hind ? len)
277       [#i #ltin #p1i (cases (sub1 i (le_S … ltin) p1i)) * 
278        #h1i #p2h1i #eqi % // % // >not_eq_to_eqb_false normalize // 
279        @(not_to_not ??? (lt_to_not_eq ? ? ltin)) // 
280       |#j #ltj #p2j (cases (sub2 j ltj (andb_true_r …p2j))) * 
281        #ltkj #p1kj #eqj % // % // 
282        (cases (le_to_or_lt_eq …(le_S_S_to_le …ltkj))) //
283        #eqkj @False_ind lapply p2j @eqb_elim 
284        normalize /2/
285       ]
286     |>bigop_Sfalse // @(Hind ? len) 
287       [@(transitive_sub … (sub_lt …) sub1) //
288       |#i #lti #p2i cases(sub2 i lti p2i) * #ltki #p1ki #eqi
289        % // % // cases(le_to_or_lt_eq …(le_S_S_to_le …ltki)) //
290        #eqki @False_ind /2/
291       ]
292     ]
293   ]
294 qed.
295
296 (* distributivity *)
297
298 record Dop (A:Type[0]) (nil:A): Type[0] ≝
299   {sum : ACop A nil; 
300    prod: A → A →A;
301    null: \forall a. prod a nil = nil; 
302    distr: ∀a,b,c:A. prod a (sum b c) = sum (prod a b) (prod a c)
303   }.
304   
305 theorem bigop_distr: ∀n,p,B,nil.∀R:Dop B nil.∀f,a.
306   let aop ≝ sum B nil R in
307   let mop ≝ prod B nil R in
308   mop a \big[aop,nil]_{i<n|p i}(f i) = 
309    \big[aop,nil]_{i<n|p i}(mop a (f i)).
310 #n #p #B #nil #R #f #a normalize (elim n) [@null]
311 #n #Hind (cases (true_or_false (p n))) #H
312   [>bigop_Strue // >bigop_Strue // >(distr B nil R) >Hind //
313   |>bigop_Sfalse // >bigop_Sfalse //
314   ]
315 qed.
316   
317 (* Sigma e Pi *)
318
319 notation "∑_{ ident i < n | p } f"
320   with precedence 80
321 for @{'bigop $n plus 0 (λ${ident i}. $p) (λ${ident i}. $f)}.
322
323 notation "∑_{ ident i < n } f"
324   with precedence 80
325 for @{'bigop $n plus 0 (λ${ident i}.true) (λ${ident i}. $f)}.
326
327 notation "∑_{ ident j ∈ [a,b[ } f"
328   with precedence 80
329 for @{'bigop ($b-$a) plus 0 (λ${ident j}.((λ${ident j}.true) (${ident j}+$a)))
330   (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
331   
332 notation "∑_{ ident j ∈ [a,b[ | p } f"
333   with precedence 80
334 for @{'bigop ($b-$a) plus 0 (λ${ident j}.((λ${ident j}.$p) (${ident j}+$a)))
335   (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
336  
337 notation "∏_{ ident i < n | p} f"
338   with precedence 80
339 for @{'bigop $n times 1 (λ${ident i}.$p) (λ${ident i}. $f)}.
340  
341 notation "∏_{ ident i < n } f"
342   with precedence 80
343 for @{'bigop $n times 1 (λ${ident i}.true) (λ${ident i}. $f)}.
344
345 notation "∏_{ ident j ∈ [a,b[ } f"
346   with precedence 80
347 for @{'bigop ($b-$a) times 1 (λ${ident j}.((λ${ident j}.true) (${ident j}+$a)))
348   (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
349   
350 notation "∏_{ ident j ∈ [a,b[ | p } f"
351   with precedence 80
352 for @{'bigop ($b-$a) times 1 (λ${ident j}.((λ${ident j}.$p) (${ident j}+$a)))
353   (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
354
355
356 (*
357     
358 definition p_ord_times \def
359 \lambda p,m,x.
360   match p_ord x p with
361   [pair q r \Rightarrow r*m+q].
362   
363 theorem  eq_p_ord_times: \forall p,m,x.
364 p_ord_times p m x = (ord_rem x p)*m+(ord x p).
365 intros.unfold p_ord_times. unfold ord_rem.
366 unfold ord.
367 elim (p_ord x p).
368 reflexivity.
369 qed.
370
371 theorem div_p_ord_times: 
372 \forall p,m,x. ord x p < m \to p_ord_times p m x / m = ord_rem x p.
373 intros.rewrite > eq_p_ord_times.
374 apply div_plus_times.
375 assumption.
376 qed.
377
378 theorem mod_p_ord_times: 
379 \forall p,m,x. ord x p < m \to p_ord_times p m x \mod m = ord x p.
380 intros.rewrite > eq_p_ord_times.
381 apply mod_plus_times.
382 assumption.
383 qed.
384
385 lemma lt_times_to_lt_O: \forall i,n,m:nat. i < n*m \to O < m.
386 intros.
387 elim (le_to_or_lt_eq O ? (le_O_n m))
388   [assumption
389   |apply False_ind.
390    rewrite < H1 in H.
391    rewrite < times_n_O in H.
392    apply (not_le_Sn_O ? H)
393   ]
394 qed.
395
396 theorem iter_p_gen_knm:
397 \forall A:Type.
398 \forall baseA: A.
399 \forall plusA: A \to A \to A. 
400 (symmetric A plusA) \to 
401 (associative A plusA) \to 
402 (\forall a:A.(plusA a  baseA) = a)\to
403 \forall g: nat \to A.
404 \forall h2:nat \to nat \to nat.
405 \forall h11,h12:nat \to nat. 
406 \forall k,n,m.
407 \forall p1,p21:nat \to bool.
408 \forall p22:nat \to nat \to bool.
409 (\forall x. x < k \to p1 x = true \to 
410 p21 (h11 x) = true \land p22 (h11 x) (h12 x) = true
411 \land h2 (h11 x) (h12 x) = x 
412 \land (h11 x) < n \land (h12 x) < m) \to
413 (\forall i,j. i < n \to j < m \to p21 i = true \to p22 i j = true \to 
414 p1 (h2 i j) = true \land 
415 h11 (h2 i j) = i \land h12 (h2 i j) = j
416 \land h2 i j < k) \to
417 iter_p_gen k p1 A g baseA plusA =
418 iter_p_gen n p21 A (\lambda x:nat.iter_p_gen m (p22 x) A (\lambda y. g (h2 x y)) baseA plusA) baseA plusA.
419 intros.
420 rewrite < (iter_p_gen2' n m p21 p22 ? ? ? ? H H1 H2).
421 apply sym_eq.
422 apply (eq_iter_p_gen_gh A baseA plusA H H1 H2 g ? (\lambda x.(h11 x)*m+(h12 x)))
423  [intros.
424   elim (H4 (i/m) (i \mod m));clear H4
425    [elim H7.clear H7.
426     elim H4.clear H4.
427     assumption
428    |apply (lt_times_to_lt_div ? ? ? H5)
429    |apply lt_mod_m_m.
430     apply (lt_times_to_lt_O ? ? ? H5)
431    |apply (andb_true_true ? ? H6)
432    |apply (andb_true_true_r ? ? H6)
433    ]
434  |intros.
435   elim (H4 (i/m) (i \mod m));clear H4
436    [elim H7.clear H7.
437     elim H4.clear H4.
438     rewrite > H10.
439     rewrite > H9.
440     apply sym_eq.
441     apply div_mod.
442     apply (lt_times_to_lt_O ? ? ? H5)
443    |apply (lt_times_to_lt_div ? ? ? H5)
444    |apply lt_mod_m_m.
445     apply (lt_times_to_lt_O ? ? ? H5)  
446    |apply (andb_true_true ? ? H6)
447    |apply (andb_true_true_r ? ? H6)
448    ]
449  |intros.
450   elim (H4 (i/m) (i \mod m));clear H4
451    [elim H7.clear H7.
452     elim H4.clear H4.
453     assumption
454    |apply (lt_times_to_lt_div ? ? ? H5)
455    |apply lt_mod_m_m.
456     apply (lt_times_to_lt_O ? ? ? H5)
457    |apply (andb_true_true ? ? H6)
458    |apply (andb_true_true_r ? ? H6)
459    ]
460  |intros.
461   elim (H3 j H5 H6).
462   elim H7.clear H7.
463   elim H9.clear H9.
464   elim H7.clear H7.
465   rewrite > div_plus_times
466    [rewrite > mod_plus_times
467      [rewrite > H9.
468       rewrite > H12.
469       reflexivity.
470      |assumption
471      ]
472    |assumption
473    ]
474  |intros.
475   elim (H3 j H5 H6).
476   elim H7.clear H7.
477   elim H9.clear H9.
478   elim H7.clear H7. 
479   rewrite > div_plus_times
480    [rewrite > mod_plus_times
481      [assumption
482      |assumption
483      ]
484    |assumption
485    ]
486  |intros.
487   elim (H3 j H5 H6).
488   elim H7.clear H7.
489   elim H9.clear H9.
490   elim H7.clear H7. 
491   apply (lt_to_le_to_lt ? ((h11 j)*m+m))
492    [apply monotonic_lt_plus_r.
493     assumption
494    |rewrite > sym_plus.
495     change with ((S (h11 j)*m) \le n*m).
496     apply monotonic_le_times_l.
497     assumption
498    ]
499  ]
500 qed.
501
502 theorem iter_p_gen_divides:
503 \forall A:Type.
504 \forall baseA: A.
505 \forall plusA: A \to A \to A. 
506 \forall n,m,p:nat.O < n \to prime p \to Not (divides p n) \to 
507 \forall g: nat \to A.
508 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a  baseA) = a)
509
510 \to
511
512 iter_p_gen (S (n*(exp p m))) (\lambda x.divides_b x (n*(exp p m))) A g baseA plusA =
513 iter_p_gen (S n) (\lambda x.divides_b x n) A
514   (\lambda x.iter_p_gen (S m) (\lambda y.true) A (\lambda y.g (x*(exp p y))) baseA plusA) baseA plusA.
515 intros.
516 cut (O < p)
517   [rewrite < (iter_p_gen2 ? ? ? ? ? ? ? ? H3 H4 H5).
518    apply (trans_eq ? ? 
519     (iter_p_gen (S n*S m) (\lambda x:nat.divides_b (x/S m) n) A
520        (\lambda x:nat.g (x/S m*(p)\sup(x\mod S m)))  baseA plusA) )
521     [apply sym_eq.
522      apply (eq_iter_p_gen_gh ? ? ? ? ? ? g ? (p_ord_times p (S m)))
523       [ assumption
524       | assumption
525       | assumption
526       |intros.
527        lapply (divides_b_true_to_lt_O ? ? H H7).
528        apply divides_to_divides_b_true
529         [rewrite > (times_n_O O).
530          apply lt_times
531           [assumption
532           |apply lt_O_exp.assumption
533           ]
534         |apply divides_times
535           [apply divides_b_true_to_divides.assumption
536           |apply (witness ? ? (p \sup (m-i \mod (S m)))).
537            rewrite < exp_plus_times.
538            apply eq_f.
539            rewrite > sym_plus.
540            apply plus_minus_m_m.
541            autobatch by le_S_S_to_le, lt_mod_m_m, lt_O_S;
542           ]
543         ]
544       |intros.
545        lapply (divides_b_true_to_lt_O ? ? H H7).
546        unfold p_ord_times.
547        rewrite > (p_ord_exp1 p ? (i \mod (S m)) (i/S m))
548         [change with ((i/S m)*S m+i \mod S m=i).
549          apply sym_eq.
550          apply div_mod.
551          apply lt_O_S
552         |assumption
553         |unfold Not.intro.
554          apply H2.
555          apply (trans_divides ? (i/ S m))
556           [assumption|
557            apply divides_b_true_to_divides;assumption]
558         |apply sym_times.
559         ]
560       |intros.
561        apply le_S_S.
562        apply le_times
563         [apply le_S_S_to_le.
564          change with ((i/S m) < S n).
565          apply (lt_times_to_lt_l m).
566          apply (le_to_lt_to_lt ? i);[2:assumption]
567          autobatch by eq_plus_to_le, div_mod, lt_O_S.
568         |apply le_exp
569           [assumption
570           |apply le_S_S_to_le.
571            apply lt_mod_m_m.
572            apply lt_O_S
573           ]
574         ]
575       |intros.
576        cut (ord j p < S m)
577         [rewrite > div_p_ord_times
578           [apply divides_to_divides_b_true
579             [apply lt_O_ord_rem
580              [elim H1.assumption
581              |apply (divides_b_true_to_lt_O ? ? ? H7).
582                rewrite > (times_n_O O).
583                apply lt_times
584                [assumption|apply lt_O_exp.assumption]
585              ]
586            |cut (n = ord_rem (n*(exp p m)) p)
587               [rewrite > Hcut2.
588                apply divides_to_divides_ord_rem
589                 [apply (divides_b_true_to_lt_O ? ? ? H7).
590                  rewrite > (times_n_O O).
591                  apply lt_times
592                  [assumption|apply lt_O_exp.assumption]
593                  |rewrite > (times_n_O O).
594                    apply lt_times
595                   [assumption|apply lt_O_exp.assumption]
596                |assumption
597                |apply divides_b_true_to_divides.
598                 assumption
599                ]
600               |unfold ord_rem.
601               rewrite > (p_ord_exp1 p ? m n)
602                 [reflexivity
603                |assumption
604                 |assumption
605                |apply sym_times
606                ]
607              ]
608             ]
609           |assumption
610           ]
611         |cut (m = ord (n*(exp p m)) p)
612           [apply le_S_S.
613            rewrite > Hcut1.
614            apply divides_to_le_ord
615             [apply (divides_b_true_to_lt_O ? ? ? H7).
616              rewrite > (times_n_O O).
617              apply lt_times
618               [assumption|apply lt_O_exp.assumption]
619             |rewrite > (times_n_O O).
620              apply lt_times
621               [assumption|apply lt_O_exp.assumption]
622             |assumption
623             |apply divides_b_true_to_divides.
624              assumption
625             ]
626           |unfold ord.
627            rewrite > (p_ord_exp1 p ? m n)
628             [reflexivity
629             |assumption
630             |assumption
631             |apply sym_times
632             ]
633           ]
634         ]
635       |intros.
636        cut (ord j p < S m)
637         [rewrite > div_p_ord_times
638           [rewrite > mod_p_ord_times
639             [rewrite > sym_times.
640              apply sym_eq.
641              apply exp_ord
642               [elim H1.assumption
643               |apply (divides_b_true_to_lt_O ? ? ? H7).
644                rewrite > (times_n_O O).
645                apply lt_times
646                 [assumption|apply lt_O_exp.assumption]
647               ]
648            |cut (m = ord (n*(exp p m)) p)
649              [apply le_S_S.
650               rewrite > Hcut2.
651               apply divides_to_le_ord
652                [apply (divides_b_true_to_lt_O ? ? ? H7).
653                 rewrite > (times_n_O O).
654                 apply lt_times
655                  [assumption|apply lt_O_exp.assumption]
656                |rewrite > (times_n_O O).
657                 apply lt_times
658                  [assumption|apply lt_O_exp.assumption]
659                |assumption
660                |apply divides_b_true_to_divides.
661                 assumption
662                ]
663              |unfold ord.
664               rewrite > (p_ord_exp1 p ? m n)
665                 [reflexivity
666                 |assumption
667                 |assumption
668                 |apply sym_times
669                 ]
670               ]
671             ]
672           |assumption
673           ]
674         |cut (m = ord (n*(exp p m)) p)
675           [apply le_S_S.
676            rewrite > Hcut1.
677            apply divides_to_le_ord
678             [apply (divides_b_true_to_lt_O ? ? ? H7).
679              rewrite > (times_n_O O).
680              apply lt_times
681               [assumption|apply lt_O_exp.assumption]
682             |rewrite > (times_n_O O).
683              apply lt_times
684               [assumption|apply lt_O_exp.assumption]
685             |assumption
686             |apply divides_b_true_to_divides.
687              assumption
688             ]
689           |unfold ord.
690            rewrite > (p_ord_exp1 p ? m n)
691             [reflexivity
692             |assumption
693             |assumption
694             |apply sym_times
695             ]
696           ]
697         ]
698       |intros.
699        rewrite > eq_p_ord_times.
700        rewrite > sym_plus.
701        apply (lt_to_le_to_lt ? (S m +ord_rem j p*S m))
702         [apply lt_plus_l.
703          apply le_S_S.
704          cut (m = ord (n*(p \sup m)) p)
705           [rewrite > Hcut1.
706            apply divides_to_le_ord
707             [apply (divides_b_true_to_lt_O ? ? ? H7).
708              rewrite > (times_n_O O).
709              apply lt_times
710               [assumption|apply lt_O_exp.assumption]
711             |rewrite > (times_n_O O).
712              apply lt_times
713               [assumption|apply lt_O_exp.assumption]
714             |assumption
715             |apply divides_b_true_to_divides.
716              assumption
717             ]
718           |unfold ord.
719            rewrite > sym_times.
720            rewrite > (p_ord_exp1 p ? m n)
721             [reflexivity
722             |assumption
723             |assumption
724             |reflexivity
725             ]
726           ]
727         |change with (S (ord_rem j p)*S m \le S n*S m).
728          apply le_times_l.
729          apply le_S_S.
730          cut (n = ord_rem (n*(p \sup m)) p)
731           [rewrite > Hcut1.
732            apply divides_to_le
733             [apply lt_O_ord_rem
734               [elim H1.assumption
735               |rewrite > (times_n_O O).
736                apply lt_times
737                 [assumption|apply lt_O_exp.assumption]
738               ]
739             |apply divides_to_divides_ord_rem
740               [apply (divides_b_true_to_lt_O ? ? ? H7).
741                rewrite > (times_n_O O).
742                apply lt_times
743                 [assumption|apply lt_O_exp.assumption]
744               |rewrite > (times_n_O O).
745                apply lt_times
746                 [assumption|apply lt_O_exp.assumption]
747               |assumption
748               |apply divides_b_true_to_divides.
749                assumption
750               ]
751             ]
752         |unfold ord_rem.
753          rewrite > sym_times.
754          rewrite > (p_ord_exp1 p ? m n)
755           [reflexivity
756           |assumption
757           |assumption
758           |reflexivity
759           ]
760         ]
761       ]
762     ]
763   |apply eq_iter_p_gen
764   
765     [intros.
766      elim (divides_b (x/S m) n);reflexivity
767     |intros.reflexivity
768     ]
769   ]
770 |elim H1.apply lt_to_le.assumption
771 ]
772 qed.
773     
774
775
776 theorem iter_p_gen_2_eq: 
777 \forall A:Type.
778 \forall baseA: A.
779 \forall plusA: A \to A \to A. 
780 (symmetric A plusA) \to 
781 (associative A plusA) \to 
782 (\forall a:A.(plusA a  baseA) = a)\to
783 \forall g: nat \to nat \to A.
784 \forall h11,h12,h21,h22: nat \to nat \to nat. 
785 \forall n1,m1,n2,m2.
786 \forall p11,p21:nat \to bool.
787 \forall p12,p22:nat \to nat \to bool.
788 (\forall i,j. i < n2 \to j < m2 \to p21 i = true \to p22 i j = true \to 
789 p11 (h11 i j) = true \land p12 (h11 i j) (h12 i j) = true
790 \land h21 (h11 i j) (h12 i j) = i \land h22 (h11 i j) (h12 i j) = j
791 \land h11 i j < n1 \land h12 i j < m1) \to
792 (\forall i,j. i < n1 \to j < m1 \to p11 i = true \to p12 i j = true \to 
793 p21 (h21 i j) = true \land p22 (h21 i j) (h22 i j) = true
794 \land h11 (h21 i j) (h22 i j) = i \land h12 (h21 i j) (h22 i j) = j
795 \land (h21 i j) < n2 \land (h22 i j) < m2) \to
796 iter_p_gen n1 p11 A 
797      (\lambda x:nat .iter_p_gen m1 (p12 x) A (\lambda y. g x y) baseA plusA) 
798      baseA plusA =
799 iter_p_gen n2 p21 A 
800     (\lambda x:nat .iter_p_gen m2 (p22 x) A  (\lambda y. g (h11 x y) (h12 x y)) baseA plusA )
801     baseA plusA.
802
803 intros.
804 rewrite < (iter_p_gen2' ? ? ? ? ? ? ? ? H H1 H2).
805 letin ha:= (\lambda x,y.(((h11 x y)*m1) + (h12 x y))).
806 letin ha12:= (\lambda x.(h21 (x/m1) (x \mod m1))).
807 letin ha22:= (\lambda x.(h22 (x/m1) (x \mod m1))).
808
809 apply (trans_eq ? ? 
810 (iter_p_gen n2 p21 A (\lambda x:nat. iter_p_gen m2 (p22 x) A
811  (\lambda y:nat.(g (((h11 x y)*m1+(h12 x y))/m1) (((h11 x y)*m1+(h12 x y))\mod m1))) baseA plusA ) baseA plusA))
812 [
813   apply (iter_p_gen_knm A baseA plusA H H1 H2 (\lambda e. (g (e/m1) (e \mod m1))) ha ha12 ha22);intros
814   [ elim (and_true ? ? H6).
815     cut(O \lt m1)
816     [ cut(x/m1 < n1)
817       [ cut((x \mod m1) < m1)
818         [ elim (H4 ? ? Hcut1 Hcut2 H7 H8).
819           elim H9.clear H9.
820           elim H11.clear H11.
821           elim H9.clear H9.
822           elim H11.clear H11.
823           split
824           [ split
825             [ split
826               [ split
827                 [ assumption
828                 | assumption
829                 ]
830               | unfold ha.
831                 unfold ha12.
832                 unfold ha22.
833                 rewrite > H14.
834                 rewrite > H13.
835                 apply sym_eq.
836                 apply div_mod.
837                 assumption
838               ]
839             | assumption
840             ]
841           | assumption
842           ]
843         | apply lt_mod_m_m.
844           assumption
845         ]
846       | apply (lt_times_n_to_lt m1)
847         [ assumption
848         | apply (le_to_lt_to_lt ? x)
849           [ apply (eq_plus_to_le ? ? (x \mod m1)).
850             apply div_mod.
851             assumption
852           | assumption
853         ]
854       ]  
855     ]
856     | apply not_le_to_lt.unfold.intro.
857       generalize in match H5.
858       apply (le_n_O_elim ? H9).
859       rewrite < times_n_O.
860       apply le_to_not_lt.
861       apply le_O_n.              
862     ]
863   | elim (H3 ? ? H5 H6 H7 H8).
864     elim H9.clear H9.
865     elim H11.clear H11.
866     elim H9.clear H9.
867     elim H11.clear H11.
868     cut(((h11 i j)*m1 + (h12 i j))/m1 = (h11 i j))
869     [ cut(((h11 i j)*m1 + (h12 i j)) \mod m1 = (h12 i j))
870       [ split
871         [ split
872           [ split
873             [ apply true_to_true_to_andb_true
874               [ rewrite > Hcut.
875                 assumption
876               | rewrite > Hcut1.
877                 rewrite > Hcut.
878                 assumption
879               ] 
880             | unfold ha.
881               unfold ha12.
882               rewrite > Hcut1.
883               rewrite > Hcut.
884               assumption
885             ]
886           | unfold ha.
887             unfold ha22.
888             rewrite > Hcut1.
889             rewrite > Hcut.
890             assumption            
891           ]
892         | cut(O \lt m1)
893           [ cut(O \lt n1)      
894             [ apply (lt_to_le_to_lt ? ((h11 i j)*m1 + m1) )
895               [ unfold ha.
896                 apply (lt_plus_r).
897                 assumption
898               | rewrite > sym_plus.
899                 rewrite > (sym_times (h11 i j) m1).
900                 rewrite > times_n_Sm.
901                 rewrite > sym_times.
902                 apply (le_times_l).
903                 assumption  
904               ]
905             | apply not_le_to_lt.unfold.intro.
906               generalize in match H12.
907               apply (le_n_O_elim ? H11).       
908               apply le_to_not_lt.
909               apply le_O_n
910             ]
911           | apply not_le_to_lt.unfold.intro.
912             generalize in match H10.
913             apply (le_n_O_elim ? H11).       
914             apply le_to_not_lt.
915             apply le_O_n
916           ]  
917         ]
918       | rewrite > (mod_plus_times m1 (h11 i j) (h12 i j)).
919         reflexivity.
920         assumption
921       ]     
922     | rewrite > (div_plus_times m1 (h11 i j) (h12 i j)).
923       reflexivity.
924       assumption
925     ]
926   ]
927 | apply (eq_iter_p_gen1)
928   [ intros. reflexivity 
929   | intros.
930     apply (eq_iter_p_gen1)
931     [ intros. reflexivity
932     | intros.
933       rewrite > (div_plus_times)
934       [ rewrite > (mod_plus_times)
935         [ reflexivity
936         | elim (H3 x x1 H5 H7 H6 H8).
937           assumption
938         ]
939       | elim (H3 x x1 H5 H7 H6 H8).       
940         assumption
941       ]
942     ]
943   ]
944 ]
945 qed.
946
947 theorem iter_p_gen_iter_p_gen: 
948 \forall A:Type.
949 \forall baseA: A.
950 \forall plusA: A \to A \to A. 
951 (symmetric A plusA) \to 
952 (associative A plusA) \to 
953 (\forall a:A.(plusA a  baseA) = a)\to
954 \forall g: nat \to nat \to A.
955 \forall n,m.
956 \forall p11,p21:nat \to bool.
957 \forall p12,p22:nat \to nat \to bool.
958 (\forall x,y. x < n \to y < m \to 
959  (p11 x \land p12 x y) = (p21 y \land p22 y x)) \to
960 iter_p_gen n p11 A 
961   (\lambda x:nat.iter_p_gen m (p12 x) A (\lambda y. g x y) baseA plusA) 
962   baseA plusA =
963 iter_p_gen m p21 A 
964   (\lambda y:nat.iter_p_gen n (p22 y) A  (\lambda x. g x y) baseA plusA )
965   baseA plusA.
966 intros.
967 apply (iter_p_gen_2_eq A baseA plusA H H1 H2 (\lambda x,y. g x y) (\lambda x,y.y) (\lambda x,y.x) (\lambda x,y.y) (\lambda x,y.x)
968        n m m n p11 p21 p12 p22)
969   [intros.split
970     [split
971       [split
972         [split
973           [split
974             [apply (andb_true_true ? (p12 j i)).
975              rewrite > H3
976               [rewrite > H6.rewrite > H7.reflexivity
977               |assumption
978               |assumption
979               ]
980             |apply (andb_true_true_r (p11 j)).
981              rewrite > H3
982               [rewrite > H6.rewrite > H7.reflexivity
983               |assumption
984               |assumption
985               ]
986             ]
987           |reflexivity
988           ]
989         |reflexivity
990         ]
991       |assumption
992       ]
993     |assumption
994     ]
995   |intros.split
996     [split
997       [split
998         [split
999           [split
1000             [apply (andb_true_true ? (p22 j i)).
1001              rewrite < H3
1002               [rewrite > H6.rewrite > H7.reflexivity
1003               |assumption
1004               |assumption
1005               ]
1006             |apply (andb_true_true_r (p21 j)).
1007              rewrite < H3
1008               [rewrite > H6.rewrite > H7.reflexivity
1009               |assumption
1010               |assumption
1011               ]
1012             ]
1013           |reflexivity
1014           ]
1015         |reflexivity
1016         ]
1017       |assumption
1018       ]
1019     |assumption
1020     ]
1021   ]
1022 qed. *)