include "arithmetics/primes.ma".
include "arithmetics/bigops.ma".
+(* Sigma e Pi *)
+
+notation "∑_{ ident i < n | p } f"
+ with precedence 80
+for @{'bigop $n plus 0 (λ${ident i}. $p) (λ${ident i}. $f)}.
+
+notation "∑_{ ident i < n } f"
+ with precedence 80
+for @{'bigop $n plus 0 (λ${ident i}.true) (λ${ident i}. $f)}.
+
+notation "∑_{ ident j ∈ [a,b[ } f"
+ with precedence 80
+for @{'bigop ($b-$a) plus 0 (λ${ident j}.((λ${ident j}.true) (${ident j}+$a)))
+ (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
+
+notation "∑_{ ident j ∈ [a,b[ | p } f"
+ with precedence 80
+for @{'bigop ($b-$a) plus 0 (λ${ident j}.((λ${ident j}.$p) (${ident j}+$a)))
+ (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
+
+notation "∏_{ ident i < n | p} f"
+ with precedence 80
+for @{'bigop $n times 1 (λ${ident i}.$p) (λ${ident i}. $f)}.
+
+notation "∏_{ ident i < n } f"
+ with precedence 80
+for @{'bigop $n times 1 (λ${ident i}.true) (λ${ident i}. $f)}.
+
+notation "∏_{ ident j ∈ [a,b[ } f"
+ with precedence 80
+for @{'bigop ($b-$a) times 1 (λ${ident j}.((λ${ident j}.true) (${ident j}+$a)))
+ (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
+
+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)))}.
+
+(* instances of associative and commutative operations *)
+
+definition plusA ≝ mk_Aop nat 0 plus (λa.refl ? a) (λn.sym_eq ??? (plus_n_O n))
+ (λa,b,c.sym_eq ??? (associative_plus a b c)).
+
+definition plusAC ≝ mk_ACop nat 0 plusA commutative_plus.
+
+definition timesA ≝ mk_Aop nat 1 times
+ (λa.sym_eq ??? (plus_n_O a)) (λn.sym_eq ??? (times_n_1 n))
+ (λa,b,c.sym_eq ??? (associative_times a b c)).
+
+definition timesAC ≝ mk_ACop nat 1 timesA commutative_times.
+
+definition natD ≝ mk_Dop nat 0 plusAC times (λn.(sym_eq ??? (times_n_O n)))
+ distributive_times_plus.
+
+(********************************************************)
+
theorem sigma_const: ∀n:nat. ∑_{i<n} 1 = n.
#n elim n // #n1 >bigop_Strue //
qed.
|#m1 #Hind >times_pi >Hind %
]
qed.
-
-(*
-theorem true_to_pi_p_Sn: ∀n,p,g.
- p n = true \to pi_p (S n) p g = (g n)*(pi_p n p g).
-intros.
-unfold pi_p.
-apply true_to_iter_p_gen_Sn.
-assumption.
-qed.
-
-theorem false_to_pi_p_Sn:
-\forall n:nat. \forall p:nat \to bool. \forall g:nat \to nat.
-p n = false \to pi_p (S n) p g = pi_p n p g.
-intros.
-unfold pi_p.
-apply false_to_iter_p_gen_Sn.
-assumption.
-qed.
-
-theorem eq_pi_p: \forall p1,p2:nat \to bool.
-\forall g1,g2: nat \to nat.\forall n.
-(\forall x. x < n \to p1 x = p2 x) \to
-(\forall x. x < n \to g1 x = g2 x) \to
-pi_p n p1 g1 = pi_p n p2 g2.
-intros.
-unfold pi_p.
-apply eq_iter_p_gen;
-assumption.
-qed.
-
-theorem eq_pi_p1: \forall p1,p2:nat \to bool.
-\forall g1,g2: nat \to nat.\forall n.
-(\forall x. x < n \to p1 x = p2 x) \to
-(\forall x. x < n \to p1 x = true \to g1 x = g2 x) \to
-pi_p n p1 g1 = pi_p n p2 g2.
-intros.
-unfold pi_p.
-apply eq_iter_p_gen1;
-assumption.
-qed.
-
-theorem pi_p_false:
-\forall g: nat \to nat.\forall n.pi_p n (\lambda x.false) g = S O.
-intros.
-unfold pi_p.
-apply iter_p_gen_false.
-qed.
-
-theorem pi_p_times: \forall n,k:nat.\forall p:nat \to bool.
-\forall g: nat \to nat.
-pi_p (k+n) p g
-= pi_p k (\lambda x.p (x+n)) (\lambda x.g (x+n)) * pi_p n p g.
-intros.
-unfold pi_p.
-apply (iter_p_gen_plusA nat n k p g (S O) times)
-[ apply sym_times.
-| intros.
- apply sym_eq.
- apply times_n_SO
-| apply associative_times
-]
-qed.
-
-theorem false_to_eq_pi_p: \forall n,m:nat.n \le m \to
-\forall p:nat \to bool.
-\forall g: nat \to nat. (\forall i:nat. n \le i \to i < m \to
-p i = false) \to pi_p m p g = pi_p n p g.
-intros.
-unfold pi_p.
-apply (false_to_eq_iter_p_gen);
-assumption.
-qed.
-
-theorem or_false_eq_SO_to_eq_pi_p:
-\forall n,m:nat.\forall p:nat \to bool.
-\forall g: nat \to nat.
-n \le m \to (\forall i:nat. n \le i \to i < m \to p i = false \lor g i = S O)
-\to pi_p m p g = pi_p n p g.
-intros.
-unfold pi_p.
-apply or_false_eq_baseA_to_eq_iter_p_gen
- [intros.simplify.rewrite < plus_n_O.reflexivity
- |assumption
- |assumption
- ]
-qed.
-
-theorem pi_p2 :
-\forall n,m:nat.
-\forall p1,p2:nat \to bool.
-\forall g: nat \to nat \to nat.
-pi_p (n*m)
- (\lambda x.andb (p1 (div x m)) (p2 (mod x m)))
- (\lambda x.g (div x m) (mod x m)) =
-pi_p n p1
- (\lambda x.pi_p m p2 (g x)).
-intros.
-unfold pi_p.
-apply (iter_p_gen2 n m p1 p2 nat g (S O) times)
-[ apply sym_times
-| apply associative_times
-| intros.
- apply sym_eq.
- apply times_n_SO
-]
-qed.
-
-theorem pi_p2' :
-\forall n,m:nat.
-\forall p1:nat \to bool.
-\forall p2:nat \to nat \to bool.
-\forall g: nat \to nat \to nat.
-pi_p (n*m)
- (\lambda x.andb (p1 (div x m)) (p2 (div x m) (mod x m)))
- (\lambda x.g (div x m) (mod x m)) =
-pi_p n p1
- (\lambda x.pi_p m (p2 x) (g x)).
-intros.
-unfold pi_p.
-apply (iter_p_gen2' n m p1 p2 nat g (S O) times)
-[ apply sym_times
-| apply associative_times
-| intros.
- apply sym_eq.
- apply times_n_SO
-]
-qed.
-
-lemma pi_p_gi: \forall g: nat \to nat.
-\forall n,i.\forall p:nat \to bool.i < n \to p i = true \to
-pi_p n p g = g i * pi_p n (\lambda x. andb (p x) (notb (eqb x i))) g.
-intros.
-unfold pi_p.
-apply (iter_p_gen_gi)
-[ apply sym_times
-| apply associative_times
-| intros.
- apply sym_eq.
- apply times_n_SO
-| assumption
-| assumption
-]
-qed.
-
-theorem eq_pi_p_gh:
-\forall g,h,h1: nat \to nat.\forall n,n1.
-\forall p1,p2:nat \to bool.
-(\forall i. i < n \to p1 i = true \to p2 (h i) = true) \to
-(\forall i. i < n \to p1 i = true \to h1 (h i) = i) \to
-(\forall i. i < n \to p1 i = true \to h i < n1) \to
-(\forall j. j < n1 \to p2 j = true \to p1 (h1 j) = true) \to
-(\forall j. j < n1 \to p2 j = true \to h (h1 j) = j) \to
-(\forall j. j < n1 \to p2 j = true \to h1 j < n) \to
-pi_p n p1 (\lambda x.g(h x)) = pi_p n1 p2 g.
-intros.
-unfold pi_p.
-apply (eq_iter_p_gen_gh nat (S O) times ? ? ? g h h1 n n1 p1 p2)
-[ apply sym_times
-| apply associative_times
-| intros.
- apply sym_eq.
- apply times_n_SO
-| assumption
-| assumption
-| assumption
-| assumption
-| assumption
-| assumption
-]
-qed.
-
-theorem exp_sigma_p: \forall n,a,p.
-pi_p n p (\lambda x.a) = (exp a (sigma_p n p (\lambda x.S O))).
-intros.
-elim n
- [reflexivity
- |apply (bool_elim ? (p n1))
- [intro.
- rewrite > true_to_pi_p_Sn
- [rewrite > true_to_sigma_p_Sn
- [simplify.
- rewrite > H.
- reflexivity.
- |assumption
- ]
- |assumption
- ]
- |intro.
- rewrite > false_to_pi_p_Sn
- [rewrite > false_to_sigma_p_Sn
- [simplify.assumption
- |assumption
- ]
- |assumption
- ]
- ]
- ]
-qed.
-
-theorem exp_sigma_p1: \forall n,a,p,f.
-pi_p n p (\lambda x.(exp a (f x))) = (exp a (sigma_p n p f)).
-intros.
-elim n
- [reflexivity
- |apply (bool_elim ? (p n1))
- [intro.
- rewrite > true_to_pi_p_Sn
- [rewrite > true_to_sigma_p_Sn
- [simplify.
- rewrite > H.
- rewrite > exp_plus_times.
- reflexivity.
- |assumption
- ]
- |assumption
- ]
- |intro.
- rewrite > false_to_pi_p_Sn
- [rewrite > false_to_sigma_p_Sn
- [simplify.assumption
- |assumption
- ]
- |assumption
- ]
- ]
- ]
-qed.
-
-theorem times_pi_p: \forall n,p,f,g.
-pi_p n p (\lambda x.f x*g x) = pi_p n p f * pi_p n p g.
-intros.
-elim n
- [simplify.reflexivity
- |apply (bool_elim ? (p n1))
- [intro.
- rewrite > true_to_pi_p_Sn
- [rewrite > true_to_pi_p_Sn
- [rewrite > true_to_pi_p_Sn
- [rewrite > H.autobatch
- |assumption
- ]
- |assumption
- ]
- |assumption
- ]
- |intro.
- rewrite > false_to_pi_p_Sn
- [rewrite > false_to_pi_p_Sn
- [rewrite > false_to_pi_p_Sn;assumption
- |assumption
- ]
- |assumption
- ]
- ]
- ]
-qed.
-
-
-theorem exp_times_pi_p: \forall n,m,k,p,f.
-pi_p n p (\lambda x.exp k (m*(f x))) =
-exp (pi_p n p (\lambda x.exp k (f x))) m.
-intros.
-apply (trans_eq ? ? (pi_p n p (\lambda x.(exp (exp k (f x)) m))))
- [apply eq_pi_p;intros
- [reflexivity
- |apply sym_eq.rewrite > sym_times.
- apply exp_exp_times
- ]
- |apply exp_pi_p
- ]
-qed.
-
-
-theorem pi_p_knm:
-\forall g: nat \to nat.
-\forall h2:nat \to nat \to nat.
-\forall h11,h12:nat \to nat.
-\forall k,n,m.
-\forall p1,p21:nat \to bool.
-\forall p22:nat \to nat \to bool.
-(\forall x. x < k \to p1 x = true \to
-p21 (h11 x) = true ∧ p22 (h11 x) (h12 x) = true
-\land h2 (h11 x) (h12 x) = x
-\land (h11 x) < n \land (h12 x) < m) \to
-(\forall i,j. i < n \to j < m \to p21 i = true \to p22 i j = true \to
-p1 (h2 i j) = true \land
-h11 (h2 i j) = i \land h12 (h2 i j) = j
-\land h2 i j < k) →
-(*
-Pi z < k | p1 z. g z =
-Pi x < n | p21 x. Pi y < m | p22 x y.g (h2 x y).
-*)
-pi_p k p1 g =
-pi_p n p21 (\lambda x:nat.pi_p m (p22 x) (\lambda y. g (h2 x y))).
-intros.
-unfold pi_p.unfold pi_p.
-apply (iter_p_gen_knm nat (S O) times sym_times assoc_times ? ? ? h11 h12)
- [intros.apply sym_eq.apply times_n_SO.
- |assumption
- |assumption
- ]
-qed.
-
-theorem pi_p_pi_p:
-\forall g: nat \to nat \to nat.
-\forall h11,h12,h21,h22: nat \to nat \to nat.
-\forall n1,m1,n2,m2.
-\forall p11,p21:nat \to bool.
-\forall p12,p22:nat \to nat \to bool.
-(\forall i,j. i < n2 \to j < m2 \to p21 i = true \to p22 i j = true \to
-p11 (h11 i j) = true \land p12 (h11 i j) (h12 i j) = true
-\land h21 (h11 i j) (h12 i j) = i \land h22 (h11 i j) (h12 i j) = j
-\land h11 i j < n1 \land h12 i j < m1) \to
-(\forall i,j. i < n1 \to j < m1 \to p11 i = true \to p12 i j = true \to
-p21 (h21 i j) = true \land p22 (h21 i j) (h22 i j) = true
-\land h11 (h21 i j) (h22 i j) = i \land h12 (h21 i j) (h22 i j) = j
-\land (h21 i j) < n2 \land (h22 i j) < m2) \to
-pi_p n1 p11
- (\lambda x:nat .pi_p m1 (p12 x) (\lambda y. g x y)) =
-pi_p n2 p21
- (\lambda x:nat .pi_p m2 (p22 x) (\lambda y. g (h11 x y) (h12 x y))).
-intros.
-unfold pi_p.unfold pi_p.
-apply (iter_p_gen_2_eq ? ? ? sym_times assoc_times ? ? ? ? h21 h22)
- [intros.apply sym_eq.apply times_n_SO.
- |assumption
- |assumption
- ]
-qed.
-
-theorem pi_p_pi_p1:
-\forall g: nat \to nat \to nat.
-\forall n,m.
-\forall p11,p21:nat \to bool.
-\forall p12,p22:nat \to nat \to bool.
-(\forall x,y. x < n \to y < m \to
- (p11 x \land p12 x y) = (p21 y \land p22 y x)) \to
-pi_p n p11 (\lambda x:nat.pi_p m (p12 x) (\lambda y. g x y)) =
-pi_p m p21 (\lambda y:nat.pi_p n (p22 y) (\lambda x. g x y)).
-intros.
-unfold pi_p.unfold pi_p.
-apply (iter_p_gen_iter_p_gen ? ? ? sym_times assoc_times)
- [intros.apply sym_eq.apply times_n_SO.
- |assumption
- ]
-qed. *)
\ No newline at end of file
|>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)}.
-
-notation "∑_{ ident i < n } f"
- with precedence 80
-for @{'bigop $n plus 0 (λ${ident i}.true) (λ${ident i}. $f)}.
-
-notation "∑_{ ident j ∈ [a,b[ } f"
- with precedence 80
-for @{'bigop ($b-$a) plus 0 (λ${ident j}.((λ${ident j}.true) (${ident j}+$a)))
- (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
-
-notation "∑_{ ident j ∈ [a,b[ | p } f"
- with precedence 80
-for @{'bigop ($b-$a) plus 0 (λ${ident j}.((λ${ident j}.$p) (${ident j}+$a)))
- (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
-
-notation "∏_{ ident i < n | p} f"
- with precedence 80
-for @{'bigop $n times 1 (λ${ident i}.$p) (λ${ident i}. $f)}.
-
-notation "∏_{ ident i < n } f"
- with precedence 80
-for @{'bigop $n times 1 (λ${ident i}.true) (λ${ident i}. $f)}.
-
-notation "∏_{ ident j ∈ [a,b[ } f"
- with precedence 80
-for @{'bigop ($b-$a) times 1 (λ${ident j}.((λ${ident j}.true) (${ident j}+$a)))
- (λ${ident j}.((λ${ident j}.$f)(${ident j}+$a)))}.
-
-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
-\lambda p,m,x.
- match p_ord x p with
- [pair q r \Rightarrow r*m+q].
-
-theorem eq_p_ord_times: \forall p,m,x.
-p_ord_times p m x = (ord_rem x p)*m+(ord x p).
-intros.unfold p_ord_times. unfold ord_rem.
-unfold ord.
-elim (p_ord x p).
-reflexivity.
-qed.
-
-theorem div_p_ord_times:
-\forall p,m,x. ord x p < m \to p_ord_times p m x / m = ord_rem x p.
-intros.rewrite > eq_p_ord_times.
-apply div_plus_times.
-assumption.
-qed.
-
-theorem mod_p_ord_times:
-\forall p,m,x. ord x p < m \to p_ord_times p m x \mod m = ord x p.
-intros.rewrite > eq_p_ord_times.
-apply mod_plus_times.
-assumption.
-qed.
-
-lemma lt_times_to_lt_O: \forall i,n,m:nat. i < n*m \to O < m.
-intros.
-elim (le_to_or_lt_eq O ? (le_O_n m))
- [assumption
- |apply False_ind.
- rewrite < H1 in H.
- rewrite < times_n_O in H.
- apply (not_le_Sn_O ? H)
- ]
-qed.
-
-theorem iter_p_gen_knm:
-\forall A:Type.
-\forall baseA: A.
-\forall plusA: A \to A \to A.
-(symmetric A plusA) \to
-(associative A plusA) \to
-(\forall a:A.(plusA a baseA) = a)\to
-\forall g: nat \to A.
-\forall h2:nat \to nat \to nat.
-\forall h11,h12:nat \to nat.
-\forall k,n,m.
-\forall p1,p21:nat \to bool.
-\forall p22:nat \to nat \to bool.
-(\forall x. x < k \to p1 x = true \to
-p21 (h11 x) = true \land p22 (h11 x) (h12 x) = true
-\land h2 (h11 x) (h12 x) = x
-\land (h11 x) < n \land (h12 x) < m) \to
-(\forall i,j. i < n \to j < m \to p21 i = true \to p22 i j = true \to
-p1 (h2 i j) = true \land
-h11 (h2 i j) = i \land h12 (h2 i j) = j
-\land h2 i j < k) \to
-iter_p_gen k p1 A g baseA plusA =
-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.
-intros.
-rewrite < (iter_p_gen2' n m p21 p22 ? ? ? ? H H1 H2).
-apply sym_eq.
-apply (eq_iter_p_gen_gh A baseA plusA H H1 H2 g ? (\lambda x.(h11 x)*m+(h12 x)))
- [intros.
- elim (H4 (i/m) (i \mod m));clear H4
- [elim H7.clear H7.
- elim H4.clear H4.
- assumption
- |apply (lt_times_to_lt_div ? ? ? H5)
- |apply lt_mod_m_m.
- apply (lt_times_to_lt_O ? ? ? H5)
- |apply (andb_true_true ? ? H6)
- |apply (andb_true_true_r ? ? H6)
- ]
- |intros.
- elim (H4 (i/m) (i \mod m));clear H4
- [elim H7.clear H7.
- elim H4.clear H4.
- rewrite > H10.
- rewrite > H9.
- apply sym_eq.
- apply div_mod.
- apply (lt_times_to_lt_O ? ? ? H5)
- |apply (lt_times_to_lt_div ? ? ? H5)
- |apply lt_mod_m_m.
- apply (lt_times_to_lt_O ? ? ? H5)
- |apply (andb_true_true ? ? H6)
- |apply (andb_true_true_r ? ? H6)
- ]
- |intros.
- elim (H4 (i/m) (i \mod m));clear H4
- [elim H7.clear H7.
- elim H4.clear H4.
- assumption
- |apply (lt_times_to_lt_div ? ? ? H5)
- |apply lt_mod_m_m.
- apply (lt_times_to_lt_O ? ? ? H5)
- |apply (andb_true_true ? ? H6)
- |apply (andb_true_true_r ? ? H6)
- ]
- |intros.
- elim (H3 j H5 H6).
- elim H7.clear H7.
- elim H9.clear H9.
- elim H7.clear H7.
- rewrite > div_plus_times
- [rewrite > mod_plus_times
- [rewrite > H9.
- rewrite > H12.
- reflexivity.
- |assumption
- ]
- |assumption
- ]
- |intros.
- elim (H3 j H5 H6).
- elim H7.clear H7.
- elim H9.clear H9.
- elim H7.clear H7.
- rewrite > div_plus_times
- [rewrite > mod_plus_times
- [assumption
- |assumption
- ]
- |assumption
- ]
- |intros.
- elim (H3 j H5 H6).
- elim H7.clear H7.
- elim H9.clear H9.
- elim H7.clear H7.
- apply (lt_to_le_to_lt ? ((h11 j)*m+m))
- [apply monotonic_lt_plus_r.
- assumption
- |rewrite > sym_plus.
- change with ((S (h11 j)*m) \le n*m).
- apply monotonic_le_times_l.
- assumption
- ]
- ]
-qed.
-
-theorem iter_p_gen_divides:
-\forall A:Type.
-\forall baseA: A.
-\forall plusA: A \to A \to A.
-\forall n,m,p:nat.O < n \to prime p \to Not (divides p n) \to
-\forall g: nat \to A.
-(symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a)
-
-\to
-
-iter_p_gen (S (n*(exp p m))) (\lambda x.divides_b x (n*(exp p m))) A g baseA plusA =
-iter_p_gen (S n) (\lambda x.divides_b x n) A
- (\lambda x.iter_p_gen (S m) (\lambda y.true) A (\lambda y.g (x*(exp p y))) baseA plusA) baseA plusA.
-intros.
-cut (O < p)
- [rewrite < (iter_p_gen2 ? ? ? ? ? ? ? ? H3 H4 H5).
- apply (trans_eq ? ?
- (iter_p_gen (S n*S m) (\lambda x:nat.divides_b (x/S m) n) A
- (\lambda x:nat.g (x/S m*(p)\sup(x\mod S m))) baseA plusA) )
- [apply sym_eq.
- apply (eq_iter_p_gen_gh ? ? ? ? ? ? g ? (p_ord_times p (S m)))
- [ assumption
- | assumption
- | assumption
- |intros.
- lapply (divides_b_true_to_lt_O ? ? H H7).
- apply divides_to_divides_b_true
- [rewrite > (times_n_O O).
- apply lt_times
- [assumption
- |apply lt_O_exp.assumption
- ]
- |apply divides_times
- [apply divides_b_true_to_divides.assumption
- |apply (witness ? ? (p \sup (m-i \mod (S m)))).
- rewrite < exp_plus_times.
- apply eq_f.
- rewrite > sym_plus.
- apply plus_minus_m_m.
- autobatch by le_S_S_to_le, lt_mod_m_m, lt_O_S;
- ]
- ]
- |intros.
- lapply (divides_b_true_to_lt_O ? ? H H7).
- unfold p_ord_times.
- rewrite > (p_ord_exp1 p ? (i \mod (S m)) (i/S m))
- [change with ((i/S m)*S m+i \mod S m=i).
- apply sym_eq.
- apply div_mod.
- apply lt_O_S
- |assumption
- |unfold Not.intro.
- apply H2.
- apply (trans_divides ? (i/ S m))
- [assumption|
- apply divides_b_true_to_divides;assumption]
- |apply sym_times.
- ]
- |intros.
- apply le_S_S.
- apply le_times
- [apply le_S_S_to_le.
- change with ((i/S m) < S n).
- apply (lt_times_to_lt_l m).
- apply (le_to_lt_to_lt ? i);[2:assumption]
- autobatch by eq_plus_to_le, div_mod, lt_O_S.
- |apply le_exp
- [assumption
- |apply le_S_S_to_le.
- apply lt_mod_m_m.
- apply lt_O_S
- ]
- ]
- |intros.
- cut (ord j p < S m)
- [rewrite > div_p_ord_times
- [apply divides_to_divides_b_true
- [apply lt_O_ord_rem
- [elim H1.assumption
- |apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- ]
- |cut (n = ord_rem (n*(exp p m)) p)
- [rewrite > Hcut2.
- apply divides_to_divides_ord_rem
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- |unfold ord_rem.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |apply sym_times
- ]
- ]
- ]
- |assumption
- ]
- |cut (m = ord (n*(exp p m)) p)
- [apply le_S_S.
- rewrite > Hcut1.
- apply divides_to_le_ord
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- |unfold ord.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |apply sym_times
- ]
- ]
- ]
- |intros.
- cut (ord j p < S m)
- [rewrite > div_p_ord_times
- [rewrite > mod_p_ord_times
- [rewrite > sym_times.
- apply sym_eq.
- apply exp_ord
- [elim H1.assumption
- |apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- ]
- |cut (m = ord (n*(exp p m)) p)
- [apply le_S_S.
- rewrite > Hcut2.
- apply divides_to_le_ord
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- |unfold ord.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |apply sym_times
- ]
- ]
- ]
- |assumption
- ]
- |cut (m = ord (n*(exp p m)) p)
- [apply le_S_S.
- rewrite > Hcut1.
- apply divides_to_le_ord
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- |unfold ord.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |apply sym_times
- ]
- ]
- ]
- |intros.
- rewrite > eq_p_ord_times.
- rewrite > sym_plus.
- apply (lt_to_le_to_lt ? (S m +ord_rem j p*S m))
- [apply lt_plus_l.
- apply le_S_S.
- cut (m = ord (n*(p \sup m)) p)
- [rewrite > Hcut1.
- apply divides_to_le_ord
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- |unfold ord.
- rewrite > sym_times.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |reflexivity
- ]
- ]
- |change with (S (ord_rem j p)*S m \le S n*S m).
- apply le_times_l.
- apply le_S_S.
- cut (n = ord_rem (n*(p \sup m)) p)
- [rewrite > Hcut1.
- apply divides_to_le
- [apply lt_O_ord_rem
- [elim H1.assumption
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- ]
- |apply divides_to_divides_ord_rem
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- ]
- |unfold ord_rem.
- rewrite > sym_times.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |reflexivity
- ]
- ]
- ]
- ]
- |apply eq_iter_p_gen
-
- [intros.
- elim (divides_b (x/S m) n);reflexivity
- |intros.reflexivity
- ]
- ]
-|elim H1.apply lt_to_le.assumption
-]
-qed.
-
-
-
-theorem iter_p_gen_2_eq:
-\forall A:Type.
-\forall baseA: A.
-\forall plusA: A \to A \to A.
-(symmetric A plusA) \to
-(associative A plusA) \to
-(\forall a:A.(plusA a baseA) = a)\to
-\forall g: nat \to nat \to A.
-\forall h11,h12,h21,h22: nat \to nat \to nat.
-\forall n1,m1,n2,m2.
-\forall p11,p21:nat \to bool.
-\forall p12,p22:nat \to nat \to bool.
-(\forall i,j. i < n2 \to j < m2 \to p21 i = true \to p22 i j = true \to
-p11 (h11 i j) = true \land p12 (h11 i j) (h12 i j) = true
-\land h21 (h11 i j) (h12 i j) = i \land h22 (h11 i j) (h12 i j) = j
-\land h11 i j < n1 \land h12 i j < m1) \to
-(\forall i,j. i < n1 \to j < m1 \to p11 i = true \to p12 i j = true \to
-p21 (h21 i j) = true \land p22 (h21 i j) (h22 i j) = true
-\land h11 (h21 i j) (h22 i j) = i \land h12 (h21 i j) (h22 i j) = j
-\land (h21 i j) < n2 \land (h22 i j) < m2) \to
-iter_p_gen n1 p11 A
- (\lambda x:nat .iter_p_gen m1 (p12 x) A (\lambda y. g x y) baseA plusA)
- baseA plusA =
-iter_p_gen n2 p21 A
- (\lambda x:nat .iter_p_gen m2 (p22 x) A (\lambda y. g (h11 x y) (h12 x y)) baseA plusA )
- baseA plusA.
-
-intros.
-rewrite < (iter_p_gen2' ? ? ? ? ? ? ? ? H H1 H2).
-letin ha:= (\lambda x,y.(((h11 x y)*m1) + (h12 x y))).
-letin ha12:= (\lambda x.(h21 (x/m1) (x \mod m1))).
-letin ha22:= (\lambda x.(h22 (x/m1) (x \mod m1))).
-
-apply (trans_eq ? ?
-(iter_p_gen n2 p21 A (\lambda x:nat. iter_p_gen m2 (p22 x) A
- (\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))
-[
- apply (iter_p_gen_knm A baseA plusA H H1 H2 (\lambda e. (g (e/m1) (e \mod m1))) ha ha12 ha22);intros
- [ elim (and_true ? ? H6).
- cut(O \lt m1)
- [ cut(x/m1 < n1)
- [ cut((x \mod m1) < m1)
- [ elim (H4 ? ? Hcut1 Hcut2 H7 H8).
- elim H9.clear H9.
- elim H11.clear H11.
- elim H9.clear H9.
- elim H11.clear H11.
- split
- [ split
- [ split
- [ split
- [ assumption
- | assumption
- ]
- | unfold ha.
- unfold ha12.
- unfold ha22.
- rewrite > H14.
- rewrite > H13.
- apply sym_eq.
- apply div_mod.
- assumption
- ]
- | assumption
- ]
- | assumption
- ]
- | apply lt_mod_m_m.
- assumption
- ]
- | apply (lt_times_n_to_lt m1)
- [ assumption
- | apply (le_to_lt_to_lt ? x)
- [ apply (eq_plus_to_le ? ? (x \mod m1)).
- apply div_mod.
- assumption
- | assumption
- ]
- ]
- ]
- | apply not_le_to_lt.unfold.intro.
- generalize in match H5.
- apply (le_n_O_elim ? H9).
- rewrite < times_n_O.
- apply le_to_not_lt.
- apply le_O_n.
- ]
- | elim (H3 ? ? H5 H6 H7 H8).
- elim H9.clear H9.
- elim H11.clear H11.
- elim H9.clear H9.
- elim H11.clear H11.
- cut(((h11 i j)*m1 + (h12 i j))/m1 = (h11 i j))
- [ cut(((h11 i j)*m1 + (h12 i j)) \mod m1 = (h12 i j))
- [ split
- [ split
- [ split
- [ apply true_to_true_to_andb_true
- [ rewrite > Hcut.
- assumption
- | rewrite > Hcut1.
- rewrite > Hcut.
- assumption
- ]
- | unfold ha.
- unfold ha12.
- rewrite > Hcut1.
- rewrite > Hcut.
- assumption
- ]
- | unfold ha.
- unfold ha22.
- rewrite > Hcut1.
- rewrite > Hcut.
- assumption
- ]
- | cut(O \lt m1)
- [ cut(O \lt n1)
- [ apply (lt_to_le_to_lt ? ((h11 i j)*m1 + m1) )
- [ unfold ha.
- apply (lt_plus_r).
- assumption
- | rewrite > sym_plus.
- rewrite > (sym_times (h11 i j) m1).
- rewrite > times_n_Sm.
- rewrite > sym_times.
- apply (le_times_l).
- assumption
- ]
- | apply not_le_to_lt.unfold.intro.
- generalize in match H12.
- apply (le_n_O_elim ? H11).
- apply le_to_not_lt.
- apply le_O_n
- ]
- | apply not_le_to_lt.unfold.intro.
- generalize in match H10.
- apply (le_n_O_elim ? H11).
- apply le_to_not_lt.
- apply le_O_n
- ]
- ]
- | rewrite > (mod_plus_times m1 (h11 i j) (h12 i j)).
- reflexivity.
- assumption
- ]
- | rewrite > (div_plus_times m1 (h11 i j) (h12 i j)).
- reflexivity.
- assumption
- ]
- ]
-| apply (eq_iter_p_gen1)
- [ intros. reflexivity
- | intros.
- apply (eq_iter_p_gen1)
- [ intros. reflexivity
- | intros.
- rewrite > (div_plus_times)
- [ rewrite > (mod_plus_times)
- [ reflexivity
- | elim (H3 x x1 H5 H7 H6 H8).
- assumption
- ]
- | elim (H3 x x1 H5 H7 H6 H8).
- assumption
- ]
- ]
- ]
-]
-qed.
-
-theorem iter_p_gen_iter_p_gen:
-\forall A:Type.
-\forall baseA: A.
-\forall plusA: A \to A \to A.
-(symmetric A plusA) \to
-(associative A plusA) \to
-(\forall a:A.(plusA a baseA) = a)\to
-\forall g: nat \to nat \to A.
-\forall n,m.
-\forall p11,p21:nat \to bool.
-\forall p12,p22:nat \to nat \to bool.
-(\forall x,y. x < n \to y < m \to
- (p11 x \land p12 x y) = (p21 y \land p22 y x)) \to
-iter_p_gen n p11 A
- (\lambda x:nat.iter_p_gen m (p12 x) A (\lambda y. g x y) baseA plusA)
- baseA plusA =
-iter_p_gen m p21 A
- (\lambda y:nat.iter_p_gen n (p22 y) A (\lambda x. g x y) baseA plusA )
- baseA plusA.
-intros.
-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)
- n m m n p11 p21 p12 p22)
- [intros.split
- [split
- [split
- [split
- [split
- [apply (andb_true_true ? (p12 j i)).
- rewrite > H3
- [rewrite > H6.rewrite > H7.reflexivity
- |assumption
- |assumption
- ]
- |apply (andb_true_true_r (p11 j)).
- rewrite > H3
- [rewrite > H6.rewrite > H7.reflexivity
- |assumption
- |assumption
- ]
- ]
- |reflexivity
- ]
- |reflexivity
- ]
- |assumption
- ]
- |assumption
- ]
- |intros.split
- [split
- [split
- [split
- [split
- [apply (andb_true_true ? (p22 j i)).
- rewrite < H3
- [rewrite > H6.rewrite > H7.reflexivity
- |assumption
- |assumption
- ]
- |apply (andb_true_true_r (p21 j)).
- rewrite < H3
- [rewrite > H6.rewrite > H7.reflexivity
- |assumption
- |assumption
- ]
- ]
- |reflexivity
- ]
- |reflexivity
- ]
- |assumption
- ]
- |assumption
- ]
- ]
-qed. *)
\ No newline at end of file
\ /
V_______________________________________________________________ *)
-include "arithmetics/sigma_pi.ma".
include "arithmetics/primes.ma".
+include "arithmetics/bigops.ma".
(* binomial coefficient *)
definition bc ≝ λn,k. n!/(k!*(n-k)!).
qed.
theorem bc_n_O: ∀n. bc n O = 1.
-#n >bceq <minus_n_O /2/
+#n >bceq <minus_n_O /2 by injective_plus_r/
qed.
theorem fact_minus: ∀n,k. k < n →
<associative_times in ⊢ (???(?(??%)?));
>(div_times_times ?? (n - k)) in ⊢ (???(??%)) ;
[|>(times_n_O 0) @lt_times //
- |@(le_plus_to_le_r k ??) <plus_minus_m_m /2/]
+ |@(le_plus_to_le_r k ??) <plus_minus_m_m /2 by lt_to_le/]
>associative_times in ⊢ (???(??(??%)));
>fact_minus // <plus_div
[<distributive_times_plus
- >commutative_plus in ⊢ (???%); <plus_n_Sm <plus_minus_m_m [|/2/] @refl
+ >commutative_plus in ⊢ (???%); <plus_n_Sm <plus_minus_m_m [|/2 by lt_to_le/] @refl
|<fact_minus // <associative_times @divides_times // @(bc2 n (S k)) //
|>associative_times >(commutative_times (S k))
- <associative_times @divides_times // @bc2 /2/
+ <associative_times @divides_times // @bc2 /2 by lt_to_le/
|>(times_n_O 0) @lt_times [@(le_1_fact (S k)) | //]
]
qed.
]
qed.
-(*
theorem binomial_law:∀a,b,n.
- (a+b)^n = Σ_{k < S n}((bc n k)*(a^(n-k))*(b^k)).
+ (a+b)^n = ∑_{k < S n}((bc n k)*(a^(n-k))*(b^k)).
#a #b #n (elim n) //
--n #n #Hind normalize in ⊢ (? ? % ?).
+-n #n #Hind normalize in ⊢ (??%?); >commutative_times
>bigop_Strue // >Hind >distributive_times_plus
-<(minus_n_n (S n)) <commutative_times <(commutative_times b)
+<(minus_n_n (S n)) <commutative_times (* <(commutative_times b) *)
(* hint??? *)
>(bigop_distr ???? natDop ? a) >(bigop_distr ???? natDop ? b)
>bigop_Strue in ⊢ (??(??%)?) // <associative_plus
#a #n cut (S a = a + 1) // #H >H
>binomial_law @same_bigop //
qed.
+definition M ≝ λm.bc (S(2*m)) m.
+
+theorem lt_M: ∀m. O < m → M m < exp 2 (2*m).
+#m #posm @(lt_times_n_to_lt_l 2)
+ |change in ⊢ (? ? %) with (exp 2 (S(2*m))).
+ change in ⊢ (? ? (? % ?)) with (1+1).
+ rewrite > exp_plus_sigma_p.
+ apply (le_to_lt_to_lt ? (sigma_p (S (S (2*m))) (λk:nat.orb (eqb k m) (eqb k (S m)))
+ (λk:nat.bc (S (2*m)) k*(1)\sup(S (2*m)-k)*(1)\sup(k))))
+ [rewrite > (sigma_p_gi ? ? m)
+ [rewrite > (sigma_p_gi ? ? (S m))
+ [rewrite > (false_to_eq_sigma_p O (S(S(2*m))))
+ [simplify in ⊢ (? ? (? ? (? ? %))).
+ simplify in ⊢ (? % ?).
+ rewrite < exp_SO_n.rewrite < exp_SO_n.
+ rewrite < exp_SO_n.rewrite < exp_SO_n.
+ rewrite < times_n_SO.rewrite < times_n_SO.
+ rewrite < times_n_SO.rewrite < times_n_SO.
+ apply le_plus
+ [unfold M.apply le_n
+ |apply le_plus_l.unfold M.
+ change in \vdash (? ? %) with (fact (S(2*m))/(fact (S m)*(fact ((2*m)-m)))).
+ simplify in \vdash (? ? (? ? (? ? (? (? % ?))))).
+ rewrite < plus_n_O.rewrite < minus_plus_m_m.
+ rewrite < sym_times in \vdash (? ? (? ? %)).
+ change in \vdash (? % ?) with (fact (S(2*m))/(fact m*(fact (S(2*m)-m)))).
+ simplify in \vdash (? (? ? (? ? (? (? (? %) ?)))) ?).
+ rewrite < plus_n_O.change in \vdash (? (? ? (? ? (? (? % ?)))) ?) with (S m + m).
+ rewrite < minus_plus_m_m.
+ apply le_n
+ ]
+ |apply le_O_n
+ |intros.
+ elim (eqb i m);elim (eqb i (S m));reflexivity
+ ]
+ |apply le_S_S.apply le_S_S.
+ apply le_times_n.
+ apply le_n_Sn
+ |rewrite > (eq_to_eqb_true ? ? (refl_eq ? (S m))).
+ rewrite > (not_eq_to_eqb_false (S m) m)
+ [reflexivity
+ |intro.apply (not_eq_n_Sn m).
+ apply sym_eq.assumption
+ ]
+ ]
+ |apply le_S.apply le_S_S.
+ apply le_times_n.
+ apply le_n_Sn
+ |rewrite > (eq_to_eqb_true ? ? (refl_eq ? (S m))).
+ reflexivity
+ ]
+ |rewrite > (bool_to_nat_to_eq_sigma_p (S(S(2*m))) ? (\lambda k.true) ?
+ (\lambda k.bool_to_nat (eqb k m\lor eqb k (S m))*(bc (S (2*m)) k*(1)\sup(S (2*m)-k)*(1)\sup(k))))
+ in \vdash (? % ?)
+ [apply lt_sigma_p
+ [intros.elim (eqb i m\lor eqb i (S m))
+ [rewrite > sym_times.rewrite < times_n_SO.apply le_n
+ |apply le_O_n
+ ]
+ |apply (ex_intro ? ? O).
+ split
+ [split[apply lt_O_S|reflexivity]
+ |rewrite > (not_eq_to_eqb_false ? ? (not_eq_O_S m)).
+ rewrite > (not_eq_to_eqb_false ? ? (lt_to_not_eq ? ? H)).
+ simplify in \vdash (? % ?).
+ rewrite < exp_SO_n.rewrite < exp_SO_n.
+ rewrite > bc_n_O.simplify.
+ apply le_n
+ ]
+ ]
+ |intros.rewrite > sym_times in \vdash (? ? ? %).
+ rewrite < times_n_SO.
+ reflexivity
+ ]
+ ]
+ ]
+qed.
+
(*
theorem exp_Sn_SSO: \forall n. exp (S n) 2 = S((exp n 2) + 2*n).
\ /
V_______________________________________________________________ *)
-(* To be ported
+include "arithmetics/primes.ma".
include "arithmetics/bigops.ma".
-definition natAop ≝ mk_Aop nat 0 plus (λa.refl ? a) (λn.sym_eq ??? (plus_n_O n))
- (λa,b,c.sym_eq ??? (associative_plus a b c)).
-
-definition natACop ≝ mk_ACop nat 0 natAop commutative_plus.
-
-definition natDop ≝ mk_Dop nat 0 natACop times (λn.(sym_eq ??? (times_n_O n)))
- distributive_times_plus.
-
-unification hint 0 ≔ ;
- S ≟ natAop
-(* ---------------------------------------- *) ⊢
- plus ≡ op ? ? S.
-
-unification hint 0 ≔ ;
- S ≟ natACop
-(* ---------------------------------------- *) ⊢
- plus ≡ op ? ? S.
-
-unification hint 0 ≔ ;
- S ≟ natDop
-(* ---------------------------------------- *) ⊢
- plus ≡ sum ? ? S.
-
-unification hint 0 ≔ ;
- S ≟ natDop
-(* ---------------------------------------- *) ⊢
- times ≡ prod ? ? S.
-
(* Sigma e Pi *)
notation "∑_{ ident i < n | p } f"
with precedence 80
-for @{'bigop $n plus 0 (λ${ident i}.$p) (λ${ident i}. $f)}.
+for @{'bigop $n plus 0 (λ${ident i}. $p) (λ${ident i}. $f)}.
notation "∑_{ ident i < n } 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
-\lambda p,m,x.
- match p_ord x p with
- [pair q r \Rightarrow r*m+q].
-
-theorem eq_p_ord_times: \forall p,m,x.
-p_ord_times p m x = (ord_rem x p)*m+(ord x p).
-intros.unfold p_ord_times. unfold ord_rem.
-unfold ord.
-elim (p_ord x p).
-reflexivity.
-qed.
+(* instances of associative and commutative operations *)
-theorem div_p_ord_times:
-\forall p,m,x. ord x p < m \to p_ord_times p m x / m = ord_rem x p.
-intros.rewrite > eq_p_ord_times.
-apply div_plus_times.
-assumption.
-qed.
+definition plusA ≝ mk_Aop nat 0 plus (λa.refl ? a) (λn.sym_eq ??? (plus_n_O n))
+ (λa,b,c.sym_eq ??? (associative_plus a b c)).
+
+definition plusAC ≝ mk_ACop nat 0 plusA commutative_plus.
-theorem mod_p_ord_times:
-\forall p,m,x. ord x p < m \to p_ord_times p m x \mod m = ord x p.
-intros.rewrite > eq_p_ord_times.
-apply mod_plus_times.
-assumption.
-qed.
+definition timesA ≝ mk_Aop nat 1 times
+ (λa.sym_eq ??? (plus_n_O a)) (λn.sym_eq ??? (times_n_1 n))
+ (λa,b,c.sym_eq ??? (associative_times a b c)).
+
+definition timesAC ≝ mk_ACop nat 1 timesA commutative_times.
-lemma lt_times_to_lt_O: \forall i,n,m:nat. i < n*m \to O < m.
-intros.
-elim (le_to_or_lt_eq O ? (le_O_n m))
- [assumption
- |apply False_ind.
- rewrite < H1 in H.
- rewrite < times_n_O in H.
- apply (not_le_Sn_O ? H)
- ]
-qed.
+definition natD ≝ mk_Dop nat 0 plusAC times (λn.(sym_eq ??? (times_n_O n)))
+ distributive_times_plus.
+
+(********************************************************)
-theorem iter_p_gen_knm:
-\forall A:Type.
-\forall baseA: A.
-\forall plusA: A \to A \to A.
-(symmetric A plusA) \to
-(associative A plusA) \to
-(\forall a:A.(plusA a baseA) = a)\to
-\forall g: nat \to A.
-\forall h2:nat \to nat \to nat.
-\forall h11,h12:nat \to nat.
-\forall k,n,m.
-\forall p1,p21:nat \to bool.
-\forall p22:nat \to nat \to bool.
-(\forall x. x < k \to p1 x = true \to
-p21 (h11 x) = true \land p22 (h11 x) (h12 x) = true
-\land h2 (h11 x) (h12 x) = x
-\land (h11 x) < n \land (h12 x) < m) \to
-(\forall i,j. i < n \to j < m \to p21 i = true \to p22 i j = true \to
-p1 (h2 i j) = true \land
-h11 (h2 i j) = i \land h12 (h2 i j) = j
-\land h2 i j < k) \to
-iter_p_gen k p1 A g baseA plusA =
-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.
-intros.
-rewrite < (iter_p_gen2' n m p21 p22 ? ? ? ? H H1 H2).
-apply sym_eq.
-apply (eq_iter_p_gen_gh A baseA plusA H H1 H2 g ? (\lambda x.(h11 x)*m+(h12 x)))
- [intros.
- elim (H4 (i/m) (i \mod m));clear H4
- [elim H7.clear H7.
- elim H4.clear H4.
- assumption
- |apply (lt_times_to_lt_div ? ? ? H5)
- |apply lt_mod_m_m.
- apply (lt_times_to_lt_O ? ? ? H5)
- |apply (andb_true_true ? ? H6)
- |apply (andb_true_true_r ? ? H6)
- ]
- |intros.
- elim (H4 (i/m) (i \mod m));clear H4
- [elim H7.clear H7.
- elim H4.clear H4.
- rewrite > H10.
- rewrite > H9.
- apply sym_eq.
- apply div_mod.
- apply (lt_times_to_lt_O ? ? ? H5)
- |apply (lt_times_to_lt_div ? ? ? H5)
- |apply lt_mod_m_m.
- apply (lt_times_to_lt_O ? ? ? H5)
- |apply (andb_true_true ? ? H6)
- |apply (andb_true_true_r ? ? H6)
- ]
- |intros.
- elim (H4 (i/m) (i \mod m));clear H4
- [elim H7.clear H7.
- elim H4.clear H4.
- assumption
- |apply (lt_times_to_lt_div ? ? ? H5)
- |apply lt_mod_m_m.
- apply (lt_times_to_lt_O ? ? ? H5)
- |apply (andb_true_true ? ? H6)
- |apply (andb_true_true_r ? ? H6)
- ]
- |intros.
- elim (H3 j H5 H6).
- elim H7.clear H7.
- elim H9.clear H9.
- elim H7.clear H7.
- rewrite > div_plus_times
- [rewrite > mod_plus_times
- [rewrite > H9.
- rewrite > H12.
- reflexivity.
- |assumption
- ]
- |assumption
- ]
- |intros.
- elim (H3 j H5 H6).
- elim H7.clear H7.
- elim H9.clear H9.
- elim H7.clear H7.
- rewrite > div_plus_times
- [rewrite > mod_plus_times
- [assumption
- |assumption
- ]
- |assumption
- ]
- |intros.
- elim (H3 j H5 H6).
- elim H7.clear H7.
- elim H9.clear H9.
- elim H7.clear H7.
- apply (lt_to_le_to_lt ? ((h11 j)*m+m))
- [apply monotonic_lt_plus_r.
- assumption
- |rewrite > sym_plus.
- change with ((S (h11 j)*m) \le n*m).
- apply monotonic_le_times_l.
- assumption
- ]
- ]
+theorem sigma_const: ∀n:nat. ∑_{i<n} 1 = n.
+#n elim n // #n1 >bigop_Strue //
qed.
-theorem iter_p_gen_divides:
-\forall A:Type.
-\forall baseA: A.
-\forall plusA: A \to A \to A.
-\forall n,m,p:nat.O < n \to prime p \to Not (divides p n) \to
-\forall g: nat \to A.
-(symmetric A plusA) \to (associative A plusA) \to (\forall a:A.(plusA a baseA) = a)
-
-\to
-
-iter_p_gen (S (n*(exp p m))) (\lambda x.divides_b x (n*(exp p m))) A g baseA plusA =
-iter_p_gen (S n) (\lambda x.divides_b x n) A
- (\lambda x.iter_p_gen (S m) (\lambda y.true) A (\lambda y.g (x*(exp p y))) baseA plusA) baseA plusA.
-intros.
-cut (O < p)
- [rewrite < (iter_p_gen2 ? ? ? ? ? ? ? ? H3 H4 H5).
- apply (trans_eq ? ?
- (iter_p_gen (S n*S m) (\lambda x:nat.divides_b (x/S m) n) A
- (\lambda x:nat.g (x/S m*(p)\sup(x\mod S m))) baseA plusA) )
- [apply sym_eq.
- apply (eq_iter_p_gen_gh ? ? ? ? ? ? g ? (p_ord_times p (S m)))
- [ assumption
- | assumption
- | assumption
- |intros.
- lapply (divides_b_true_to_lt_O ? ? H H7).
- apply divides_to_divides_b_true
- [rewrite > (times_n_O O).
- apply lt_times
- [assumption
- |apply lt_O_exp.assumption
- ]
- |apply divides_times
- [apply divides_b_true_to_divides.assumption
- |apply (witness ? ? (p \sup (m-i \mod (S m)))).
- rewrite < exp_plus_times.
- apply eq_f.
- rewrite > sym_plus.
- apply plus_minus_m_m.
- autobatch by le_S_S_to_le, lt_mod_m_m, lt_O_S;
- ]
- ]
- |intros.
- lapply (divides_b_true_to_lt_O ? ? H H7).
- unfold p_ord_times.
- rewrite > (p_ord_exp1 p ? (i \mod (S m)) (i/S m))
- [change with ((i/S m)*S m+i \mod S m=i).
- apply sym_eq.
- apply div_mod.
- apply lt_O_S
- |assumption
- |unfold Not.intro.
- apply H2.
- apply (trans_divides ? (i/ S m))
- [assumption|
- apply divides_b_true_to_divides;assumption]
- |apply sym_times.
- ]
- |intros.
- apply le_S_S.
- apply le_times
- [apply le_S_S_to_le.
- change with ((i/S m) < S n).
- apply (lt_times_to_lt_l m).
- apply (le_to_lt_to_lt ? i);[2:assumption]
- autobatch by eq_plus_to_le, div_mod, lt_O_S.
- |apply le_exp
- [assumption
- |apply le_S_S_to_le.
- apply lt_mod_m_m.
- apply lt_O_S
- ]
- ]
- |intros.
- cut (ord j p < S m)
- [rewrite > div_p_ord_times
- [apply divides_to_divides_b_true
- [apply lt_O_ord_rem
- [elim H1.assumption
- |apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- ]
- |cut (n = ord_rem (n*(exp p m)) p)
- [rewrite > Hcut2.
- apply divides_to_divides_ord_rem
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- |unfold ord_rem.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |apply sym_times
- ]
- ]
- ]
- |assumption
- ]
- |cut (m = ord (n*(exp p m)) p)
- [apply le_S_S.
- rewrite > Hcut1.
- apply divides_to_le_ord
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- |unfold ord.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |apply sym_times
- ]
- ]
- ]
- |intros.
- cut (ord j p < S m)
- [rewrite > div_p_ord_times
- [rewrite > mod_p_ord_times
- [rewrite > sym_times.
- apply sym_eq.
- apply exp_ord
- [elim H1.assumption
- |apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- ]
- |cut (m = ord (n*(exp p m)) p)
- [apply le_S_S.
- rewrite > Hcut2.
- apply divides_to_le_ord
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- |unfold ord.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |apply sym_times
- ]
- ]
- ]
- |assumption
- ]
- |cut (m = ord (n*(exp p m)) p)
- [apply le_S_S.
- rewrite > Hcut1.
- apply divides_to_le_ord
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- |unfold ord.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |apply sym_times
- ]
- ]
- ]
- |intros.
- rewrite > eq_p_ord_times.
- rewrite > sym_plus.
- apply (lt_to_le_to_lt ? (S m +ord_rem j p*S m))
- [apply lt_plus_l.
- apply le_S_S.
- cut (m = ord (n*(p \sup m)) p)
- [rewrite > Hcut1.
- apply divides_to_le_ord
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- |unfold ord.
- rewrite > sym_times.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |reflexivity
- ]
- ]
- |change with (S (ord_rem j p)*S m \le S n*S m).
- apply le_times_l.
- apply le_S_S.
- cut (n = ord_rem (n*(p \sup m)) p)
- [rewrite > Hcut1.
- apply divides_to_le
- [apply lt_O_ord_rem
- [elim H1.assumption
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- ]
- |apply divides_to_divides_ord_rem
- [apply (divides_b_true_to_lt_O ? ? ? H7).
- rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |rewrite > (times_n_O O).
- apply lt_times
- [assumption|apply lt_O_exp.assumption]
- |assumption
- |apply divides_b_true_to_divides.
- assumption
- ]
- ]
- |unfold ord_rem.
- rewrite > sym_times.
- rewrite > (p_ord_exp1 p ? m n)
- [reflexivity
- |assumption
- |assumption
- |reflexivity
- ]
- ]
- ]
- ]
- |apply eq_iter_p_gen
-
- [intros.
- elim (divides_b (x/S m) n);reflexivity
- |intros.reflexivity
+(* monotonicity; these roperty should be expressed at a more
+genral level *)
+
+theorem le_pi:
+∀n.∀p:nat → bool.∀g1,g2:nat → nat.
+ (∀i.i<n → p i = true → g1 i ≤ g2 i ) →
+ ∏_{i < n | p i} (g1 i) ≤ ∏_{i < n | p i} (g2 i).
+#n #p #g1 #g2 elim n
+ [#_ @le_n
+ |#n1 #Hind #Hle cases (true_or_false (p n1)) #Hcase
+ [ >bigop_Strue // >bigop_Strue // @le_times
+ [@Hle // |@Hind #i #lti #Hpi @Hle [@lt_to_le @le_S_S @lti|@Hpi]]
+ |>bigop_Sfalse // >bigop_Sfalse // @Hind
+ #i #lti #Hpi @Hle [@lt_to_le @le_S_S @lti|@Hpi]
]
]
-|elim H1.apply lt_to_le.assumption
-]
qed.
+theorem exp_sigma: ∀n,a,p.
+ ∏_{i < n | p i} a = exp a (∑_{i < n | p i} 1).
+#n #a #p elim n // #n1 cases (true_or_false (p n1)) #Hcase
+ [>bigop_Strue // >bigop_Strue // |>bigop_Sfalse // >bigop_Sfalse //]
+qed.
-
-theorem iter_p_gen_2_eq:
-\forall A:Type.
-\forall baseA: A.
-\forall plusA: A \to A \to A.
-(symmetric A plusA) \to
-(associative A plusA) \to
-(\forall a:A.(plusA a baseA) = a)\to
-\forall g: nat \to nat \to A.
-\forall h11,h12,h21,h22: nat \to nat \to nat.
-\forall n1,m1,n2,m2.
-\forall p11,p21:nat \to bool.
-\forall p12,p22:nat \to nat \to bool.
-(\forall i,j. i < n2 \to j < m2 \to p21 i = true \to p22 i j = true \to
-p11 (h11 i j) = true \land p12 (h11 i j) (h12 i j) = true
-\land h21 (h11 i j) (h12 i j) = i \land h22 (h11 i j) (h12 i j) = j
-\land h11 i j < n1 \land h12 i j < m1) \to
-(\forall i,j. i < n1 \to j < m1 \to p11 i = true \to p12 i j = true \to
-p21 (h21 i j) = true \land p22 (h21 i j) (h22 i j) = true
-\land h11 (h21 i j) (h22 i j) = i \land h12 (h21 i j) (h22 i j) = j
-\land (h21 i j) < n2 \land (h22 i j) < m2) \to
-iter_p_gen n1 p11 A
- (\lambda x:nat .iter_p_gen m1 (p12 x) A (\lambda y. g x y) baseA plusA)
- baseA plusA =
-iter_p_gen n2 p21 A
- (\lambda x:nat .iter_p_gen m2 (p22 x) A (\lambda y. g (h11 x y) (h12 x y)) baseA plusA )
- baseA plusA.
-
-intros.
-rewrite < (iter_p_gen2' ? ? ? ? ? ? ? ? H H1 H2).
-letin ha:= (\lambda x,y.(((h11 x y)*m1) + (h12 x y))).
-letin ha12:= (\lambda x.(h21 (x/m1) (x \mod m1))).
-letin ha22:= (\lambda x.(h22 (x/m1) (x \mod m1))).
-
-apply (trans_eq ? ?
-(iter_p_gen n2 p21 A (\lambda x:nat. iter_p_gen m2 (p22 x) A
- (\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))
-[
- apply (iter_p_gen_knm A baseA plusA H H1 H2 (\lambda e. (g (e/m1) (e \mod m1))) ha ha12 ha22);intros
- [ elim (and_true ? ? H6).
- cut(O \lt m1)
- [ cut(x/m1 < n1)
- [ cut((x \mod m1) < m1)
- [ elim (H4 ? ? Hcut1 Hcut2 H7 H8).
- elim H9.clear H9.
- elim H11.clear H11.
- elim H9.clear H9.
- elim H11.clear H11.
- split
- [ split
- [ split
- [ split
- [ assumption
- | assumption
- ]
- | unfold ha.
- unfold ha12.
- unfold ha22.
- rewrite > H14.
- rewrite > H13.
- apply sym_eq.
- apply div_mod.
- assumption
- ]
- | assumption
- ]
- | assumption
- ]
- | apply lt_mod_m_m.
- assumption
- ]
- | apply (lt_times_n_to_lt m1)
- [ assumption
- | apply (le_to_lt_to_lt ? x)
- [ apply (eq_plus_to_le ? ? (x \mod m1)).
- apply div_mod.
- assumption
- | assumption
- ]
- ]
- ]
- | apply not_le_to_lt.unfold.intro.
- generalize in match H5.
- apply (le_n_O_elim ? H9).
- rewrite < times_n_O.
- apply le_to_not_lt.
- apply le_O_n.
- ]
- | elim (H3 ? ? H5 H6 H7 H8).
- elim H9.clear H9.
- elim H11.clear H11.
- elim H9.clear H9.
- elim H11.clear H11.
- cut(((h11 i j)*m1 + (h12 i j))/m1 = (h11 i j))
- [ cut(((h11 i j)*m1 + (h12 i j)) \mod m1 = (h12 i j))
- [ split
- [ split
- [ split
- [ apply true_to_true_to_andb_true
- [ rewrite > Hcut.
- assumption
- | rewrite > Hcut1.
- rewrite > Hcut.
- assumption
- ]
- | unfold ha.
- unfold ha12.
- rewrite > Hcut1.
- rewrite > Hcut.
- assumption
- ]
- | unfold ha.
- unfold ha22.
- rewrite > Hcut1.
- rewrite > Hcut.
- assumption
- ]
- | cut(O \lt m1)
- [ cut(O \lt n1)
- [ apply (lt_to_le_to_lt ? ((h11 i j)*m1 + m1) )
- [ unfold ha.
- apply (lt_plus_r).
- assumption
- | rewrite > sym_plus.
- rewrite > (sym_times (h11 i j) m1).
- rewrite > times_n_Sm.
- rewrite > sym_times.
- apply (le_times_l).
- assumption
- ]
- | apply not_le_to_lt.unfold.intro.
- generalize in match H12.
- apply (le_n_O_elim ? H11).
- apply le_to_not_lt.
- apply le_O_n
- ]
- | apply not_le_to_lt.unfold.intro.
- generalize in match H10.
- apply (le_n_O_elim ? H11).
- apply le_to_not_lt.
- apply le_O_n
- ]
- ]
- | rewrite > (mod_plus_times m1 (h11 i j) (h12 i j)).
- reflexivity.
- assumption
- ]
- | rewrite > (div_plus_times m1 (h11 i j) (h12 i j)).
- reflexivity.
- assumption
- ]
- ]
-| apply (eq_iter_p_gen1)
- [ intros. reflexivity
- | intros.
- apply (eq_iter_p_gen1)
- [ intros. reflexivity
- | intros.
- rewrite > (div_plus_times)
- [ rewrite > (mod_plus_times)
- [ reflexivity
- | elim (H3 x x1 H5 H7 H6 H8).
- assumption
- ]
- | elim (H3 x x1 H5 H7 H6 H8).
- assumption
- ]
- ]
+theorem times_pi: ∀n,p,f,g.
+ ∏_{i<n|p i}(f i*g i) = ∏_{i<n|p i}(f i) * ∏_{i<n|p i}(g i).
+#n #p #f #g elim n // #n1 cases (true_or_false (p n1)) #Hcase #Hind
+ [>bigop_Strue // >bigop_Strue // >bigop_Strue //
+ |>bigop_Sfalse // >bigop_Sfalse // >bigop_Sfalse //
]
-]
qed.
-theorem iter_p_gen_iter_p_gen:
-\forall A:Type.
-\forall baseA: A.
-\forall plusA: A \to A \to A.
-(symmetric A plusA) \to
-(associative A plusA) \to
-(\forall a:A.(plusA a baseA) = a)\to
-\forall g: nat \to nat \to A.
-\forall n,m.
-\forall p11,p21:nat \to bool.
-\forall p12,p22:nat \to nat \to bool.
-(\forall x,y. x < n \to y < m \to
- (p11 x \land p12 x y) = (p21 y \land p22 y x)) \to
-iter_p_gen n p11 A
- (\lambda x:nat.iter_p_gen m (p12 x) A (\lambda y. g x y) baseA plusA)
- baseA plusA =
-iter_p_gen m p21 A
- (\lambda y:nat.iter_p_gen n (p22 y) A (\lambda x. g x y) baseA plusA )
- baseA plusA.
-intros.
-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)
- n m m n p11 p21 p12 p22)
- [intros.split
- [split
- [split
- [split
- [split
- [apply (andb_true_true ? (p12 j i)).
- rewrite > H3
- [rewrite > H6.rewrite > H7.reflexivity
- |assumption
- |assumption
- ]
- |apply (andb_true_true_r (p11 j)).
- rewrite > H3
- [rewrite > H6.rewrite > H7.reflexivity
- |assumption
- |assumption
- ]
- ]
- |reflexivity
- ]
- |reflexivity
- ]
- |assumption
- ]
- |assumption
- ]
- |intros.split
- [split
- [split
- [split
- [split
- [apply (andb_true_true ? (p22 j i)).
- rewrite < H3
- [rewrite > H6.rewrite > H7.reflexivity
- |assumption
- |assumption
- ]
- |apply (andb_true_true_r (p21 j)).
- rewrite < H3
- [rewrite > H6.rewrite > H7.reflexivity
- |assumption
- |assumption
- ]
- ]
- |reflexivity
- ]
- |reflexivity
- ]
- |assumption
- ]
- |assumption
- ]
+theorem pi_1: ∀n,p.
+ ∏_{i < n | p i} 1 = 1.
+#n #p elim n // #n1 #Hind cases (true_or_false (p n1)) #Hc
+ [>bigop_Strue >Hind // |>bigop_Sfalse // ]
+qed.
+
+theorem exp_pi: ∀n,m,p,f.
+ ∏_{i < n | p i}(exp (f i) m) = exp (∏_{i < n | p i}(f i)) m.
+#n #m #p #f elim m
+ [@pi_1
+ |#m1 #Hind >times_pi >Hind %
]
-qed. *)*)
\ No newline at end of file
+qed.