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)).
161 [ rewrite < (plus_n_O n).
163 rewrite > H in \vdash (? ? ? %).
166 | apply (bool_elim ? (p (n1+n)))
168 rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
169 rewrite > (true_to_iter_p_gen_Sn n1 (\lambda x.p (x+n)) ? ? ? ? H4).
170 rewrite > (H2 (g (n1 + n)) ? ?).
174 rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
175 rewrite > (false_to_iter_p_gen_Sn n1 (\lambda x.p (x+n)) ? ? ? ? H4).
181 theorem false_to_eq_iter_p_gen: \forall A:Type. \forall n,m:nat.\forall p:nat \to bool.
182 \forall g: nat \to A. \forall baseA:A. \forall plusA: A \to A \to A.
183 n \le m \to (\forall i:nat. n \le i \to i < m \to p i = false)
184 \to iter_p_gen m p A g baseA plusA = iter_p_gen n p A g baseA plusA.
204 (* a therem slightly more general than the previous one *)
205 theorem or_false_eq_baseA_to_eq_iter_p_gen: \forall A:Type. \forall n,m:nat.\forall p:nat \to bool.
206 \forall g: nat \to A. \forall baseA:A. \forall plusA: A \to A \to A.
207 (\forall a. plusA baseA a = a) \to
208 n \le m \to (\forall i:nat. n \le i \to i < m \to p i = false \lor g i = baseA)
209 \to iter_p_gen m p A g baseA plusA = iter_p_gen n p A g baseA plusA.
213 |apply (bool_elim ? (p n1));intro
216 apply not_eq_true_false.
220 |rewrite > true_to_iter_p_gen_Sn
226 |apply le_S.assumption
233 |rewrite > false_to_iter_p_gen_Sn
237 |apply le_S.assumption
245 theorem iter_p_gen2 :
247 \forall p1,p2:nat \to bool.
249 \forall g: nat \to nat \to A.
251 \forall plusA: A \to A \to A.
252 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a)
255 (\lambda x.andb (p1 (div x m)) (p2 (mod x m)))
257 (\lambda x.g (div x m) (mod x m))
261 (\lambda x.iter_p_gen m p2 A (g x) baseA plusA)
267 | apply (bool_elim ? (p1 n1))
269 rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
270 simplify in \vdash (? ? (? % ? ? ? ? ?) ?).
271 rewrite > iter_p_gen_plusA
274 [ apply eq_iter_p_gen
277 rewrite > (div_plus_times ? ? ? H5).
278 rewrite > (mod_plus_times ? ? ? H5).
284 rewrite > (div_plus_times ? ? ? H5).
285 rewrite > (mod_plus_times ? ? ? H5).
287 simplify.reflexivity.
296 rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
297 simplify in \vdash (? ? (? % ? ? ? ? ?) ?).
298 rewrite > iter_p_gen_plusA
300 apply (trans_eq ? ? (plusA baseA
301 (iter_p_gen n1 p1 A (\lambda x:nat.iter_p_gen m p2 A (g x) baseA plusA) baseA plusA )))
303 [ rewrite > (eq_iter_p_gen ? (\lambda x.false) A ? (\lambda x:nat.g ((x+n1*m)/m) ((x+n1*m)\mod m)))
304 [ apply iter_p_gen_false
307 rewrite > (div_plus_times ? ? ? H5).
308 rewrite > (mod_plus_times ? ? ? H5).
311 | intros.reflexivity.
327 theorem iter_p_gen2':
329 \forall p1: nat \to bool.
330 \forall p2: nat \to nat \to bool.
332 \forall g: nat \to nat \to A.
334 \forall plusA: A \to A \to A.
335 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a)
338 (\lambda x.andb (p1 (div x m)) (p2 (div x m)(mod x m)))
340 (\lambda x.g (div x m) (mod x m))
344 (\lambda x.iter_p_gen m (p2 x) A (g x) baseA plusA)
350 | apply (bool_elim ? (p1 n1))
352 rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
353 simplify in \vdash (? ? (? % ? ? ? ? ?) ?).
354 rewrite > iter_p_gen_plusA
357 [ apply eq_iter_p_gen
360 rewrite > (div_plus_times ? ? ? H5).
361 rewrite > (mod_plus_times ? ? ? H5).
367 rewrite > (div_plus_times ? ? ? H5).
368 rewrite > (mod_plus_times ? ? ? H5).
370 simplify.reflexivity.
379 rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H4).
380 simplify in \vdash (? ? (? % ? ? ? ? ?) ?).
381 rewrite > iter_p_gen_plusA
383 apply (trans_eq ? ? (plusA baseA
384 (iter_p_gen n1 p1 A (\lambda x:nat.iter_p_gen m (p2 x) A (g x) baseA plusA) baseA plusA )))
386 [ rewrite > (eq_iter_p_gen ? (\lambda x.false) A ? (\lambda x:nat.g ((x+n1*m)/m) ((x+n1*m)\mod m)))
387 [ apply iter_p_gen_false
390 rewrite > (div_plus_times ? ? ? H5).
391 rewrite > (mod_plus_times ? ? ? H5).
394 | intros.reflexivity.
412 \forall g: nat \to A.
414 \forall plusA: A \to A \to A.
416 \forall p:nat \to bool.
417 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a)
420 i < n \to p i = true \to
421 (iter_p_gen n p A g baseA plusA) =
422 (plusA (g i) (iter_p_gen n (\lambda x:nat. andb (p x) (notb (eqb x i))) A g baseA plusA)).
426 apply (not_le_Sn_O i).
428 | apply (bool_elim ? (p n1));intro
429 [ elim (le_to_or_lt_eq i n1)
430 [ rewrite > true_to_iter_p_gen_Sn
431 [ rewrite > true_to_iter_p_gen_Sn
432 [ rewrite < (H2 (g i) ? ?).
433 rewrite > (H1 (g i) (g n1)).
434 rewrite > (H2 (g n1) ? ?).
445 | rewrite > H6.simplify.
446 change with (notb (eqb n1 i) = notb false).
448 apply not_eq_to_eqb_false.
450 apply (lt_to_not_eq ? ? H7).
451 apply sym_eq.assumption
455 | rewrite > true_to_iter_p_gen_Sn
459 | rewrite > false_to_iter_p_gen_Sn
460 [ apply eq_iter_p_gen
464 change with (notb false = notb (eqb x n1)).
467 apply not_eq_to_eqb_false.
468 apply (lt_to_not_eq ? ? H8)
475 rewrite > (eq_to_eqb_true ? ? (refl_eq ? n1)).
481 | apply le_S_S_to_le.
484 | rewrite > false_to_iter_p_gen_Sn
485 [ elim (le_to_or_lt_eq i n1)
486 [ rewrite > false_to_iter_p_gen_Sn
494 | rewrite > H6.reflexivity
497 apply not_eq_true_false.
501 | apply le_S_S_to_le.
510 (* invariance under permutation; single sum *)
511 theorem eq_iter_p_gen_gh:
514 \forall plusA: A \to A \to A.
515 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a) \to
516 \forall g: nat \to A.
517 \forall h,h1: nat \to nat.
519 \forall p1,p2:nat \to bool.
520 (\forall i. i < n \to p1 i = true \to p2 (h i) = true) \to
521 (\forall i. i < n \to p1 i = true \to h1 (h i) = i) \to
522 (\forall i. i < n \to p1 i = true \to h i < n1) \to
523 (\forall j. j < n1 \to p2 j = true \to p1 (h1 j) = true) \to
524 (\forall j. j < n1 \to p2 j = true \to h (h1 j) = j) \to
525 (\forall j. j < n1 \to p2 j = true \to h1 j < n) \to
527 iter_p_gen n p1 A (\lambda x.g(h x)) baseA plusA =
528 iter_p_gen n1 p2 A g baseA plusA.
531 [ generalize in match H8.
534 | apply (bool_elim ? (p2 n2));intro
536 apply (not_le_Sn_O (h1 n2)).
541 | rewrite > false_to_iter_p_gen_Sn
553 | apply (bool_elim ? (p1 n1));intro
554 [ rewrite > true_to_iter_p_gen_Sn
555 [ rewrite > (iter_p_gen_gi A g baseA plusA n2 (h n1))
562 change with ((\not eqb (h i) (h n1))= \not false).
564 apply not_eq_to_eqb_false.
567 apply (lt_to_not_eq ? ? H11).
579 | apply le_S.assumption
590 [ apply le_S.assumption
596 | generalize in match H12.
605 | generalize in match H12.
612 elim (le_to_or_lt_eq (h1 j) n1)
614 | generalize in match H12.
624 | apply eqb_false_to_not_eq.
625 generalize in match H14.
627 [ apply sym_eq.assumption
633 apply not_eq_true_false.
634 apply sym_eq.assumption
636 | apply le_S_S_to_le.
639 | generalize in match H12.
662 | rewrite > false_to_iter_p_gen_Sn
665 apply H4[apply le_S.assumption|assumption]
667 apply H5[apply le_S.assumption|assumption]
669 apply H6[apply le_S.assumption|assumption]
671 apply H7[assumption|assumption]
673 apply H8[assumption|assumption]
675 elim (le_to_or_lt_eq (h1 j) n1)
677 | absurd (j = (h n1))
685 apply not_eq_true_false.
694 | apply le_S_S_to_le.
707 theorem eq_iter_p_gen_pred:
710 \forall plusA: A \to A \to A.
713 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a) \to
714 iter_p_gen (S n) (\lambda i.p (pred i)) A (\lambda i.g(pred i)) baseA plusA =
715 plusA (iter_p_gen n p A g baseA plusA) (g O).
718 [rewrite > true_to_iter_p_gen_Sn
722 |apply (bool_elim ? (p n1));intro
723 [rewrite > true_to_iter_p_gen_Sn
724 [rewrite > true_to_iter_p_gen_Sn in ⊢ (? ? ? %)
725 [rewrite > H2 in ⊢ (? ? ? %).
726 apply eq_f.assumption
731 |rewrite > false_to_iter_p_gen_Sn
732 [rewrite > false_to_iter_p_gen_Sn in ⊢ (? ? ? %);assumption
739 definition p_ord_times \def
742 [pair q r \Rightarrow r*m+q].
744 theorem eq_p_ord_times: \forall p,m,x.
745 p_ord_times p m x = (ord_rem x p)*m+(ord x p).
746 intros.unfold p_ord_times. unfold ord_rem.
752 theorem div_p_ord_times:
753 \forall p,m,x. ord x p < m \to p_ord_times p m x / m = ord_rem x p.
754 intros.rewrite > eq_p_ord_times.
755 apply div_plus_times.
759 theorem mod_p_ord_times:
760 \forall p,m,x. ord x p < m \to p_ord_times p m x \mod m = ord x p.
761 intros.rewrite > eq_p_ord_times.
762 apply mod_plus_times.
766 lemma lt_times_to_lt_O: \forall i,n,m:nat. i < n*m \to O < m.
768 elim (le_to_or_lt_eq O ? (le_O_n m))
772 rewrite < times_n_O in H.
773 apply (not_le_Sn_O ? H)
777 theorem iter_p_gen_knm:
780 \forall plusA: A \to A \to A.
781 (symmetric A plusA) \to
782 (associative A plusA) \to
783 (\forall a:A.(plusA a baseA) = a)\to
784 \forall g: nat \to A.
785 \forall h2:nat \to nat \to nat.
786 \forall h11,h12:nat \to nat.
788 \forall p1,p21:nat \to bool.
789 \forall p22:nat \to nat \to bool.
790 (\forall x. x < k \to p1 x = true \to
791 p21 (h11 x) = true \land p22 (h11 x) (h12 x) = true
792 \land h2 (h11 x) (h12 x) = x
793 \land (h11 x) < n \land (h12 x) < m) \to
794 (\forall i,j. i < n \to j < m \to p21 i = true \to p22 i j = true \to
795 p1 (h2 i j) = true \land
796 h11 (h2 i j) = i \land h12 (h2 i j) = j
797 \land h2 i j < k) \to
798 iter_p_gen k p1 A g baseA plusA =
799 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.
801 rewrite < (iter_p_gen2' n m p21 p22 ? ? ? ? H H1 H2).
803 apply (eq_iter_p_gen_gh A baseA plusA H H1 H2 g ? (\lambda x.(h11 x)*m+(h12 x)))
805 elim (H4 (i/m) (i \mod m));clear H4
809 |apply (lt_times_to_lt_div ? ? ? H5)
811 apply (lt_times_to_lt_O ? ? ? H5)
812 |apply (andb_true_true ? ? H6)
813 |apply (andb_true_true_r ? ? H6)
816 elim (H4 (i/m) (i \mod m));clear H4
823 apply (lt_times_to_lt_O ? ? ? H5)
824 |apply (lt_times_to_lt_div ? ? ? H5)
826 apply (lt_times_to_lt_O ? ? ? H5)
827 |apply (andb_true_true ? ? H6)
828 |apply (andb_true_true_r ? ? H6)
831 elim (H4 (i/m) (i \mod m));clear H4
835 |apply (lt_times_to_lt_div ? ? ? H5)
837 apply (lt_times_to_lt_O ? ? ? H5)
838 |apply (andb_true_true ? ? H6)
839 |apply (andb_true_true_r ? ? H6)
846 rewrite > div_plus_times
847 [rewrite > mod_plus_times
860 rewrite > div_plus_times
861 [rewrite > mod_plus_times
872 apply (lt_to_le_to_lt ? ((h11 j)*m+m))
873 [apply monotonic_lt_plus_r.
876 change with ((S (h11 j)*m) \le n*m).
877 apply monotonic_le_times_l.
883 theorem iter_p_gen_divides:
886 \forall plusA: A \to A \to A.
887 \forall n,m,p:nat.O < n \to prime p \to Not (divides p n) \to
888 \forall g: nat \to A.
889 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a)
893 iter_p_gen (S (n*(exp p m))) (\lambda x.divides_b x (n*(exp p m))) A g baseA plusA =
894 iter_p_gen (S n) (\lambda x.divides_b x n) A
895 (\lambda x.iter_p_gen (S m) (\lambda y.true) A (\lambda y.g (x*(exp p y))) baseA plusA) baseA plusA.
898 [rewrite < (iter_p_gen2 ? ? ? ? ? ? ? ? H3 H4 H5).
900 (iter_p_gen (S n*S m) (\lambda x:nat.divides_b (x/S m) n) A
901 (\lambda x:nat.g (x/S m*(p)\sup(x\mod S m))) baseA plusA) )
903 apply (eq_iter_p_gen_gh ? ? ? ? ? ? g ? (p_ord_times p (S m)))
908 lapply (divides_b_true_to_lt_O ? ? H H7).
909 apply divides_to_divides_b_true
910 [rewrite > (times_n_O O).
913 |apply lt_O_exp.assumption
916 [apply divides_b_true_to_divides.assumption
917 |apply (witness ? ? (p \sup (m-i \mod (S m)))).
918 rewrite < exp_plus_times.
921 apply plus_minus_m_m.
926 lapply (divides_b_true_to_lt_O ? ? H H7).
928 rewrite > (p_ord_exp1 p ? (i \mod (S m)) (i/S m))
929 [change with ((i/S m)*S m+i \mod S m=i).
936 apply (trans_divides ? (i/ S m))
938 apply divides_b_true_to_divides;assumption]
945 change with ((i/S m) < S n).
946 apply (lt_times_to_lt_l m).
947 apply (le_to_lt_to_lt ? i)
948 [autobatch|assumption]
958 [rewrite > div_p_ord_times
959 [apply divides_to_divides_b_true
962 |apply (divides_b_true_to_lt_O ? ? ? H7).
963 rewrite > (times_n_O O).
965 [assumption|apply lt_O_exp.assumption]
967 |cut (n = ord_rem (n*(exp p m)) p)
969 apply divides_to_divides_ord_rem
970 [apply (divides_b_true_to_lt_O ? ? ? H7).
971 rewrite > (times_n_O O).
973 [assumption|apply lt_O_exp.assumption]
974 |rewrite > (times_n_O O).
976 [assumption|apply lt_O_exp.assumption]
978 |apply divides_b_true_to_divides.
982 rewrite > (p_ord_exp1 p ? m n)
992 |cut (m = ord (n*(exp p m)) p)
995 apply divides_to_le_ord
996 [apply (divides_b_true_to_lt_O ? ? ? H7).
997 rewrite > (times_n_O O).
999 [assumption|apply lt_O_exp.assumption]
1000 |rewrite > (times_n_O O).
1002 [assumption|apply lt_O_exp.assumption]
1004 |apply divides_b_true_to_divides.
1008 rewrite > (p_ord_exp1 p ? m n)
1018 [rewrite > div_p_ord_times
1019 [rewrite > mod_p_ord_times
1020 [rewrite > sym_times.
1024 |apply (divides_b_true_to_lt_O ? ? ? H7).
1025 rewrite > (times_n_O O).
1027 [assumption|apply lt_O_exp.assumption]
1029 |cut (m = ord (n*(exp p m)) p)
1032 apply divides_to_le_ord
1033 [apply (divides_b_true_to_lt_O ? ? ? H7).
1034 rewrite > (times_n_O O).
1036 [assumption|apply lt_O_exp.assumption]
1037 |rewrite > (times_n_O O).
1039 [assumption|apply lt_O_exp.assumption]
1041 |apply divides_b_true_to_divides.
1045 rewrite > (p_ord_exp1 p ? m n)
1055 |cut (m = ord (n*(exp p m)) p)
1058 apply divides_to_le_ord
1059 [apply (divides_b_true_to_lt_O ? ? ? H7).
1060 rewrite > (times_n_O O).
1062 [assumption|apply lt_O_exp.assumption]
1063 |rewrite > (times_n_O O).
1065 [assumption|apply lt_O_exp.assumption]
1067 |apply divides_b_true_to_divides.
1071 rewrite > (p_ord_exp1 p ? m n)
1080 rewrite > eq_p_ord_times.
1082 apply (lt_to_le_to_lt ? (S m +ord_rem j p*S m))
1085 cut (m = ord (n*(p \sup m)) p)
1087 apply divides_to_le_ord
1088 [apply (divides_b_true_to_lt_O ? ? ? H7).
1089 rewrite > (times_n_O O).
1091 [assumption|apply lt_O_exp.assumption]
1092 |rewrite > (times_n_O O).
1094 [assumption|apply lt_O_exp.assumption]
1096 |apply divides_b_true_to_divides.
1100 rewrite > sym_times.
1101 rewrite > (p_ord_exp1 p ? m n)
1108 |change with (S (ord_rem j p)*S m \le S n*S m).
1111 cut (n = ord_rem (n*(p \sup m)) p)
1116 |rewrite > (times_n_O O).
1118 [assumption|apply lt_O_exp.assumption]
1120 |apply divides_to_divides_ord_rem
1121 [apply (divides_b_true_to_lt_O ? ? ? H7).
1122 rewrite > (times_n_O O).
1124 [assumption|apply lt_O_exp.assumption]
1125 |rewrite > (times_n_O O).
1127 [assumption|apply lt_O_exp.assumption]
1129 |apply divides_b_true_to_divides.
1134 rewrite > sym_times.
1135 rewrite > (p_ord_exp1 p ? m n)
1144 |apply eq_iter_p_gen
1147 elim (divides_b (x/S m) n);reflexivity
1151 |elim H1.apply lt_to_le.assumption
1155 (*distributive property for iter_p_gen*)
1156 theorem distributive_times_plus_iter_p_gen: \forall A:Type.
1157 \forall plusA: A \to A \to A. \forall basePlusA: A.
1158 \forall timesA: A \to A \to A.
1159 \forall n:nat. \forall k:A. \forall p:nat \to bool. \forall g:nat \to A.
1160 (symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a basePlusA) = a) \to
1161 (symmetric A timesA) \to (distributive A timesA plusA) \to
1162 (\forall a:A. (timesA a basePlusA) = basePlusA)
1166 ((timesA k (iter_p_gen n p A g basePlusA plusA)) =
1167 (iter_p_gen n p A (\lambda i:nat.(timesA k (g i))) basePlusA plusA)).
1172 | cut( (p n1) = true \lor (p n1) = false)
1174 [ rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H7).
1175 rewrite > (true_to_iter_p_gen_Sn ? ? ? ? ? ? H7) in \vdash (? ? ? %).
1177 rewrite > (H3 k (g n1)).
1180 | rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H7).
1181 rewrite > (false_to_iter_p_gen_Sn ? ? ? ? ? ? H7) in \vdash (? ? ? %).
1194 (* old version - proved without theorem iter_p_gen_knm
1195 theorem iter_p_gen_2_eq:
1198 \forall plusA: A \to A \to A.
1199 (symmetric A plusA) \to
1200 (associative A plusA) \to
1201 (\forall a:A.(plusA a baseA) = a)\to
1202 \forall g: nat \to nat \to A.
1203 \forall h11,h12,h21,h22: nat \to nat \to nat.
1204 \forall n1,m1,n2,m2.
1205 \forall p11,p21:nat \to bool.
1206 \forall p12,p22:nat \to nat \to bool.
1207 (\forall i,j. i < n2 \to j < m2 \to p21 i = true \to p22 i j = true \to
1208 p11 (h11 i j) = true \land p12 (h11 i j) (h12 i j) = true
1209 \land h21 (h11 i j) (h12 i j) = i \land h22 (h11 i j) (h12 i j) = j
1210 \land h11 i j < n1 \land h12 i j < m1) \to
1211 (\forall i,j. i < n1 \to j < m1 \to p11 i = true \to p12 i j = true \to
1212 p21 (h21 i j) = true \land p22 (h21 i j) (h22 i j) = true
1213 \land h11 (h21 i j) (h22 i j) = i \land h12 (h21 i j) (h22 i j) = j
1214 \land (h21 i j) < n2 \land (h22 i j) < m2) \to
1216 (\lambda x:nat .iter_p_gen m1 (p12 x) A (\lambda y. g x y) baseA plusA)
1219 (\lambda x:nat .iter_p_gen m2 (p22 x) A (\lambda y. g (h11 x y) (h12 x y)) baseA plusA )
1222 rewrite < (iter_p_gen2' ? ? ? ? ? ? ? ? H H1 H2).
1223 rewrite < (iter_p_gen2' ? ? ? ? ? ? ? ? H H1 H2).
1225 letin h := (\lambda x.(h11 (x/m2) (x\mod m2))*m1 + (h12 (x/m2) (x\mod m2))).
1226 letin h1 := (\lambda x.(h21 (x/m1) (x\mod m1))*m2 + (h22 (x/m1) (x\mod m1))).
1228 (iter_p_gen (n2*m2) (\lambda x:nat.p21 (x/m2)\land p22 (x/m2) (x\mod m2)) A
1229 (\lambda x:nat.g ((h x)/m1) ((h x)\mod m1)) baseA plusA ))
1231 apply eq_iter_p_gen1
1236 [cut (x \mod m2 < m2)
1237 [elim (and_true ? ? H6).
1238 elim (H3 ? ? Hcut1 Hcut2 H7 H8).
1245 apply div_plus_times.
1248 apply mod_plus_times.
1254 |apply (lt_times_n_to_lt m2)
1256 |apply (le_to_lt_to_lt ? x)
1257 [apply (eq_plus_to_le ? ? (x \mod m2)).
1264 |apply not_le_to_lt.unfold.intro.
1265 generalize in match H5.
1266 apply (le_n_O_elim ? H7).
1267 rewrite < times_n_O.
1272 |apply (eq_iter_p_gen_gh ? ? ? H H1 H2 ? h h1);intros
1275 [cut (i \mod m2 < m2)
1276 [elim (and_true ? ? H6).
1277 elim (H3 ? ? Hcut1 Hcut2 H7 H8).
1282 cut ((h11 (i/m2) (i\mod m2)*m1+h12 (i/m2) (i\mod m2))/m1 =
1283 h11 (i/m2) (i\mod m2))
1284 [cut ((h11 (i/m2) (i\mod m2)*m1+h12 (i/m2) (i\mod m2))\mod m1 =
1285 h12 (i/m2) (i\mod m2))
1291 |apply mod_plus_times.
1294 |apply div_plus_times.
1300 |apply (lt_times_n_to_lt m2)
1302 |apply (le_to_lt_to_lt ? i)
1303 [apply (eq_plus_to_le ? ? (i \mod m2)).
1310 |apply not_le_to_lt.unfold.intro.
1311 generalize in match H5.
1312 apply (le_n_O_elim ? H7).
1313 rewrite < times_n_O.
1319 [cut (i \mod m2 < m2)
1320 [elim (and_true ? ? H6).
1321 elim (H3 ? ? Hcut1 Hcut2 H7 H8).
1326 cut ((h11 (i/m2) (i\mod m2)*m1+h12 (i/m2) (i\mod m2))/m1 =
1327 h11 (i/m2) (i\mod m2))
1328 [cut ((h11 (i/m2) (i\mod m2)*m1+h12 (i/m2) (i\mod m2))\mod m1 =
1329 h12 (i/m2) (i\mod m2))
1337 |apply mod_plus_times.
1340 |apply div_plus_times.
1346 |apply (lt_times_n_to_lt m2)
1348 |apply (le_to_lt_to_lt ? i)
1349 [apply (eq_plus_to_le ? ? (i \mod m2)).
1356 |apply not_le_to_lt.unfold.intro.
1357 generalize in match H5.
1358 apply (le_n_O_elim ? H7).
1359 rewrite < times_n_O.
1365 [cut (i \mod m2 < m2)
1366 [elim (and_true ? ? H6).
1367 elim (H3 ? ? Hcut1 Hcut2 H7 H8).
1372 apply lt_times_plus_times
1373 [assumption|assumption]
1377 |apply (lt_times_n_to_lt m2)
1379 |apply (le_to_lt_to_lt ? i)
1380 [apply (eq_plus_to_le ? ? (i \mod m2)).
1387 |apply not_le_to_lt.unfold.intro.
1388 generalize in match H5.
1389 apply (le_n_O_elim ? H7).
1390 rewrite < times_n_O.
1396 [cut (j \mod m1 < m1)
1397 [elim (and_true ? ? H6).
1398 elim (H4 ? ? Hcut1 Hcut2 H7 H8).
1403 cut ((h21 (j/m1) (j\mod m1)*m2+h22 (j/m1) (j\mod m1))/m2 =
1404 h21 (j/m1) (j\mod m1))
1405 [cut ((h21 (j/m1) (j\mod m1)*m2+h22 (j/m1) (j\mod m1))\mod m2 =
1406 h22 (j/m1) (j\mod m1))
1412 |apply mod_plus_times.
1415 |apply div_plus_times.
1421 |apply (lt_times_n_to_lt m1)
1423 |apply (le_to_lt_to_lt ? j)
1424 [apply (eq_plus_to_le ? ? (j \mod m1)).
1431 |apply not_le_to_lt.unfold.intro.
1432 generalize in match H5.
1433 apply (le_n_O_elim ? H7).
1434 rewrite < times_n_O.
1440 [cut (j \mod m1 < m1)
1441 [elim (and_true ? ? H6).
1442 elim (H4 ? ? Hcut1 Hcut2 H7 H8).
1447 cut ((h21 (j/m1) (j\mod m1)*m2+h22 (j/m1) (j\mod m1))/m2 =
1448 h21 (j/m1) (j\mod m1))
1449 [cut ((h21 (j/m1) (j\mod m1)*m2+h22 (j/m1) (j\mod m1))\mod m2 =
1450 h22 (j/m1) (j\mod m1))
1458 |apply mod_plus_times.
1461 |apply div_plus_times.
1467 |apply (lt_times_n_to_lt m1)
1469 |apply (le_to_lt_to_lt ? j)
1470 [apply (eq_plus_to_le ? ? (j \mod m1)).
1477 |apply not_le_to_lt.unfold.intro.
1478 generalize in match H5.
1479 apply (le_n_O_elim ? H7).
1480 rewrite < times_n_O.
1486 [cut (j \mod m1 < m1)
1487 [elim (and_true ? ? H6).
1488 elim (H4 ? ? Hcut1 Hcut2 H7 H8).
1493 apply (lt_times_plus_times ? ? ? m2)
1494 [assumption|assumption]
1498 |apply (lt_times_n_to_lt m1)
1500 |apply (le_to_lt_to_lt ? j)
1501 [apply (eq_plus_to_le ? ? (j \mod m1)).
1508 |apply not_le_to_lt.unfold.intro.
1509 generalize in match H5.
1510 apply (le_n_O_elim ? H7).
1511 rewrite < times_n_O.
1520 theorem iter_p_gen_2_eq:
1523 \forall plusA: A \to A \to A.
1524 (symmetric A plusA) \to
1525 (associative A plusA) \to
1526 (\forall a:A.(plusA a baseA) = a)\to
1527 \forall g: nat \to nat \to A.
1528 \forall h11,h12,h21,h22: nat \to nat \to nat.
1529 \forall n1,m1,n2,m2.
1530 \forall p11,p21:nat \to bool.
1531 \forall p12,p22:nat \to nat \to bool.
1532 (\forall i,j. i < n2 \to j < m2 \to p21 i = true \to p22 i j = true \to
1533 p11 (h11 i j) = true \land p12 (h11 i j) (h12 i j) = true
1534 \land h21 (h11 i j) (h12 i j) = i \land h22 (h11 i j) (h12 i j) = j
1535 \land h11 i j < n1 \land h12 i j < m1) \to
1536 (\forall i,j. i < n1 \to j < m1 \to p11 i = true \to p12 i j = true \to
1537 p21 (h21 i j) = true \land p22 (h21 i j) (h22 i j) = true
1538 \land h11 (h21 i j) (h22 i j) = i \land h12 (h21 i j) (h22 i j) = j
1539 \land (h21 i j) < n2 \land (h22 i j) < m2) \to
1541 (\lambda x:nat .iter_p_gen m1 (p12 x) A (\lambda y. g x y) baseA plusA)
1544 (\lambda x:nat .iter_p_gen m2 (p22 x) A (\lambda y. g (h11 x y) (h12 x y)) baseA plusA )
1548 rewrite < (iter_p_gen2' ? ? ? ? ? ? ? ? H H1 H2).
1549 letin ha:= (\lambda x,y.(((h11 x y)*m1) + (h12 x y))).
1550 letin ha12:= (\lambda x.(h21 (x/m1) (x \mod m1))).
1551 letin ha22:= (\lambda x.(h22 (x/m1) (x \mod m1))).
1554 (iter_p_gen n2 p21 A (\lambda x:nat. iter_p_gen m2 (p22 x) A
1555 (\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))
1557 apply (iter_p_gen_knm A baseA plusA H H1 H2 (\lambda e. (g (e/m1) (e \mod m1))) ha ha12 ha22);intros
1558 [ elim (and_true ? ? H6).
1561 [ cut((x \mod m1) < m1)
1562 [ 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
1631 [ apply (lt_to_le_to_lt ? ((h11 i j)*m1 + m1) )
1632 [ apply (lt_plus_r).
1634 | rewrite > sym_plus.
1635 rewrite > (sym_times (h11 i j) m1).
1636 rewrite > times_n_Sm.
1637 rewrite > sym_times.
1641 | apply not_le_to_lt.unfold.intro.
1642 generalize in match H12.
1643 apply (le_n_O_elim ? H11).
1647 | apply not_le_to_lt.unfold.intro.
1648 generalize in match H10.
1649 apply (le_n_O_elim ? H11).
1654 | rewrite > (mod_plus_times m1 (h11 i j) (h12 i j)).
1658 | rewrite > (div_plus_times m1 (h11 i j) (h12 i j)).
1663 | apply (eq_iter_p_gen1)
1664 [ intros. reflexivity
1666 apply (eq_iter_p_gen1)
1667 [ intros. reflexivity
1669 rewrite > (div_plus_times)
1670 [ rewrite > (mod_plus_times)
1672 | elim (H3 x x1 H5 H7 H6 H8).
1675 | elim (H3 x x1 H5 H7 H6 H8).
1683 theorem iter_p_gen_iter_p_gen:
1686 \forall plusA: A \to A \to A.
1687 (symmetric A plusA) \to
1688 (associative A plusA) \to
1689 (\forall a:A.(plusA a baseA) = a)\to
1690 \forall g: nat \to nat \to A.
1692 \forall p11,p21:nat \to bool.
1693 \forall p12,p22:nat \to nat \to bool.
1694 (\forall x,y. x < n \to y < m \to
1695 (p11 x \land p12 x y) = (p21 y \land p22 y x)) \to
1697 (\lambda x:nat.iter_p_gen m (p12 x) A (\lambda y. g x y) baseA plusA)
1700 (\lambda y:nat.iter_p_gen n (p22 y) A (\lambda x. g x y) baseA plusA )
1703 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)
1704 n m m n p11 p21 p12 p22)
1710 [apply (andb_true_true ? (p12 j i)).
1712 [rewrite > H6.rewrite > H7.reflexivity
1716 |apply (andb_true_true_r (p11 j)).
1718 [rewrite > H6.rewrite > H7.reflexivity
1736 [apply (andb_true_true ? (p22 j i)).
1738 [rewrite > H6.rewrite > H7.reflexivity
1742 |apply (andb_true_true_r (p21 j)).
1744 [rewrite > H6.rewrite > H7.reflexivity