1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "nat/div_and_mod_diseq.ma".
16 include "nat/primes.ma".
19 (*a generic definition of summatory indexed over natural numbers:
20 * n:N is the advanced end in the range of the sommatory
21 * p: N -> bool is a predicate. if (p i) = true, then (g i) is summed, else it isn't
22 * A is the type of the elements of the sum.
23 * g:nat -> A, is the function which associates the nth element of the sum to the nth natural numbers
24 * baseA (of type A) is the neutral element of A for plusA operation
25 * plusA: A -> A -> A is the sum over elements in A.
28 let rec iter_p_gen (n:nat) (p: nat \to bool) (A:Type) (g: nat \to A)
29 (baseA: A) (plusA: A \to A \to A) \def
34 [true \Rightarrow (plusA (g k) (iter_p_gen k p A g baseA plusA))
35 |false \Rightarrow iter_p_gen k p A g baseA plusA]
38 theorem true_to_iter_p_gen_Sn:
39 \forall n:nat. \forall p:nat \to bool. \forall A:Type. \forall g:nat \to A.
40 \forall baseA:A. \forall plusA: A \to A \to A.
41 p n = true \to iter_p_gen (S n) p A g baseA plusA =
42 (plusA (g n) (iter_p_gen n p A g baseA plusA)).
49 theorem false_to_iter_p_gen_Sn:
50 \forall n:nat. \forall p:nat \to bool. \forall A:Type. \forall g:nat \to A.
51 \forall baseA:A. \forall plusA: A \to A \to A.
52 p n = false \to iter_p_gen (S n) p A g baseA plusA = iter_p_gen n p A g baseA plusA.
59 theorem eq_iter_p_gen: \forall p1,p2:nat \to bool. \forall A:Type.
60 \forall g1,g2: nat \to A. \forall baseA: A.
61 \forall plusA: A \to A \to A. \forall n:nat.
62 (\forall x. x < n \to p1 x = p2 x) \to
63 (\forall x. x < n \to g1 x = g2 x) \to
64 iter_p_gen n p1 A g1 baseA plusA = iter_p_gen n p2 A g2 baseA plusA.
68 | apply (bool_elim ? (p1 n1))
70 rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H3).
71 rewrite > true_to_iter_p_gen_Sn
73 [ apply H2.apply le_n.
75 [ intros.apply H1.apply le_S.assumption
76 | intros.apply H2.apply le_S.assumption
79 | rewrite < H3.apply sym_eq.apply H1.apply le_n
82 rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H3).
83 rewrite > false_to_iter_p_gen_Sn
85 [ intros.apply H1.apply le_S.assumption
86 | intros.apply H2.apply le_S.assumption
88 | rewrite < H3.apply sym_eq.apply H1.apply le_n
94 theorem eq_iter_p_gen1: \forall p1,p2:nat \to bool. \forall A:Type.
95 \forall g1,g2: nat \to A. \forall baseA: A.
96 \forall plusA: A \to A \to A.\forall n:nat.
97 (\forall x. x < n \to p1 x = p2 x) \to
98 (\forall x. x < n \to p1 x = true \to g1 x = g2 x) \to
99 iter_p_gen n p1 A g1 baseA plusA = iter_p_gen n p2 A g2 baseA plusA.
103 | apply (bool_elim ? (p1 n1))
105 rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H3).
106 rewrite > true_to_iter_p_gen_Sn
113 [ intros.apply H1.apply le_S.assumption
115 [ apply le_S.assumption
121 apply sym_eq.apply H1.apply le_n
124 rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H3).
125 rewrite > false_to_iter_p_gen_Sn
127 [ intros.apply H1.apply le_S.assumption
129 [ apply le_S.assumption
133 | rewrite < H3.apply sym_eq.
140 theorem iter_p_gen_false: \forall A:Type. \forall g: nat \to A. \forall baseA:A.
141 \forall plusA: A \to A \to A. \forall n.
142 iter_p_gen n (\lambda x.false) A g baseA plusA = baseA.
151 theorem iter_p_gen_plusA: \forall A:Type. \forall n,k:nat.\forall p:nat \to bool.
152 \forall g: nat \to A. \forall baseA:A. \forall plusA: A \to A \to A.
153 (symmetric A plusA) \to (\forall a:A. (plusA a baseA) = a) \to (associative A plusA)
155 iter_p_gen (k + n) p A g baseA plusA
156 = (plusA (iter_p_gen k (\lambda x.p (x+n)) A (\lambda x.g (x+n)) baseA plusA)
157 (iter_p_gen n p A g baseA plusA)).
162 rewrite > H in \vdash (? ? ? %).
165 | apply (bool_elim ? (p (n1+n)))
167 rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
168 rewrite > (true_to_iter_p_gen_Sn n1 (\lambda x.p (x+n)) ? ? ? ? H4).
169 rewrite > (H2 (g (n1 + n)) ? ?).
173 rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
174 rewrite > (false_to_iter_p_gen_Sn n1 (\lambda x.p (x+n)) ? ? ? ? H4).
180 theorem false_to_eq_iter_p_gen: \forall A:Type. \forall n,m:nat.\forall p:nat \to bool.
181 \forall g: nat \to A. \forall baseA:A. \forall plusA: A \to A \to A.
182 n \le m \to (\forall i:nat. n \le i \to i < m \to p i = false)
183 \to iter_p_gen m p A g baseA plusA = iter_p_gen n p A g baseA plusA.
203 (* a therem slightly more general than the previous one *)
204 theorem or_false_eq_baseA_to_eq_iter_p_gen: \forall A:Type. \forall n,m:nat.\forall p:nat \to bool.
205 \forall g: nat \to A. \forall baseA:A. \forall plusA: A \to A \to A.
206 (\forall a. plusA baseA a = a) \to
207 n \le m \to (\forall i:nat. n \le i \to i < m \to p i = false \lor g i = baseA)
208 \to iter_p_gen m p A g baseA plusA = iter_p_gen n p A g baseA plusA.
212 |apply (bool_elim ? (p n1));intro
215 apply not_eq_true_false.
219 |rewrite > true_to_iter_p_gen_Sn
225 |apply le_S.assumption
232 |rewrite > false_to_iter_p_gen_Sn
236 |apply le_S.assumption
244 theorem iter_p_gen2 :
246 \forall p1,p2:nat \to bool.
248 \forall g: nat \to nat \to A.
250 \forall plusA: A \to A \to A.
251 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a)
254 (\lambda x.andb (p1 (div x m)) (p2 (mod x m)))
256 (\lambda x.g (div x m) (mod x m))
260 (\lambda x.iter_p_gen m p2 A (g x) baseA plusA)
266 | apply (bool_elim ? (p1 n1))
268 rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
269 simplify in \vdash (? ? (? % ? ? ? ? ?) ?).
270 rewrite > iter_p_gen_plusA
273 [ apply eq_iter_p_gen
276 rewrite > (div_plus_times ? ? ? H5).
277 rewrite > (mod_plus_times ? ? ? H5).
283 rewrite > (div_plus_times ? ? ? H5).
284 rewrite > (mod_plus_times ? ? ? H5).
294 rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
295 simplify in \vdash (? ? (? % ? ? ? ? ?) ?).
296 rewrite > iter_p_gen_plusA
298 apply (trans_eq ? ? (plusA baseA
299 (iter_p_gen n1 p1 A (\lambda x:nat.iter_p_gen m p2 A (g x) baseA plusA) baseA plusA )))
301 [ rewrite > (eq_iter_p_gen ? (\lambda x.false) A ? (\lambda x:nat.g ((x+n1*m)/m) ((x+n1*m)\mod m)))
302 [ apply iter_p_gen_false
305 rewrite > (div_plus_times ? ? ? H5).
306 rewrite > (mod_plus_times ? ? ? H5).
309 | intros.reflexivity.
325 theorem iter_p_gen2':
327 \forall p1: nat \to bool.
328 \forall p2: nat \to nat \to bool.
330 \forall g: nat \to nat \to A.
332 \forall plusA: A \to A \to A.
333 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a)
336 (\lambda x.andb (p1 (div x m)) (p2 (div x m)(mod x m)))
338 (\lambda x.g (div x m) (mod x m))
342 (\lambda x.iter_p_gen m (p2 x) A (g x) baseA plusA)
348 | apply (bool_elim ? (p1 n1))
350 rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
351 simplify in \vdash (? ? (? % ? ? ? ? ?) ?).
352 rewrite > iter_p_gen_plusA
355 [ apply eq_iter_p_gen
358 rewrite > (div_plus_times ? ? ? H5).
359 rewrite > (mod_plus_times ? ? ? H5).
365 rewrite > (div_plus_times ? ? ? H5).
366 rewrite > (mod_plus_times ? ? ? H5).
376 rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
377 simplify in \vdash (? ? (? % ? ? ? ? ?) ?).
378 rewrite > iter_p_gen_plusA
380 apply (trans_eq ? ? (plusA baseA
381 (iter_p_gen n1 p1 A (\lambda x:nat.iter_p_gen m (p2 x) A (g x) baseA plusA) baseA plusA )))
383 [ rewrite > (eq_iter_p_gen ? (\lambda x.false) A ? (\lambda x:nat.g ((x+n1*m)/m) ((x+n1*m)\mod m)))
384 [ apply iter_p_gen_false
387 rewrite > (div_plus_times ? ? ? H5).
388 rewrite > (mod_plus_times ? ? ? H5).
391 | intros.reflexivity.
409 \forall g: nat \to A.
411 \forall plusA: A \to A \to A.
413 \forall p:nat \to bool.
414 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a)
417 i < n \to p i = true \to
418 (iter_p_gen n p A g baseA plusA) =
419 (plusA (g i) (iter_p_gen n (\lambda x:nat. andb (p x) (notb (eqb x i))) A g baseA plusA)).
423 apply (not_le_Sn_O i).
425 | apply (bool_elim ? (p n1));intro
426 [ elim (le_to_or_lt_eq i n1)
427 [ rewrite > true_to_iter_p_gen_Sn
428 [ rewrite > true_to_iter_p_gen_Sn
429 [ rewrite < (H2 (g i) ? ?).
430 rewrite > (H1 (g i) (g n1)).
431 rewrite > (H2 (g n1) ? ?).
442 | rewrite > H6.simplify.
443 change with (notb (eqb n1 i) = notb false).
445 apply not_eq_to_eqb_false.
447 apply (lt_to_not_eq ? ? H7).
448 apply sym_eq.assumption
452 | rewrite > true_to_iter_p_gen_Sn
456 | rewrite > false_to_iter_p_gen_Sn
457 [ apply eq_iter_p_gen
461 change with (notb false = notb (eqb x n1)).
464 apply not_eq_to_eqb_false.
465 apply (lt_to_not_eq ? ? H8)
472 rewrite > (eq_to_eqb_true ? ? (refl_eq ? n1)).
478 | apply le_S_S_to_le.
481 | rewrite > false_to_iter_p_gen_Sn
482 [ elim (le_to_or_lt_eq i n1)
483 [ rewrite > false_to_iter_p_gen_Sn
491 | rewrite > H6.reflexivity
494 apply not_eq_true_false.
498 | apply le_S_S_to_le.
507 (* invariance under permutation; single sum *)
508 theorem eq_iter_p_gen_gh:
511 \forall plusA: A \to A \to A.
512 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a) \to
513 \forall g: nat \to A.
514 \forall h,h1: nat \to nat.
516 \forall p1,p2:nat \to bool.
517 (\forall i. i < n \to p1 i = true \to p2 (h i) = true) \to
518 (\forall i. i < n \to p1 i = true \to h1 (h i) = i) \to
519 (\forall i. i < n \to p1 i = true \to h i < n1) \to
520 (\forall j. j < n1 \to p2 j = true \to p1 (h1 j) = true) \to
521 (\forall j. j < n1 \to p2 j = true \to h (h1 j) = j) \to
522 (\forall j. j < n1 \to p2 j = true \to h1 j < n) \to
524 iter_p_gen n p1 A (\lambda x.g(h x)) baseA plusA =
525 iter_p_gen n1 p2 A g baseA plusA.
528 [ generalize in match H8.
531 | apply (bool_elim ? (p2 n2));intro
533 apply (not_le_Sn_O (h1 n2)).
538 | rewrite > false_to_iter_p_gen_Sn
550 | apply (bool_elim ? (p1 n1));intro
551 [ rewrite > true_to_iter_p_gen_Sn
552 [ rewrite > (iter_p_gen_gi A g baseA plusA n2 (h n1))
559 change with ((\not eqb (h i) (h n1))= \not false).
561 apply not_eq_to_eqb_false.
564 apply (lt_to_not_eq ? ? H11).
576 | apply le_S.assumption
587 [ apply le_S.assumption
593 | generalize in match H12.
602 | generalize in match H12.
609 elim (le_to_or_lt_eq (h1 j) n1)
611 | generalize in match H12.
621 | apply eqb_false_to_not_eq.
622 generalize in match H14.
624 [ apply sym_eq.assumption
630 apply not_eq_true_false.
631 apply sym_eq.assumption
633 | apply le_S_S_to_le.
636 | generalize in match H12.
659 | rewrite > false_to_iter_p_gen_Sn
662 apply H4[apply le_S.assumption|assumption]
664 apply H5[apply le_S.assumption|assumption]
666 apply H6[apply le_S.assumption|assumption]
668 apply H7[assumption|assumption]
670 apply H8[assumption|assumption]
672 elim (le_to_or_lt_eq (h1 j) n1)
674 | absurd (j = (h n1))
682 apply not_eq_true_false.
691 | apply le_S_S_to_le.
704 theorem eq_iter_p_gen_pred:
707 \forall plusA: A \to A \to A.
710 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a) \to
711 iter_p_gen (S n) (\lambda i.p (pred i)) A (\lambda i.g(pred i)) baseA plusA =
712 plusA (iter_p_gen n p A g baseA plusA) (g O).
715 [rewrite > true_to_iter_p_gen_Sn
719 |apply (bool_elim ? (p n1));intro
720 [rewrite > true_to_iter_p_gen_Sn
721 [rewrite > true_to_iter_p_gen_Sn in ⊢ (? ? ? %)
722 [rewrite > H2 in ⊢ (? ? ? %).
723 apply eq_f.assumption
728 |rewrite > false_to_iter_p_gen_Sn
729 [rewrite > false_to_iter_p_gen_Sn in ⊢ (? ? ? %);assumption
736 definition p_ord_times \def
739 [pair q r \Rightarrow r*m+q].
741 theorem eq_p_ord_times: \forall p,m,x.
742 p_ord_times p m x = (ord_rem x p)*m+(ord x p).
743 intros.unfold p_ord_times. unfold ord_rem.
749 theorem div_p_ord_times:
750 \forall p,m,x. ord x p < m \to p_ord_times p m x / m = ord_rem x p.
751 intros.rewrite > eq_p_ord_times.
752 apply div_plus_times.
756 theorem mod_p_ord_times:
757 \forall p,m,x. ord x p < m \to p_ord_times p m x \mod m = ord x p.
758 intros.rewrite > eq_p_ord_times.
759 apply mod_plus_times.
763 lemma lt_times_to_lt_O: \forall i,n,m:nat. i < n*m \to O < m.
765 elim (le_to_or_lt_eq O ? (le_O_n m))
769 rewrite < times_n_O in H.
770 apply (not_le_Sn_O ? H)
774 theorem iter_p_gen_knm:
777 \forall plusA: A \to A \to A.
778 (symmetric A plusA) \to
779 (associative A plusA) \to
780 (\forall a:A.(plusA a baseA) = a)\to
781 \forall g: nat \to A.
782 \forall h2:nat \to nat \to nat.
783 \forall h11,h12:nat \to nat.
785 \forall p1,p21:nat \to bool.
786 \forall p22:nat \to nat \to bool.
787 (\forall x. x < k \to p1 x = true \to
788 p21 (h11 x) = true \land p22 (h11 x) (h12 x) = true
789 \land h2 (h11 x) (h12 x) = x
790 \land (h11 x) < n \land (h12 x) < m) \to
791 (\forall i,j. i < n \to j < m \to p21 i = true \to p22 i j = true \to
792 p1 (h2 i j) = true \land
793 h11 (h2 i j) = i \land h12 (h2 i j) = j
794 \land h2 i j < k) \to
795 iter_p_gen k p1 A g baseA plusA =
796 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.
798 rewrite < (iter_p_gen2' n m p21 p22 ? ? ? ? H H1 H2).
800 apply (eq_iter_p_gen_gh A baseA plusA H H1 H2 g ? (\lambda x.(h11 x)*m+(h12 x)))
802 elim (H4 (i/m) (i \mod m));clear H4
806 |apply (lt_times_to_lt_div ? ? ? H5)
808 apply (lt_times_to_lt_O ? ? ? H5)
809 |apply (andb_true_true ? ? H6)
810 |apply (andb_true_true_r ? ? H6)
813 elim (H4 (i/m) (i \mod m));clear H4
820 apply (lt_times_to_lt_O ? ? ? H5)
821 |apply (lt_times_to_lt_div ? ? ? H5)
823 apply (lt_times_to_lt_O ? ? ? H5)
824 |apply (andb_true_true ? ? H6)
825 |apply (andb_true_true_r ? ? H6)
828 elim (H4 (i/m) (i \mod m));clear H4
832 |apply (lt_times_to_lt_div ? ? ? H5)
834 apply (lt_times_to_lt_O ? ? ? H5)
835 |apply (andb_true_true ? ? H6)
836 |apply (andb_true_true_r ? ? H6)
843 rewrite > div_plus_times
844 [rewrite > mod_plus_times
857 rewrite > div_plus_times
858 [rewrite > mod_plus_times
869 apply (lt_to_le_to_lt ? ((h11 j)*m+m))
870 [apply monotonic_lt_plus_r.
873 change with ((S (h11 j)*m) \le n*m).
874 apply monotonic_le_times_l.
880 theorem iter_p_gen_divides:
883 \forall plusA: A \to A \to A.
884 \forall n,m,p:nat.O < n \to prime p \to Not (divides p n) \to
885 \forall g: nat \to A.
886 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a)
890 iter_p_gen (S (n*(exp p m))) (\lambda x.divides_b x (n*(exp p m))) A g baseA plusA =
891 iter_p_gen (S n) (\lambda x.divides_b x n) A
892 (\lambda x.iter_p_gen (S m) (\lambda y.true) A (\lambda y.g (x*(exp p y))) baseA plusA) baseA plusA.
895 [rewrite < (iter_p_gen2 ? ? ? ? ? ? ? ? H3 H4 H5).
897 (iter_p_gen (S n*S m) (\lambda x:nat.divides_b (x/S m) n) A
898 (\lambda x:nat.g (x/S m*(p)\sup(x\mod S m))) baseA plusA) )
900 apply (eq_iter_p_gen_gh ? ? ? ? ? ? g ? (p_ord_times p (S m)))
905 lapply (divides_b_true_to_lt_O ? ? H H7).
906 apply divides_to_divides_b_true
907 [rewrite > (times_n_O O).
910 |apply lt_O_exp.assumption
913 [apply divides_b_true_to_divides.assumption
914 |apply (witness ? ? (p \sup (m-i \mod (S m)))).
915 rewrite < exp_plus_times.
918 apply plus_minus_m_m.
923 lapply (divides_b_true_to_lt_O ? ? H H7).
925 rewrite > (p_ord_exp1 p ? (i \mod (S m)) (i/S m))
926 [change with ((i/S m)*S m+i \mod S m=i).
933 apply (trans_divides ? (i/ S m))
935 apply divides_b_true_to_divides;assumption]
942 change with ((i/S m) < S n).
943 apply (lt_times_to_lt_l m).
944 apply (le_to_lt_to_lt ? i)
945 [autobatch|assumption]
955 [rewrite > div_p_ord_times
956 [apply divides_to_divides_b_true
959 |apply (divides_b_true_to_lt_O ? ? ? H7).
960 rewrite > (times_n_O O).
962 [assumption|apply lt_O_exp.assumption]
964 |cut (n = ord_rem (n*(exp p m)) p)
966 apply divides_to_divides_ord_rem
967 [apply (divides_b_true_to_lt_O ? ? ? H7).
968 rewrite > (times_n_O O).
970 [assumption|apply lt_O_exp.assumption]
971 |rewrite > (times_n_O O).
973 [assumption|apply lt_O_exp.assumption]
975 |apply divides_b_true_to_divides.
979 rewrite > (p_ord_exp1 p ? m n)
989 |cut (m = ord (n*(exp p m)) p)
992 apply divides_to_le_ord
993 [apply (divides_b_true_to_lt_O ? ? ? H7).
994 rewrite > (times_n_O O).
996 [assumption|apply lt_O_exp.assumption]
997 |rewrite > (times_n_O O).
999 [assumption|apply lt_O_exp.assumption]
1001 |apply divides_b_true_to_divides.
1005 rewrite > (p_ord_exp1 p ? m n)
1015 [rewrite > div_p_ord_times
1016 [rewrite > mod_p_ord_times
1017 [rewrite > sym_times.
1021 |apply (divides_b_true_to_lt_O ? ? ? H7).
1022 rewrite > (times_n_O O).
1024 [assumption|apply lt_O_exp.assumption]
1026 |cut (m = ord (n*(exp p m)) p)
1029 apply divides_to_le_ord
1030 [apply (divides_b_true_to_lt_O ? ? ? H7).
1031 rewrite > (times_n_O O).
1033 [assumption|apply lt_O_exp.assumption]
1034 |rewrite > (times_n_O O).
1036 [assumption|apply lt_O_exp.assumption]
1038 |apply divides_b_true_to_divides.
1042 rewrite > (p_ord_exp1 p ? m n)
1052 |cut (m = ord (n*(exp p m)) p)
1055 apply divides_to_le_ord
1056 [apply (divides_b_true_to_lt_O ? ? ? H7).
1057 rewrite > (times_n_O O).
1059 [assumption|apply lt_O_exp.assumption]
1060 |rewrite > (times_n_O O).
1062 [assumption|apply lt_O_exp.assumption]
1064 |apply divides_b_true_to_divides.
1068 rewrite > (p_ord_exp1 p ? m n)
1077 rewrite > eq_p_ord_times.
1079 apply (lt_to_le_to_lt ? (S m +ord_rem j p*S m))
1082 cut (m = ord (n*(p \sup m)) p)
1084 apply divides_to_le_ord
1085 [apply (divides_b_true_to_lt_O ? ? ? H7).
1086 rewrite > (times_n_O O).
1088 [assumption|apply lt_O_exp.assumption]
1089 |rewrite > (times_n_O O).
1091 [assumption|apply lt_O_exp.assumption]
1093 |apply divides_b_true_to_divides.
1097 rewrite > sym_times.
1098 rewrite > (p_ord_exp1 p ? m n)
1105 |change with (S (ord_rem j p)*S m \le S n*S m).
1108 cut (n = ord_rem (n*(p \sup m)) p)
1113 |rewrite > (times_n_O O).
1115 [assumption|apply lt_O_exp.assumption]
1117 |apply divides_to_divides_ord_rem
1118 [apply (divides_b_true_to_lt_O ? ? ? H7).
1119 rewrite > (times_n_O O).
1121 [assumption|apply lt_O_exp.assumption]
1122 |rewrite > (times_n_O O).
1124 [assumption|apply lt_O_exp.assumption]
1126 |apply divides_b_true_to_divides.
1131 rewrite > sym_times.
1132 rewrite > (p_ord_exp1 p ? m n)
1141 |apply eq_iter_p_gen
1144 elim (divides_b (x/S m) n);reflexivity
1148 |elim H1.apply lt_to_le.assumption
1152 (*distributive property for iter_p_gen*)
1153 theorem distributive_times_plus_iter_p_gen: \forall A:Type.
1154 \forall plusA: A \to A \to A. \forall basePlusA: A.
1155 \forall timesA: A \to A \to A.
1156 \forall n:nat. \forall k:A. \forall p:nat \to bool. \forall g:nat \to A.
1157 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a basePlusA) = a) \to
1158 (symmetric A timesA) \to (distributive A timesA plusA) \to
1159 (\forall a:A. (timesA a basePlusA) = basePlusA)
1163 ((timesA k (iter_p_gen n p A g basePlusA plusA)) =
1164 (iter_p_gen n p A (\lambda i:nat.(timesA k (g i))) basePlusA plusA)).
1169 | cut( (p n1) = true \lor (p n1) = false)
1171 [ rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H7).
1172 rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H7) in \vdash (? ? ? %).
1174 rewrite > (H3 k (g n1)).
1177 | rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H7).
1178 rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H7) in \vdash (? ? ? %).
1191 (* old version - proved without theorem iter_p_gen_knm
1192 theorem iter_p_gen_2_eq:
1195 \forall plusA: A \to A \to A.
1196 (symmetric A plusA) \to
1197 (associative A plusA) \to
1198 (\forall a:A.(plusA a baseA) = a)\to
1199 \forall g: nat \to nat \to A.
1200 \forall h11,h12,h21,h22: nat \to nat \to nat.
1201 \forall n1,m1,n2,m2.
1202 \forall p11,p21:nat \to bool.
1203 \forall p12,p22:nat \to nat \to bool.
1204 (\forall i,j. i < n2 \to j < m2 \to p21 i = true \to p22 i j = true \to
1205 p11 (h11 i j) = true \land p12 (h11 i j) (h12 i j) = true
1206 \land h21 (h11 i j) (h12 i j) = i \land h22 (h11 i j) (h12 i j) = j
1207 \land h11 i j < n1 \land h12 i j < m1) \to
1208 (\forall i,j. i < n1 \to j < m1 \to p11 i = true \to p12 i j = true \to
1209 p21 (h21 i j) = true \land p22 (h21 i j) (h22 i j) = true
1210 \land h11 (h21 i j) (h22 i j) = i \land h12 (h21 i j) (h22 i j) = j
1211 \land (h21 i j) < n2 \land (h22 i j) < m2) \to
1213 (\lambda x:nat .iter_p_gen m1 (p12 x) A (\lambda y. g x y) baseA plusA)
1216 (\lambda x:nat .iter_p_gen m2 (p22 x) A (\lambda y. g (h11 x y) (h12 x y)) baseA plusA )
1219 rewrite < (iter_p_gen2' ? ? ? ? ? ? ? ? H H1 H2).
1220 rewrite < (iter_p_gen2' ? ? ? ? ? ? ? ? H H1 H2).
1222 letin h := (\lambda x.(h11 (x/m2) (x\mod m2))*m1 + (h12 (x/m2) (x\mod m2))).
1223 letin h1 := (\lambda x.(h21 (x/m1) (x\mod m1))*m2 + (h22 (x/m1) (x\mod m1))).
1225 (iter_p_gen (n2*m2) (\lambda x:nat.p21 (x/m2)\land p22 (x/m2) (x\mod m2)) A
1226 (\lambda x:nat.g ((h x)/m1) ((h x)\mod m1)) baseA plusA ))
1228 apply eq_iter_p_gen1
1233 [cut (x \mod m2 < m2)
1234 [elim (and_true ? ? H6).
1235 elim (H3 ? ? Hcut1 Hcut2 H7 H8).
1242 apply div_plus_times.
1245 apply mod_plus_times.
1251 |apply (lt_times_n_to_lt m2)
1253 |apply (le_to_lt_to_lt ? x)
1254 [apply (eq_plus_to_le ? ? (x \mod m2)).
1261 |apply not_le_to_lt.unfold.intro.
1262 generalize in match H5.
1263 apply (le_n_O_elim ? H7).
1264 rewrite < times_n_O.
1269 |apply (eq_iter_p_gen_gh ? ? ? H H1 H2 ? h h1);intros
1272 [cut (i \mod m2 < m2)
1273 [elim (and_true ? ? H6).
1274 elim (H3 ? ? Hcut1 Hcut2 H7 H8).
1279 cut ((h11 (i/m2) (i\mod m2)*m1+h12 (i/m2) (i\mod m2))/m1 =
1280 h11 (i/m2) (i\mod m2))
1281 [cut ((h11 (i/m2) (i\mod m2)*m1+h12 (i/m2) (i\mod m2))\mod m1 =
1282 h12 (i/m2) (i\mod m2))
1288 |apply mod_plus_times.
1291 |apply div_plus_times.
1297 |apply (lt_times_n_to_lt m2)
1299 |apply (le_to_lt_to_lt ? i)
1300 [apply (eq_plus_to_le ? ? (i \mod m2)).
1307 |apply not_le_to_lt.unfold.intro.
1308 generalize in match H5.
1309 apply (le_n_O_elim ? H7).
1310 rewrite < times_n_O.
1316 [cut (i \mod m2 < m2)
1317 [elim (and_true ? ? H6).
1318 elim (H3 ? ? Hcut1 Hcut2 H7 H8).
1323 cut ((h11 (i/m2) (i\mod m2)*m1+h12 (i/m2) (i\mod m2))/m1 =
1324 h11 (i/m2) (i\mod m2))
1325 [cut ((h11 (i/m2) (i\mod m2)*m1+h12 (i/m2) (i\mod m2))\mod m1 =
1326 h12 (i/m2) (i\mod m2))
1334 |apply mod_plus_times.
1337 |apply div_plus_times.
1343 |apply (lt_times_n_to_lt m2)
1345 |apply (le_to_lt_to_lt ? i)
1346 [apply (eq_plus_to_le ? ? (i \mod m2)).
1353 |apply not_le_to_lt.unfold.intro.
1354 generalize in match H5.
1355 apply (le_n_O_elim ? H7).
1356 rewrite < times_n_O.
1362 [cut (i \mod m2 < m2)
1363 [elim (and_true ? ? H6).
1364 elim (H3 ? ? Hcut1 Hcut2 H7 H8).
1369 apply lt_times_plus_times
1370 [assumption|assumption]
1374 |apply (lt_times_n_to_lt m2)
1376 |apply (le_to_lt_to_lt ? i)
1377 [apply (eq_plus_to_le ? ? (i \mod m2)).
1384 |apply not_le_to_lt.unfold.intro.
1385 generalize in match H5.
1386 apply (le_n_O_elim ? H7).
1387 rewrite < times_n_O.
1393 [cut (j \mod m1 < m1)
1394 [elim (and_true ? ? H6).
1395 elim (H4 ? ? Hcut1 Hcut2 H7 H8).
1400 cut ((h21 (j/m1) (j\mod m1)*m2+h22 (j/m1) (j\mod m1))/m2 =
1401 h21 (j/m1) (j\mod m1))
1402 [cut ((h21 (j/m1) (j\mod m1)*m2+h22 (j/m1) (j\mod m1))\mod m2 =
1403 h22 (j/m1) (j\mod m1))
1409 |apply mod_plus_times.
1412 |apply div_plus_times.
1418 |apply (lt_times_n_to_lt m1)
1420 |apply (le_to_lt_to_lt ? j)
1421 [apply (eq_plus_to_le ? ? (j \mod m1)).
1428 |apply not_le_to_lt.unfold.intro.
1429 generalize in match H5.
1430 apply (le_n_O_elim ? H7).
1431 rewrite < times_n_O.
1437 [cut (j \mod m1 < m1)
1438 [elim (and_true ? ? H6).
1439 elim (H4 ? ? Hcut1 Hcut2 H7 H8).
1444 cut ((h21 (j/m1) (j\mod m1)*m2+h22 (j/m1) (j\mod m1))/m2 =
1445 h21 (j/m1) (j\mod m1))
1446 [cut ((h21 (j/m1) (j\mod m1)*m2+h22 (j/m1) (j\mod m1))\mod m2 =
1447 h22 (j/m1) (j\mod m1))
1455 |apply mod_plus_times.
1458 |apply div_plus_times.
1464 |apply (lt_times_n_to_lt m1)
1466 |apply (le_to_lt_to_lt ? j)
1467 [apply (eq_plus_to_le ? ? (j \mod m1)).
1474 |apply not_le_to_lt.unfold.intro.
1475 generalize in match H5.
1476 apply (le_n_O_elim ? H7).
1477 rewrite < times_n_O.
1483 [cut (j \mod m1 < m1)
1484 [elim (and_true ? ? H6).
1485 elim (H4 ? ? Hcut1 Hcut2 H7 H8).
1490 apply (lt_times_plus_times ? ? ? m2)
1491 [assumption|assumption]
1495 |apply (lt_times_n_to_lt m1)
1497 |apply (le_to_lt_to_lt ? j)
1498 [apply (eq_plus_to_le ? ? (j \mod m1)).
1505 |apply not_le_to_lt.unfold.intro.
1506 generalize in match H5.
1507 apply (le_n_O_elim ? H7).
1508 rewrite < times_n_O.
1517 theorem iter_p_gen_2_eq:
1520 \forall plusA: A \to A \to A.
1521 (symmetric A plusA) \to
1522 (associative A plusA) \to
1523 (\forall a:A.(plusA a baseA) = a)\to
1524 \forall g: nat \to nat \to A.
1525 \forall h11,h12,h21,h22: nat \to nat \to nat.
1526 \forall n1,m1,n2,m2.
1527 \forall p11,p21:nat \to bool.
1528 \forall p12,p22:nat \to nat \to bool.
1529 (\forall i,j. i < n2 \to j < m2 \to p21 i = true \to p22 i j = true \to
1530 p11 (h11 i j) = true \land p12 (h11 i j) (h12 i j) = true
1531 \land h21 (h11 i j) (h12 i j) = i \land h22 (h11 i j) (h12 i j) = j
1532 \land h11 i j < n1 \land h12 i j < m1) \to
1533 (\forall i,j. i < n1 \to j < m1 \to p11 i = true \to p12 i j = true \to
1534 p21 (h21 i j) = true \land p22 (h21 i j) (h22 i j) = true
1535 \land h11 (h21 i j) (h22 i j) = i \land h12 (h21 i j) (h22 i j) = j
1536 \land (h21 i j) < n2 \land (h22 i j) < m2) \to
1538 (\lambda x:nat .iter_p_gen m1 (p12 x) A (\lambda y. g x y) baseA plusA)
1541 (\lambda x:nat .iter_p_gen m2 (p22 x) A (\lambda y. g (h11 x y) (h12 x y)) baseA plusA )
1545 rewrite < (iter_p_gen2' ? ? ? ? ? ? ? ? H H1 H2).
1546 letin ha:= (\lambda x,y.(((h11 x y)*m1) + (h12 x y))).
1547 letin ha12:= (\lambda x.(h21 (x/m1) (x \mod m1))).
1548 letin ha22:= (\lambda x.(h22 (x/m1) (x \mod m1))).
1551 (iter_p_gen n2 p21 A (\lambda x:nat. iter_p_gen m2 (p22 x) A
1552 (\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))
1554 apply (iter_p_gen_knm A baseA plusA H H1 H2 (\lambda e. (g (e/m1) (e \mod m1))) ha ha12 ha22);intros
1555 [ elim (and_true ? ? H6).
1558 [ cut((x \mod m1) < m1)
1559 [ elim (H4 ? ? Hcut1 Hcut2 H7 H8).
1587 | apply (lt_times_n_to_lt m1)
1589 | apply (le_to_lt_to_lt ? x)
1590 [ apply (eq_plus_to_le ? ? (x \mod m1)).
1597 | apply not_le_to_lt.unfold.intro.
1598 generalize in match H5.
1599 apply (le_n_O_elim ? H9).
1600 rewrite < times_n_O.
1604 | elim (H3 ? ? H5 H6 H7 H8).
1609 cut(((h11 i j)*m1 + (h12 i j))/m1 = (h11 i j))
1610 [ cut(((h11 i j)*m1 + (h12 i j)) \mod m1 = (h12 i j))
1614 [ apply true_to_true_to_andb_true
1635 [ apply (lt_to_le_to_lt ? ((h11 i j)*m1 + m1) )
1639 | rewrite > sym_plus.
1640 rewrite > (sym_times (h11 i j) m1).
1641 rewrite > times_n_Sm.
1642 rewrite > sym_times.
1646 | apply not_le_to_lt.unfold.intro.
1647 generalize in match H12.
1648 apply (le_n_O_elim ? H11).
1652 | apply not_le_to_lt.unfold.intro.
1653 generalize in match H10.
1654 apply (le_n_O_elim ? H11).
1659 | rewrite > (mod_plus_times m1 (h11 i j) (h12 i j)).
1663 | rewrite > (div_plus_times m1 (h11 i j) (h12 i j)).
1668 | apply (eq_iter_p_gen1)
1669 [ intros. reflexivity
1671 apply (eq_iter_p_gen1)
1672 [ intros. reflexivity
1674 rewrite > (div_plus_times)
1675 [ rewrite > (mod_plus_times)
1677 | elim (H3 x x1 H5 H7 H6 H8).
1680 | elim (H3 x x1 H5 H7 H6 H8).
1688 theorem iter_p_gen_iter_p_gen:
1691 \forall plusA: A \to A \to A.
1692 (symmetric A plusA) \to
1693 (associative A plusA) \to
1694 (\forall a:A.(plusA a baseA) = a)\to
1695 \forall g: nat \to nat \to A.
1697 \forall p11,p21:nat \to bool.
1698 \forall p12,p22:nat \to nat \to bool.
1699 (\forall x,y. x < n \to y < m \to
1700 (p11 x \land p12 x y) = (p21 y \land p22 y x)) \to
1702 (\lambda x:nat.iter_p_gen m (p12 x) A (\lambda y. g x y) baseA plusA)
1705 (\lambda y:nat.iter_p_gen n (p22 y) A (\lambda x. g x y) baseA plusA )
1708 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)
1709 n m m n p11 p21 p12 p22)
1715 [apply (andb_true_true ? (p12 j i)).
1717 [rewrite > H6.rewrite > H7.reflexivity
1721 |apply (andb_true_true_r (p11 j)).
1723 [rewrite > H6.rewrite > H7.reflexivity
1741 [apply (andb_true_true ? (p22 j i)).
1743 [rewrite > H6.rewrite > H7.reflexivity
1747 |apply (andb_true_true_r (p21 j)).
1749 [rewrite > H6.rewrite > H7.reflexivity