--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/Q/q".
+
+include "Z/compare.ma".
+include "Z/plus.ma".
+
+(* a fraction is a list of Z-coefficients for primes, in natural
+order. The last coefficient must eventually be different from 0 *)
+
+inductive fraction : Set \def
+ pp : nat \to fraction
+| nn: nat \to fraction
+| cons : Z \to fraction \to fraction.
+
+inductive ratio : Set \def
+ one : ratio
+ | frac : fraction \to ratio.
+
+(* a rational number is either O or a ratio with a sign *)
+inductive Q : Set \def
+ OQ : Q
+ | Qpos : ratio \to Q
+ | Qneg : ratio \to Q.
+
+(* double elimination principles *)
+theorem fraction_elim2:
+\forall R:fraction \to fraction \to Prop.
+(\forall n:nat.\forall g:fraction.R (pp n) g) \to
+(\forall n:nat.\forall g:fraction.R (nn n) g) \to
+(\forall x:Z.\forall f:fraction.\forall m:nat.R (cons x f) (pp m)) \to
+(\forall x:Z.\forall f:fraction.\forall m:nat.R (cons x f) (nn m)) \to
+(\forall x,y:Z.\forall f,g:fraction.R f g \to R (cons x f) (cons y g)) \to
+\forall f,g:fraction. R f g.
+intros 7.
+elim f
+[ apply H
+| apply H1
+| elim g
+ [ apply H2
+ | apply H3
+ | auto
+ (*apply H4.
+ apply H5*)
+ ]
+]
+qed.
+
+(* boolean equality *)
+let rec eqfb f g \def
+match f with
+[ (pp n) \Rightarrow
+ match g with
+ [ (pp m) \Rightarrow eqb n m
+ | (nn m) \Rightarrow false
+ | (cons y g1) \Rightarrow false]
+| (nn n) \Rightarrow
+ match g with
+ [ (pp m) \Rightarrow false
+ | (nn m) \Rightarrow eqb n m
+ | (cons y g1) \Rightarrow false]
+| (cons x f1) \Rightarrow
+ match g with
+ [ (pp m) \Rightarrow false
+ | (nn m) \Rightarrow false
+ | (cons y g1) \Rightarrow andb (eqZb x y) (eqfb f1 g1)]].
+
+(* discrimination *)
+definition aux \def
+ \lambda f. match f with
+ [ (pp n) \Rightarrow n
+ | (nn n) \Rightarrow n
+ | (cons x f) \Rightarrow O].
+
+definition fhd \def
+\lambda f. match f with
+ [ (pp n) \Rightarrow (pos n)
+ | (nn n) \Rightarrow (neg n)
+ | (cons x f) \Rightarrow x].
+
+definition ftl \def
+\lambda f. match f with
+ [ (pp n) \Rightarrow (pp n)
+ | (nn n) \Rightarrow (nn n)
+ | (cons x f) \Rightarrow f].
+
+theorem injective_pp : injective nat fraction pp.
+unfold injective.
+intros.
+change with ((aux(pp x)) = (aux (pp y))).
+auto.
+(*apply eq_f.
+assumption.*)
+qed.
+
+theorem injective_nn : injective nat fraction nn.
+unfold injective.
+intros.
+change with ((aux (nn x)) = (aux (nn y))).
+auto.
+(*apply eq_f.
+assumption.*)
+qed.
+
+theorem eq_cons_to_eq1: \forall f,g:fraction.\forall x,y:Z.
+(cons x f) = (cons y g) \to x = y.
+intros.
+change with ((fhd (cons x f)) = (fhd (cons y g))).
+auto.
+(*apply eq_f.assumption.*)
+qed.
+
+theorem eq_cons_to_eq2: \forall x,y:Z.\forall f,g:fraction.
+(cons x f) = (cons y g) \to f = g.
+intros.
+change with ((ftl (cons x f)) = (ftl (cons y g))).
+auto.
+(*apply eq_f.assumption.*)
+qed.
+
+theorem not_eq_pp_nn: \forall n,m:nat. pp n \neq nn m.
+intros.
+unfold Not.
+intro.
+change with match (pp n) with
+[ (pp n) \Rightarrow False
+| (nn n) \Rightarrow True
+| (cons x f) \Rightarrow True].
+rewrite > H.
+simplify.
+exact I.
+qed.
+
+theorem not_eq_pp_cons:
+\forall n:nat.\forall x:Z. \forall f:fraction.
+pp n \neq cons x f.
+intros.
+unfold Not.
+intro.
+change with match (pp n) with
+[ (pp n) \Rightarrow False
+| (nn n) \Rightarrow True
+| (cons x f) \Rightarrow True].
+rewrite > H.
+simplify.
+exact I.
+qed.
+
+theorem not_eq_nn_cons:
+\forall n:nat.\forall x:Z. \forall f:fraction.
+nn n \neq cons x f.
+intros.
+unfold Not.
+intro.
+change with match (nn n) with
+[ (pp n) \Rightarrow True
+| (nn n) \Rightarrow False
+| (cons x f) \Rightarrow True].
+rewrite > H.
+simplify.
+exact I.
+qed.
+
+theorem decidable_eq_fraction: \forall f,g:fraction.
+decidable (f = g).
+intros.
+unfold decidable.
+apply (fraction_elim2 (\lambda f,g. f=g \lor (f=g \to False)))
+[ intros.
+ elim g1
+ [ elim ((decidable_eq_nat n n1) : n=n1 \lor (n=n1 \to False))
+ [ auto
+ (*left.
+ apply eq_f.
+ assumption*)
+ | right.
+ intro.
+ auto
+ (*apply H.
+ apply injective_pp.
+ assumption*)
+ ]
+ | auto
+ (*right.
+ apply not_eq_pp_nn*)
+ | auto
+ (*right.
+ apply not_eq_pp_cons*)
+ ]
+| intros.
+ elim g1
+ [ right.
+ intro.
+ apply (not_eq_pp_nn n1 n).
+ auto
+ (*apply sym_eq.
+ assumption*)
+ | elim ((decidable_eq_nat n n1) : n=n1 \lor (n=n1 \to False))
+ [ auto
+ (*left.
+ apply eq_f.
+ assumption*)
+ | right.
+ intro.
+ auto
+ (*apply H.
+ apply injective_nn.
+ assumption*)
+ ]
+ | auto
+ (*right.
+ apply not_eq_nn_cons*)
+ ]
+| intros.
+ right.
+ intro.
+ apply (not_eq_pp_cons m x f1).
+ auto
+ (*apply sym_eq.
+ assumption*)
+| intros.
+ right.
+ intro.
+ apply (not_eq_nn_cons m x f1).
+ auto
+ (*apply sym_eq.
+ assumption*)
+| intros.
+ elim H
+ [ elim ((decidable_eq_Z x y) : x=y \lor (x=y \to False))
+ [ auto
+ (*left.
+ apply eq_f2;
+ assumption*)
+ | right.
+ intro.
+ auto
+ (*apply H2.
+ apply (eq_cons_to_eq1 f1 g1).
+ assumption*)
+ ]
+ | right.
+ intro.
+ auto
+ (*apply H1.
+ apply (eq_cons_to_eq2 x y f1 g1).
+ assumption*)
+ ]
+]
+qed.
+
+theorem eqfb_to_Prop: \forall f,g:fraction.
+match (eqfb f g) with
+[true \Rightarrow f=g
+|false \Rightarrow f \neq g].
+intros.
+apply (fraction_elim2
+(\lambda f,g.match (eqfb f g) with
+[true \Rightarrow f=g
+|false \Rightarrow f \neq g]))
+[ intros.
+ elim g1
+ [ simplify.
+ apply eqb_elim
+ [ intro.
+ simplify.
+ auto
+ (*apply eq_f.
+ assumption*)
+ | intro.
+ simplify.
+ unfold Not.
+ intro.
+ auto
+ (*apply H.
+ apply injective_pp.
+ assumption*)
+ ]
+ | simplify.
+ apply not_eq_pp_nn
+ | simplify.
+ apply not_eq_pp_cons
+ ]
+| intros.
+ elim g1
+ [ simplify.
+ unfold Not.
+ intro.
+ apply (not_eq_pp_nn n1 n).
+ auto
+ (*apply sym_eq.
+ assumption*)
+ | simplify.
+ apply eqb_elim
+ [ intro.
+ simplify.
+ auto
+ (*apply eq_f.
+ assumption*)
+ | intro.
+ simplify.
+ unfold Not.
+ intro.
+ auto
+ (*apply H.
+ apply injective_nn.
+ assumption*)
+ ]
+ | simplify.
+ apply not_eq_nn_cons
+ ]
+| intros.
+ simplify.
+ unfold Not.
+ intro.
+ apply (not_eq_pp_cons m x f1).
+ auto
+ (*apply sym_eq.
+ assumption*)
+| intros.
+ simplify.
+ unfold Not.
+ intro.
+ apply (not_eq_nn_cons m x f1).
+ auto
+ (*apply sym_eq.
+ assumption*)
+| intros.
+ simplify.
+ apply eqZb_elim
+ [ intro.
+ generalize in match H.
+ elim (eqfb f1 g1)
+ [ simplify.
+ apply eq_f2
+ [ assumption
+ | (*qui auto non chiude il goal*)
+ apply H2
+ ]
+ | simplify.
+ unfold Not.
+ intro.
+ apply H2.
+ auto
+ (*apply (eq_cons_to_eq2 x y).
+ assumption*)
+ ]
+ | intro.
+ simplify.
+ unfold Not.
+ intro.
+ auto
+ (*apply H1.
+ apply (eq_cons_to_eq1 f1 g1).
+ assumption*)
+ ]
+]
+qed.
+
+let rec finv f \def
+ match f with
+ [ (pp n) \Rightarrow (nn n)
+ | (nn n) \Rightarrow (pp n)
+ | (cons x g) \Rightarrow (cons (Zopp x) (finv g))].
+
+definition Z_to_ratio :Z \to ratio \def
+\lambda x:Z. match x with
+[ OZ \Rightarrow one
+| (pos n) \Rightarrow frac (pp n)
+| (neg n) \Rightarrow frac (nn n)].
+
+let rec ftimes f g \def
+ match f with
+ [ (pp n) \Rightarrow
+ match g with
+ [(pp m) \Rightarrow Z_to_ratio (pos n + pos m)
+ | (nn m) \Rightarrow Z_to_ratio (pos n + neg m)
+ | (cons y g1) \Rightarrow frac (cons (pos n + y) g1)]
+ | (nn n) \Rightarrow
+ match g with
+ [(pp m) \Rightarrow Z_to_ratio (neg n + pos m)
+ | (nn m) \Rightarrow Z_to_ratio (neg n + neg m)
+ | (cons y g1) \Rightarrow frac (cons (neg n + y) g1)]
+ | (cons x f1) \Rightarrow
+ match g with
+ [ (pp m) \Rightarrow frac (cons (x + pos m) f1)
+ | (nn m) \Rightarrow frac (cons (x + neg m) f1)
+ | (cons y g1) \Rightarrow
+ match ftimes f1 g1 with
+ [ one \Rightarrow Z_to_ratio (x + y)
+ | (frac h) \Rightarrow frac (cons (x + y) h)]]].
+
+theorem symmetric2_ftimes: symmetric2 fraction ratio ftimes.
+unfold symmetric2.
+intros.
+apply (fraction_elim2 (\lambda f,g.ftimes f g = ftimes g f))
+[ intros.
+ elim g
+ [ change with (Z_to_ratio (pos n + pos n1) = Z_to_ratio (pos n1 + pos n)).
+ auto
+ (*apply eq_f.
+ apply sym_Zplus*)
+ | change with (Z_to_ratio (pos n + neg n1) = Z_to_ratio (neg n1 + pos n)).
+ auto
+ (*apply eq_f.
+ apply sym_Zplus*)
+ | change with (frac (cons (pos n + z) f) = frac (cons (z + pos n) f)).
+ auto
+ (*rewrite < sym_Zplus.
+ reflexivity*)
+ ]
+| intros.
+ elim g
+ [ change with (Z_to_ratio (neg n + pos n1) = Z_to_ratio (pos n1 + neg n)).
+ auto
+ (*apply eq_f.
+ apply sym_Zplus*)
+ | change with (Z_to_ratio (neg n + neg n1) = Z_to_ratio (neg n1 + neg n)).
+ auto
+ (*apply eq_f.
+ apply sym_Zplus*)
+ | change with (frac (cons (neg n + z) f) = frac (cons (z + neg n) f)).
+ auto
+ (*rewrite < sym_Zplus.
+ reflexivity*)
+ ]
+| intros.
+ change with (frac (cons (x1 + pos m) f) = frac (cons (pos m + x1) f)).
+ auto
+ (*rewrite < sym_Zplus.
+ reflexivity*)
+| intros.
+ change with (frac (cons (x1 + neg m) f) = frac (cons (neg m + x1) f)).
+ auto
+ (*rewrite < sym_Zplus.
+ reflexivity*)
+| intros.
+ (*CSC: simplify does something nasty here because of a redex in Zplus *)
+ change with
+ (match ftimes f g with
+ [ one \Rightarrow Z_to_ratio (x1 + y1)
+ | (frac h) \Rightarrow frac (cons (x1 + y1) h)] =
+ match ftimes g f with
+ [ one \Rightarrow Z_to_ratio (y1 + x1)
+ | (frac h) \Rightarrow frac (cons (y1 + x1) h)]).
+ rewrite < H.
+ rewrite < sym_Zplus.
+ reflexivity
+]
+qed.
+
+theorem ftimes_finv : \forall f:fraction. ftimes f (finv f) = one.
+intro.
+elim f
+[ change with (Z_to_ratio (pos n + - (pos n)) = one).
+ auto
+ (*rewrite > Zplus_Zopp.
+ reflexivity*)
+| change with (Z_to_ratio (neg n + - (neg n)) = one).
+ auto
+ (*rewrite > Zplus_Zopp.
+ reflexivity*)
+|
+ (*CSC: simplify does something nasty here because of a redex in Zplus *)
+ (* again: we would need something to help finding the right change *)
+ change with
+ (match ftimes f1 (finv f1) with
+ [ one \Rightarrow Z_to_ratio (z + - z)
+ | (frac h) \Rightarrow frac (cons (z + - z) h)] = one).
+ rewrite > H.
+ rewrite > Zplus_Zopp.
+ reflexivity
+]
+qed.
+
+definition rtimes : ratio \to ratio \to ratio \def
+\lambda r,s:ratio.
+ match r with
+ [one \Rightarrow s
+ | (frac f) \Rightarrow
+ match s with
+ [one \Rightarrow frac f
+ | (frac g) \Rightarrow ftimes f g]].
+
+theorem symmetric_rtimes : symmetric ratio rtimes.
+change with (\forall r,s:ratio. rtimes r s = rtimes s r).
+intros.
+elim r
+[ elim s;reflexivity
+| elim s
+ [ reflexivity
+ | simplify.
+ apply symmetric2_ftimes
+ ]
+]
+qed.
+
+definition rinv : ratio \to ratio \def
+\lambda r:ratio.
+ match r with
+ [one \Rightarrow one
+ | (frac f) \Rightarrow frac (finv f)].
+
+theorem rtimes_rinv: \forall r:ratio. rtimes r (rinv r) = one.
+intro.
+elim r
+[ reflexivity
+| simplify.
+ apply ftimes_finv
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/Z/compare".
+
+include "Z/orders.ma".
+include "nat/compare.ma".
+
+(* boolean equality *)
+definition eqZb : Z \to Z \to bool \def
+\lambda x,y:Z.
+ match x with
+ [ OZ \Rightarrow
+ match y with
+ [ OZ \Rightarrow true
+ | (pos q) \Rightarrow false
+ | (neg q) \Rightarrow false]
+ | (pos p) \Rightarrow
+ match y with
+ [ OZ \Rightarrow false
+ | (pos q) \Rightarrow eqb p q
+ | (neg q) \Rightarrow false]
+ | (neg p) \Rightarrow
+ match y with
+ [ OZ \Rightarrow false
+ | (pos q) \Rightarrow false
+ | (neg q) \Rightarrow eqb p q]].
+
+theorem eqZb_to_Prop:
+\forall x,y:Z.
+match eqZb x y with
+[ true \Rightarrow x=y
+| false \Rightarrow x \neq y].
+intros.
+elim x
+[ elim y
+ [ simplify.
+ reflexivity
+ | simplify.
+ apply not_eq_OZ_pos
+ | simplify.
+ apply not_eq_OZ_neg
+ ]
+| elim y
+ [ simplify.
+ unfold Not.
+ intro.
+ apply (not_eq_OZ_pos n).
+ auto
+ (*apply sym_eq.
+ assumption*)
+ | simplify.
+ apply eqb_elim
+ [ intro.
+ simplify.
+ auto
+ (*apply eq_f.
+ assumption*)
+ | intro.
+ simplify.
+ unfold Not.
+ intro.
+ auto
+ (*apply H.
+ apply inj_pos.
+ assumption*)
+ ]
+ | simplify.
+ apply not_eq_pos_neg
+ ]
+| elim y
+ [ simplify.
+ unfold Not.
+ intro.
+ apply (not_eq_OZ_neg n).
+ auto
+ (*apply sym_eq.
+ assumption*)
+ | simplify.
+ unfold Not.
+ intro.
+ apply (not_eq_pos_neg n1 n).
+ auto
+ (*apply sym_eq.
+ assumption*)
+ | simplify.
+ apply eqb_elim
+ [ intro.
+ simplify.
+ auto
+ (*apply eq_f.
+ assumption*)
+ | intro.
+ simplify.
+ unfold Not.
+ intro.
+ auto
+ (*apply H.
+ apply inj_neg.
+ assumption*)
+ ]
+ ]
+]
+qed.
+
+theorem eqZb_elim: \forall x,y:Z.\forall P:bool \to Prop.
+(x=y \to (P true)) \to (x \neq y \to (P false)) \to P (eqZb x y).
+intros.
+cut
+(match (eqZb x y) with
+[ true \Rightarrow x=y
+| false \Rightarrow x \neq y] \to P (eqZb x y))
+[ apply Hcut.
+ (*NB qui auto non chiude il goal*)
+ apply eqZb_to_Prop
+| elim (eqZb)
+ [ (*NB qui auto non chiude il goal*)
+ apply (H H2)
+ | (*NB qui auto non chiude il goal*)
+ apply (H1 H2)
+ ]
+]
+qed.
+
+definition Z_compare : Z \to Z \to compare \def
+\lambda x,y:Z.
+ match x with
+ [ OZ \Rightarrow
+ match y with
+ [ OZ \Rightarrow EQ
+ | (pos m) \Rightarrow LT
+ | (neg m) \Rightarrow GT ]
+ | (pos n) \Rightarrow
+ match y with
+ [ OZ \Rightarrow GT
+ | (pos m) \Rightarrow (nat_compare n m)
+ | (neg m) \Rightarrow GT]
+ | (neg n) \Rightarrow
+ match y with
+ [ OZ \Rightarrow LT
+ | (pos m) \Rightarrow LT
+ | (neg m) \Rightarrow nat_compare m n ]].
+
+theorem Z_compare_to_Prop :
+\forall x,y:Z. match (Z_compare x y) with
+[ LT \Rightarrow x < y
+| EQ \Rightarrow x=y
+| GT \Rightarrow y < x].
+intros.
+elim x
+[ elim y
+ [ simplify.
+ apply refl_eq
+ | simplify.
+ exact I
+ | simplify.
+ exact I
+ ]
+| elim y
+ [ simplify.
+ exact I
+ | simplify.
+ cut (match (nat_compare n n1) with
+ [ LT \Rightarrow n<n1
+ | EQ \Rightarrow n=n1
+ | GT \Rightarrow n1<n] \to
+ match (nat_compare n n1) with
+ [ LT \Rightarrow (S n) \leq n1
+ | EQ \Rightarrow pos n = pos n1
+ | GT \Rightarrow (S n1) \leq n])
+ [ apply Hcut.
+ (*NB qui auto non chiude il goal*)
+ apply nat_compare_to_Prop
+ | elim (nat_compare n n1)
+ [ simplify.
+ (*NB qui auto non chiude il goal*)
+ exact H
+ | simplify.
+ apply eq_f.
+ (*NB qui auto non chiude il goal*)
+ exact H
+ | simplify.
+ (*NB qui auto non chiude il goal*)
+ exact H
+ ]
+ ]
+ | simplify.
+ exact I
+ ]
+| elim y
+ [ simplify.
+ exact I
+ | simplify.
+ exact I
+ | simplify.
+ cut (match (nat_compare n1 n) with
+ [ LT \Rightarrow n1<n
+ | EQ \Rightarrow n1=n
+ | GT \Rightarrow n<n1] \to
+ match (nat_compare n1 n) with
+ [ LT \Rightarrow (S n1) \leq n
+ | EQ \Rightarrow neg n = neg n1
+ | GT \Rightarrow (S n) \leq n1])
+ [ apply Hcut.
+ (*NB qui auto non chiude il goal*)
+ apply nat_compare_to_Prop
+ | elim (nat_compare n1 n)
+ [ simplify.
+ (*NB qui auto non chiude il goal*)
+ exact H
+ | simplify.
+ apply eq_f.
+ apply sym_eq.
+ (*NB qui auto non chiude il goal*)
+ exact H
+ | simplify.
+ (*NB qui auto non chiude il goal*)
+ exact H
+ ]
+ ]
+ ]
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/Z/orders".
+
+include "Z/z.ma".
+include "nat/orders.ma".
+
+definition Zle : Z \to Z \to Prop \def
+\lambda x,y:Z.
+ match x with
+ [ OZ \Rightarrow
+ match y with
+ [ OZ \Rightarrow True
+ | (pos m) \Rightarrow True
+ | (neg m) \Rightarrow False ]
+ | (pos n) \Rightarrow
+ match y with
+ [ OZ \Rightarrow False
+ | (pos m) \Rightarrow n \leq m
+ | (neg m) \Rightarrow False ]
+ | (neg n) \Rightarrow
+ match y with
+ [ OZ \Rightarrow True
+ | (pos m) \Rightarrow True
+ | (neg m) \Rightarrow m \leq n ]].
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "integer 'less or equal to'" 'leq x y = (cic:/matita/Z/orders/Zle.con x y).
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "integer 'neither less nor equal to'" 'nleq x y =
+ (cic:/matita/logic/connectives/Not.con (cic:/matita/Z/orders/Zle.con x y)).
+
+definition Zlt : Z \to Z \to Prop \def
+\lambda x,y:Z.
+ match x with
+ [ OZ \Rightarrow
+ match y with
+ [ OZ \Rightarrow False
+ | (pos m) \Rightarrow True
+ | (neg m) \Rightarrow False ]
+ | (pos n) \Rightarrow
+ match y with
+ [ OZ \Rightarrow False
+ | (pos m) \Rightarrow n<m
+ | (neg m) \Rightarrow False ]
+ | (neg n) \Rightarrow
+ match y with
+ [ OZ \Rightarrow True
+ | (pos m) \Rightarrow True
+ | (neg m) \Rightarrow m<n ]].
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "integer 'less than'" 'lt x y = (cic:/matita/Z/orders/Zlt.con x y).
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "integer 'not less than'" 'nless x y =
+ (cic:/matita/logic/connectives/Not.con (cic:/matita/Z/orders/Zlt.con x y)).
+
+theorem irreflexive_Zlt: irreflexive Z Zlt.
+unfold irreflexive.
+unfold Not.
+intro.
+elim x
+[ (*qui auto non chiude il goal*)
+ exact H
+| cut (neg n < neg n \to False)
+ [ apply Hcut.
+ (*qui auto non chiude il goal*)
+ apply H
+ | auto
+ (*simplify.
+ unfold lt.
+ apply not_le_Sn_n*)
+ ]
+| cut (pos n < pos n \to False)
+ [ apply Hcut.
+ (*qui auto non chiude il goal*)
+ apply H
+ | auto
+ (*simplify.
+ unfold lt.
+ apply not_le_Sn_n*)
+ ]
+]
+qed.
+
+theorem irrefl_Zlt: irreflexive Z Zlt
+\def irreflexive_Zlt.
+
+theorem Zlt_neg_neg_to_lt:
+\forall n,m:nat. neg n < neg m \to m < n.
+intros.
+(*qui auto non chiude il goal*)
+apply H.
+qed.
+
+theorem lt_to_Zlt_neg_neg: \forall n,m:nat.m < n \to neg n < neg m.
+intros.
+simplify.
+apply H.
+qed.
+
+theorem Zlt_pos_pos_to_lt:
+\forall n,m:nat. pos n < pos m \to n < m.
+intros.
+(*qui auto non chiude il goal*)
+apply H.
+qed.
+
+theorem lt_to_Zlt_pos_pos: \forall n,m:nat.n < m \to pos n < pos m.
+intros.
+simplify.
+apply H.
+qed.
+
+theorem Zlt_to_Zle: \forall x,y:Z. x < y \to Zsucc x \leq y.
+intros 2.
+elim x
+[ (* goal: x=OZ *)
+ cut (OZ < y \to Zsucc OZ \leq y)
+ [ auto
+ (*apply Hcut.
+ assumption*)
+ | simplify.
+ elim y
+ [ simplify.
+ (*qui auto non chiude il goal*)
+ exact H1
+ | simplify.
+ apply le_O_n
+ | simplify.
+ (*qui auto non chiude il goal*)
+ exact H1
+ ]
+ ]
+| (* goal: x=pos *)
+ exact H
+| (* goal: x=neg *)
+ cut (neg n < y \to Zsucc (neg n) \leq y)
+ [ auto
+ (*apply Hcut.
+ assumption*)
+ | elim n
+ [ cut (neg O < y \to Zsucc (neg O) \leq y)
+ [ auto
+ (*apply Hcut.
+ assumption*)
+ | simplify.
+ elim y
+ [ simplify.
+ exact I
+ | simplify.
+ exact I
+ | simplify.
+ (*qui auto non chiude il goal*)
+ apply (not_le_Sn_O n1 H2)
+ ]
+ ]
+ | cut (neg (S n1) < y \to (Zsucc (neg (S n1))) \leq y)
+ [ auto
+ (*apply Hcut.
+ assumption*)
+ | simplify.
+ elim y
+ [ simplify.
+ exact I
+ | simplify.
+ exact I
+ | simplify.
+ (*qui auto non chiude il goal*)
+ apply (le_S_S_to_le n2 n1 H3)
+ ]
+ ]
+ ]
+ ]
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/Z/plus".
+
+include "Z/z.ma".
+include "nat/minus.ma".
+
+definition Zplus :Z \to Z \to Z \def
+\lambda x,y.
+ match x with
+ [ OZ \Rightarrow y
+ | (pos m) \Rightarrow
+ match y with
+ [ OZ \Rightarrow x
+ | (pos n) \Rightarrow (pos (pred ((S m)+(S n))))
+ | (neg n) \Rightarrow
+ match nat_compare m n with
+ [ LT \Rightarrow (neg (pred (n-m)))
+ | EQ \Rightarrow OZ
+ | GT \Rightarrow (pos (pred (m-n)))] ]
+ | (neg m) \Rightarrow
+ match y with
+ [ OZ \Rightarrow x
+ | (pos n) \Rightarrow
+ match nat_compare m n with
+ [ LT \Rightarrow (pos (pred (n-m)))
+ | EQ \Rightarrow OZ
+ | GT \Rightarrow (neg (pred (m-n)))]
+ | (neg n) \Rightarrow (neg (pred ((S m)+(S n))))] ].
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "integer plus" 'plus x y = (cic:/matita/Z/plus/Zplus.con x y).
+
+theorem Zplus_z_OZ: \forall z:Z. z+OZ = z.
+intro.
+elim z;auto.
+ (*simplify;reflexivity.*)
+qed.
+
+(* theorem symmetric_Zplus: symmetric Z Zplus. *)
+
+theorem sym_Zplus : \forall x,y:Z. x+y = y+x.
+intros.
+elim x
+[ auto
+ (*rewrite > Zplus_z_OZ.
+ reflexivity*)
+| elim y
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | simplify.
+ auto
+ (*rewrite < plus_n_Sm.
+ rewrite < plus_n_Sm.
+ rewrite < sym_plus.
+ reflexivity*)
+ | simplify.
+ rewrite > nat_compare_n_m_m_n.
+ simplify.
+ elim nat_compare;auto
+ (*[ simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ ]*)
+ ]
+| elim y
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | simplify.
+ rewrite > nat_compare_n_m_m_n.
+ simplify.
+ elim nat_compare;auto
+ (*[ simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ ]*)
+ | simplify.
+ auto
+ (*rewrite < plus_n_Sm.
+ rewrite < plus_n_Sm.
+ rewrite < sym_plus.
+ reflexivity*)
+ ]
+]
+qed.
+
+theorem Zpred_Zplus_neg_O : \forall z:Z. Zpred z = (neg O)+z.
+intros.
+elim z
+[ auto
+ (*simplify.
+ reflexivity*)
+| elim n;auto
+ (*[ simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ ]*)
+| simplify.
+ reflexivity
+]
+qed.
+
+theorem Zsucc_Zplus_pos_O : \forall z:Z. Zsucc z = (pos O)+z.
+intros.
+elim z
+[ auto
+ (*simplify.
+ reflexivity*)
+| auto
+ (*simplify.
+ reflexivity*)
+| elim n;auto
+ (*[ simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ ]*)
+]
+qed.
+
+theorem Zplus_pos_pos:
+\forall n,m. (pos n)+(pos m) = (Zsucc (pos n))+(Zpred (pos m)).
+intros.
+elim n
+[ elim m;auto
+ (*[ simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ ]*)
+| elim m
+ [ auto
+ (*simplify.
+ rewrite < plus_n_Sm.
+ rewrite < plus_n_O.
+ reflexivity*)
+ | simplify.
+ auto
+ (*rewrite < plus_n_Sm.
+ rewrite < plus_n_Sm.
+ reflexivity*)
+ ]
+]
+qed.
+
+theorem Zplus_pos_neg:
+\forall n,m. (pos n)+(neg m) = (Zsucc (pos n))+(Zpred (neg m)).
+intros.
+reflexivity.
+qed.
+
+theorem Zplus_neg_pos :
+\forall n,m. (neg n)+(pos m) = (Zsucc (neg n))+(Zpred (pos m)).
+intros.
+elim n
+[ elim m;auto
+ (*[ simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ ]*)
+| elim m;auto
+ (*[ simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ ]*)
+]
+qed.
+
+theorem Zplus_neg_neg:
+\forall n,m. (neg n)+(neg m) = (Zsucc (neg n))+(Zpred (neg m)).
+intros.
+elim n
+[ auto
+ (*elim m
+ [ simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ ]*)
+| elim m
+ [ auto
+ (*simplify.
+ rewrite > plus_n_Sm.
+ reflexivity*)
+ | simplify.
+ auto
+ (*rewrite > plus_n_Sm.
+ reflexivity*)
+ ]
+]
+qed.
+
+theorem Zplus_Zsucc_Zpred:
+\forall x,y. x+y = (Zsucc x)+(Zpred y).
+intros.
+elim x
+[ auto
+ (*elim y
+ [ simplify.
+ reflexivity
+ | rewrite < Zsucc_Zplus_pos_O.
+ rewrite > Zsucc_Zpred.
+ reflexivity
+ | simplify.
+ reflexivity
+ ]*)
+| elim y;auto
+ (*[ simplify.
+ reflexivity
+ | apply Zplus_pos_pos
+ | apply Zplus_pos_neg
+ ]*)
+| elim y;auto
+ (*[ rewrite < sym_Zplus.
+ rewrite < (sym_Zplus (Zpred OZ)).
+ rewrite < Zpred_Zplus_neg_O.
+ rewrite > Zpred_Zsucc.
+ simplify.
+ reflexivity
+ | apply Zplus_neg_pos
+ | rewrite < Zplus_neg_neg.
+ reflexivity
+ ]*)
+]
+qed.
+
+theorem Zplus_Zsucc_pos_pos :
+\forall n,m. (Zsucc (pos n))+(pos m) = Zsucc ((pos n)+(pos m)).
+intros.
+reflexivity.
+qed.
+
+theorem Zplus_Zsucc_pos_neg:
+\forall n,m. (Zsucc (pos n))+(neg m) = (Zsucc ((pos n)+(neg m))).
+intros.
+apply (nat_elim2
+(\lambda n,m. (Zsucc (pos n))+(neg m) = (Zsucc ((pos n)+(neg m)))))
+[ intro.
+ elim n1;auto
+ (*[ simplify.
+ reflexivity
+ | elim n2; simplify; reflexivity
+ ]*)
+| intros.
+ auto
+ (*elim n1;simplify;reflexivity*)
+| intros.
+ rewrite < (Zplus_pos_neg ? m1).
+ elim H.
+ reflexivity
+]
+qed.
+
+theorem Zplus_Zsucc_neg_neg :
+\forall n,m. Zsucc (neg n) + neg m = Zsucc (neg n + neg m).
+intros.
+apply (nat_elim2
+(\lambda n,m. Zsucc (neg n) + neg m = Zsucc (neg n + neg m)))
+[ intros.
+ auto
+ (*elim n1
+ [ simplify.
+ reflexivity
+ | elim n2;simplify;reflexivity
+ ]*)
+| intros.
+ auto
+ (*elim n1;simplify;reflexivity*)
+| intros.
+ auto.
+ (*rewrite < (Zplus_neg_neg ? m1).
+ reflexivity*)
+]
+qed.
+
+theorem Zplus_Zsucc_neg_pos:
+\forall n,m. Zsucc (neg n)+(pos m) = Zsucc ((neg n)+(pos m)).
+intros.
+apply (nat_elim2
+(\lambda n,m. Zsucc (neg n) + (pos m) = Zsucc (neg n + pos m)))
+[ intros.
+ auto
+ (*elim n1
+ [ simplify.
+ reflexivity
+ | elim n2;simplify;reflexivity
+ ]*)
+| intros.
+ auto
+ (*elim n1;simplify;reflexivity*)
+| intros.
+ auto
+ (*rewrite < H.
+ rewrite < (Zplus_neg_pos ? (S m1)).
+ reflexivity*)
+]
+qed.
+
+theorem Zplus_Zsucc : \forall x,y:Z. (Zsucc x)+y = Zsucc (x+y).
+intros.
+elim x
+[ auto
+ (*elim y
+ [ simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ | rewrite < Zsucc_Zplus_pos_O.
+ reflexivity
+ ]*)
+| elim y;auto
+ (*[ rewrite < (sym_Zplus OZ).
+ reflexivity
+ | apply Zplus_Zsucc_pos_pos
+ | apply Zplus_Zsucc_pos_neg
+ ]*)
+| elim y;auto
+ (*[ rewrite < sym_Zplus.
+ rewrite < (sym_Zplus OZ).
+ simplify.
+ reflexivity
+ | apply Zplus_Zsucc_neg_pos
+ | apply Zplus_Zsucc_neg_neg
+ ]*)
+]
+qed.
+
+theorem Zplus_Zpred: \forall x,y:Z. (Zpred x)+y = Zpred (x+y).
+intros.
+cut (Zpred (x+y) = Zpred ((Zsucc (Zpred x))+y));auto.
+(*[ rewrite > Hcut.
+ rewrite > Zplus_Zsucc.
+ rewrite > Zpred_Zsucc.
+ reflexivity
+| rewrite > Zsucc_Zpred.
+ reflexivity
+]*)
+qed.
+
+
+theorem associative_Zplus: associative Z Zplus.
+change with (\forall x,y,z:Z. (x + y) + z = x + (y + z)).
+(* simplify. *)
+intros.
+elim x
+[ auto
+ (*simplify.
+ reflexivity*)
+| elim n
+ [ rewrite < Zsucc_Zplus_pos_O.
+ auto
+ (*rewrite < Zsucc_Zplus_pos_O.
+ rewrite > Zplus_Zsucc.
+ reflexivity*)
+ | rewrite > (Zplus_Zsucc (pos n1)).
+ rewrite > (Zplus_Zsucc (pos n1)).
+ auto
+ (*rewrite > (Zplus_Zsucc ((pos n1)+y)).
+ apply eq_f.
+ assumption*)
+ ]
+| elim n
+ [ rewrite < (Zpred_Zplus_neg_O (y+z)).
+ auto
+ (*rewrite < (Zpred_Zplus_neg_O y).
+ rewrite < Zplus_Zpred.
+ reflexivity*)
+ | rewrite > (Zplus_Zpred (neg n1)).
+ rewrite > (Zplus_Zpred (neg n1)).
+ auto
+ (*rewrite > (Zplus_Zpred ((neg n1)+y)).
+ apply eq_f.
+ assumption*)
+ ]
+]
+qed.
+
+variant assoc_Zplus : \forall x,y,z:Z. (x+y)+z = x+(y+z)
+\def associative_Zplus.
+
+(* Zopp *)
+definition Zopp : Z \to Z \def
+\lambda x:Z. match x with
+[ OZ \Rightarrow OZ
+| (pos n) \Rightarrow (neg n)
+| (neg n) \Rightarrow (pos n) ].
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "integer unary minus" 'uminus x = (cic:/matita/Z/plus/Zopp.con x).
+
+theorem Zopp_Zplus: \forall x,y:Z. -(x+y) = -x + -y.
+intros.
+elim x
+[ elim y;auto
+ (*simplify;reflexivity*)
+| elim y
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | auto
+ (*simplify.
+ reflexivity*)
+ | simplify.
+ apply nat_compare_elim;
+ intro;auto (*simplify;reflexivity*)
+ ]
+| elim y
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | simplify.
+ apply nat_compare_elim;
+ intro;auto
+ (*simplify;reflexivity*)
+ | auto
+ (*simplify.
+ reflexivity*)
+ ]
+]
+qed.
+
+theorem Zopp_Zopp: \forall x:Z. --x = x.
+intro.
+elim x;reflexivity.
+qed.
+
+theorem Zplus_Zopp: \forall x:Z. x+ -x = OZ.
+intro.
+elim x
+[ apply refl_eq
+| simplify.
+ rewrite > nat_compare_n_n.
+ auto
+ (*simplify.
+ apply refl_eq*)
+| simplify.
+ rewrite > nat_compare_n_n.
+ auto
+ (*simplify.
+ apply refl_eq*)
+]
+qed.
+
+(* minus *)
+definition Zminus : Z \to Z \to Z \def \lambda x,y:Z. x + (-y).
+
+interpretation "integer minus" 'minus x y = (cic:/matita/Z/plus/Zminus.con x y).
--- /dev/null
+(**************************************************************************)
+(* __ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/Z/times".
+
+include "nat/lt_arith.ma".
+include "Z/plus.ma".
+
+definition Ztimes :Z \to Z \to Z \def
+\lambda x,y.
+ match x with
+ [ OZ \Rightarrow OZ
+ | (pos m) \Rightarrow
+ match y with
+ [ OZ \Rightarrow OZ
+ | (pos n) \Rightarrow (pos (pred ((S m) * (S n))))
+ | (neg n) \Rightarrow (neg (pred ((S m) * (S n))))]
+ | (neg m) \Rightarrow
+ match y with
+ [ OZ \Rightarrow OZ
+ | (pos n) \Rightarrow (neg (pred ((S m) * (S n))))
+ | (neg n) \Rightarrow (pos (pred ((S m) * (S n))))]].
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "integer times" 'times x y = (cic:/matita/Z/times/Ztimes.con x y).
+
+theorem Ztimes_z_OZ: \forall z:Z. z*OZ = OZ.
+intro.
+elim z;auto.
+ (*simplify;reflexivity.*)
+qed.
+
+theorem Ztimes_neg_Zopp: \forall n:nat.\forall x:Z.
+neg n * x = - (pos n * x).
+intros.
+elim x;auto.
+ (*simplify;reflexivity.*)
+qed.
+
+theorem symmetric_Ztimes : symmetric Z Ztimes.
+change with (\forall x,y:Z. x*y = y*x).
+intros.
+elim x
+[ auto
+ (*rewrite > Ztimes_z_OZ.
+ reflexivity*)
+| elim y
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | change with (pos (pred ((S n) * (S n1))) = pos (pred ((S n1) * (S n)))).
+ auto
+ (*rewrite < sym_times.
+ reflexivity*)
+ | change with (neg (pred ((S n) * (S n1))) = neg (pred ((S n1) * (S n)))).
+ auto
+ (*rewrite < sym_times.
+ reflexivity*)
+ ]
+| elim y
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | change with (neg (pred ((S n) * (S n1))) = neg (pred ((S n1) * (S n)))).
+ auto
+ (*rewrite < sym_times.
+ reflexivity*)
+ | change with (pos (pred ((S n) * (S n1))) = pos (pred ((S n1) * (S n)))).
+ auto
+ (*rewrite < sym_times.
+ reflexivity*)
+ ]
+]
+qed.
+
+variant sym_Ztimes : \forall x,y:Z. x*y = y*x
+\def symmetric_Ztimes.
+
+theorem associative_Ztimes: associative Z Ztimes.
+unfold associative.
+intros.
+elim x
+[ auto
+ (*simplify.
+ reflexivity*)
+| elim y
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | elim z
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | change with
+ (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
+ pos (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
+ rewrite < S_pred
+ [ rewrite < S_pred;auto
+ (*[ rewrite < assoc_times.
+ reflexivity
+ | apply lt_O_times_S_S
+ ]*)
+ | apply lt_O_times_S_S
+ ]
+ | change with
+ (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
+ neg (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
+ rewrite < S_pred
+ [ rewrite < S_pred;auto
+ (*[ rewrite < assoc_times.
+ reflexivity
+ | apply lt_O_times_S_S
+ ]*)
+ | apply lt_O_times_S_S
+ ]
+ ]
+ | elim z
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | change with
+ (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
+ neg (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
+ rewrite < S_pred
+ [ rewrite < S_pred;auto
+ (*[ rewrite < assoc_times.
+ reflexivity
+ | apply lt_O_times_S_S
+ ]*)
+ | apply lt_O_times_S_S
+ ]
+ | change with
+ (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
+ pos(pred ((S n) * (S (pred ((S n1) * (S n2))))))).
+ rewrite < S_pred
+ [ rewrite < S_pred;auto
+ (*[ rewrite < assoc_times.
+ reflexivity
+ | apply lt_O_times_S_S
+ ]*)
+ | apply lt_O_times_S_S
+ ]
+ ]
+ ]
+| elim y
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | elim z
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | change with
+ (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
+ neg (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
+ rewrite < S_pred
+ [ rewrite < S_pred;auto
+ (*[ rewrite < assoc_times.
+ reflexivity
+ | apply lt_O_times_S_S
+ ]*)
+ | apply lt_O_times_S_S
+ ]
+ | change with
+ (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
+ pos (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
+ rewrite < S_pred
+ [ rewrite < S_pred;auto
+ (*[ rewrite < assoc_times.
+ reflexivity
+ | apply lt_O_times_S_S
+ ]*)
+ | apply lt_O_times_S_S
+ ]
+ ]
+ | elim z
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | change with
+ (pos (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
+ pos (pred ((S n) * (S (pred ((S n1) * (S n2))))))).
+ rewrite < S_pred
+ [ rewrite < S_pred;auto
+ (*[ rewrite < assoc_times.
+ reflexivity
+ | apply lt_O_times_S_S
+ ]*)
+ | apply lt_O_times_S_S
+ ]
+ | change with
+ (neg (pred ((S (pred ((S n) * (S n1)))) * (S n2))) =
+ neg(pred ((S n) * (S (pred ((S n1) * (S n2))))))).
+ rewrite < S_pred
+ [ rewrite < S_pred;auto
+ (*[ rewrite < assoc_times.
+ reflexivity
+ | apply lt_O_times_S_S
+ ]*)
+ | apply lt_O_times_S_S
+ ]
+ ]
+ ]
+]
+qed.
+
+variant assoc_Ztimes : \forall x,y,z:Z.
+(x * y) * z = x * (y * z) \def
+associative_Ztimes.
+
+lemma times_minus1: \forall n,p,q:nat. lt q p \to
+(S n) * (S (pred ((S p) - (S q)))) =
+pred ((S n) * (S p)) - pred ((S n) * (S q)).
+intros.
+rewrite < S_pred
+[ rewrite > minus_pred_pred
+ [ auto
+ (*rewrite < distr_times_minus.
+ reflexivity*)
+
+ (* we now close all positivity conditions *)
+ | apply lt_O_times_S_S
+ | apply lt_O_times_S_S
+ ]
+| auto
+ (*simplify.
+ unfold lt.
+ apply le_SO_minus.
+ exact H*)
+]
+qed.
+
+lemma Ztimes_Zplus_pos_neg_pos: \forall n,p,q:nat.
+(pos n)*((neg p)+(pos q)) = (pos n)*(neg p)+ (pos n)*(pos q).
+intros.
+simplify.
+change in match (p + n * (S p)) with (pred ((S n) * (S p))).
+change in match (q + n * (S q)) with (pred ((S n) * (S q))).
+rewrite < nat_compare_pred_pred
+[ rewrite < nat_compare_times_l.
+ rewrite < nat_compare_S_S.
+ apply (nat_compare_elim p q)
+ [ intro.
+ (* uff *)
+ change with (pos (pred ((S n) * (S (pred ((S q) - (S p)))))) =
+ pos (pred ((pred ((S n) * (S q))) - (pred ((S n) * (S p)))))).
+ rewrite < (times_minus1 n q p H).
+ reflexivity
+ | intro.
+ auto
+ (*rewrite < H.
+ simplify.
+ reflexivity*)
+ | intro.
+ change with (neg (pred ((S n) * (S (pred ((S p) - (S q)))))) =
+ neg (pred ((pred ((S n) * (S p))) - (pred ((S n) * (S q)))))).
+ rewrite < (times_minus1 n p q H).
+ reflexivity
+ ]
+ (* two more positivity conditions from nat_compare_pred_pred *)
+
+| apply lt_O_times_S_S
+| apply lt_O_times_S_S
+]
+qed.
+
+lemma Ztimes_Zplus_pos_pos_neg: \forall n,p,q:nat.
+(pos n)*((pos p)+(neg q)) = (pos n)*(pos p)+ (pos n)*(neg q).
+intros.
+auto.
+(*rewrite < sym_Zplus.
+rewrite > Ztimes_Zplus_pos_neg_pos.
+apply sym_Zplus.*)
+qed.
+
+lemma distributive2_Ztimes_pos_Zplus:
+distributive2 nat Z (\lambda n,z. (pos n) * z) Zplus.
+change with (\forall n,y,z.
+(pos n) * (y + z) = (pos n) * y + (pos n) * z).
+intros.
+elim y
+[ reflexivity
+| elim z
+ [ reflexivity
+ | change with
+ (pos (pred ((S n) * ((S n1) + (S n2)))) =
+ pos (pred ((S n) * (S n1) + (S n) * (S n2)))).
+ auto
+ (*rewrite < distr_times_plus.
+ reflexivity*)
+ | apply Ztimes_Zplus_pos_pos_neg
+ ]
+| elim z
+ [ reflexivity
+ | apply Ztimes_Zplus_pos_neg_pos
+ | change with
+ (neg (pred ((S n) * ((S n1) + (S n2)))) =
+ neg (pred ((S n) * (S n1) + (S n) * (S n2)))).
+ auto
+ (*rewrite < distr_times_plus.
+ reflexivity*)
+ ]
+]
+qed.
+
+variant distr_Ztimes_Zplus_pos: \forall n,y,z.
+(pos n) * (y + z) = ((pos n) * y + (pos n) * z) \def
+distributive2_Ztimes_pos_Zplus.
+
+lemma distributive2_Ztimes_neg_Zplus :
+distributive2 nat Z (\lambda n,z. (neg n) * z) Zplus.
+change with (\forall n,y,z.
+(neg n) * (y + z) = (neg n) * y + (neg n) * z).
+intros.
+rewrite > Ztimes_neg_Zopp.
+rewrite > distr_Ztimes_Zplus_pos.
+auto.
+(*rewrite > Zopp_Zplus.
+rewrite < Ztimes_neg_Zopp.
+rewrite < Ztimes_neg_Zopp.
+reflexivity.*)
+qed.
+
+variant distr_Ztimes_Zplus_neg: \forall n,y,z.
+(neg n) * (y + z) = (neg n) * y + (neg n) * z \def
+distributive2_Ztimes_neg_Zplus.
+
+theorem distributive_Ztimes_Zplus: distributive Z Ztimes Zplus.
+change with (\forall x,y,z:Z. x * (y + z) = x*y + x*z).
+intros.
+elim x;auto.
+(*[ simplify.
+ reflexivity
+| apply distr_Ztimes_Zplus_pos
+| apply distr_Ztimes_Zplus_neg
+]*)
+qed.
+
+variant distr_Ztimes_Zplus: \forall x,y,z.
+x * (y + z) = x*y + x*z \def
+distributive_Ztimes_Zplus.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/Z/z".
+
+include "datatypes/bool.ma".
+include "nat/nat.ma".
+
+inductive Z : Set \def
+ OZ : Z
+| pos : nat \to Z
+| neg : nat \to Z.
+
+definition Z_of_nat \def
+\lambda n. match n with
+[ O \Rightarrow OZ
+| (S n)\Rightarrow pos n].
+
+coercion cic:/matita/Z/z/Z_of_nat.con.
+
+definition neg_Z_of_nat \def
+\lambda n. match n with
+[ O \Rightarrow OZ
+| (S n)\Rightarrow neg n].
+
+definition abs \def
+\lambda z.
+ match z with
+[ OZ \Rightarrow O
+| (pos n) \Rightarrow (S n)
+| (neg n) \Rightarrow (S n)].
+
+definition OZ_test \def
+\lambda z.
+match z with
+[ OZ \Rightarrow true
+| (pos n) \Rightarrow false
+| (neg n) \Rightarrow false].
+
+theorem OZ_test_to_Prop :\forall z:Z.
+match OZ_test z with
+[true \Rightarrow z=OZ
+|false \Rightarrow z \neq OZ].
+intros.
+elim z
+[ (*qui auto non chiude il goal*)
+ simplify.
+ reflexivity
+| simplify.
+ unfold Not.
+ intros (H).
+ (*qui auto non chiude il goal*)
+ destruct H
+| simplify.
+ unfold Not.
+ intros (H).
+ (*qui auto non chiude il goal*)
+ destruct H
+]
+qed.
+
+(* discrimination *)
+theorem injective_pos: injective nat Z pos.
+unfold injective.
+intros.
+apply inj_S.
+change with (abs (pos x) = abs (pos y)).
+auto.
+(*apply eq_f.
+assumption.*)
+qed.
+
+variant inj_pos : \forall n,m:nat. pos n = pos m \to n = m
+\def injective_pos.
+
+theorem injective_neg: injective nat Z neg.
+unfold injective.
+intros.
+apply inj_S.
+change with (abs (neg x) = abs (neg y)).
+auto.
+(*apply eq_f.
+assumption.*)
+qed.
+
+variant inj_neg : \forall n,m:nat. neg n = neg m \to n = m
+\def injective_neg.
+
+theorem not_eq_OZ_pos: \forall n:nat. OZ \neq pos n.
+unfold Not.
+intros (n H).
+(*qui auto non chiude il goal*)
+destruct H.
+qed.
+
+theorem not_eq_OZ_neg :\forall n:nat. OZ \neq neg n.
+unfold Not.
+intros (n H).
+(*qui auto non chiude il goal*)
+destruct H.
+qed.
+
+theorem not_eq_pos_neg :\forall n,m:nat. pos n \neq neg m.
+unfold Not.
+intros (n m H).
+(*qui auto non chiude il goal*)
+destruct H.
+qed.
+
+theorem decidable_eq_Z : \forall x,y:Z. decidable (x=y).
+intros.
+unfold decidable.
+elim x
+[ (* goal: x=OZ *)
+ elim y
+ [ (* goal: x=OZ y=OZ *)
+ auto
+ (*left.
+ reflexivity*)
+ | (* goal: x=OZ 2=2 *)
+ auto
+ (*right.
+ apply not_eq_OZ_pos*)
+ | (* goal: x=OZ 2=3 *)
+ auto
+ (*right.
+ apply not_eq_OZ_neg*)
+ ]
+| (* goal: x=pos *)
+ elim y
+ [ (* goal: x=pos y=OZ *)
+ right.
+ unfold Not.
+ intro.
+ apply (not_eq_OZ_pos n).
+ auto
+ (*symmetry.
+ assumption*)
+ | (* goal: x=pos y=pos *)
+ elim (decidable_eq_nat n n1:((n=n1) \lor ((n=n1) \to False)))
+ [ auto
+ (*left.
+ apply eq_f.
+ assumption*)
+ | right.
+ unfold Not.
+ intros (H_inj).
+ auto
+ (*apply H.
+ destruct H_inj.
+ assumption*)
+ ]
+ | (* goal: x=pos y=neg *)
+ auto
+ (*right.
+ unfold Not.
+ intro.
+ apply (not_eq_pos_neg n n1).
+ assumption*)
+ ]
+| (* goal: x=neg *)
+ elim y
+ [ (* goal: x=neg y=OZ *)
+ right.
+ unfold Not.
+ intro.
+ apply (not_eq_OZ_neg n).
+ auto
+ (*symmetry.
+ assumption*)
+ | (* goal: x=neg y=pos *)
+ right.
+ unfold Not.
+ intro.
+ apply (not_eq_pos_neg n1 n).
+ auto
+ (*symmetry.
+ assumption*)
+ | (* goal: x=neg y=neg *)
+ elim (decidable_eq_nat n n1:((n=n1) \lor ((n=n1) \to False)))
+ [ auto
+ (*left.
+ apply eq_f.
+ assumption*)
+ | right.
+ unfold Not.
+ intro.
+ auto
+ (*apply H.
+ apply injective_neg.
+ assumption*)
+ ]
+ ]
+]
+qed.
+
+(* end discrimination *)
+
+definition Zsucc \def
+\lambda z. match z with
+[ OZ \Rightarrow pos O
+| (pos n) \Rightarrow pos (S n)
+| (neg n) \Rightarrow
+ match n with
+ [ O \Rightarrow OZ
+ | (S p) \Rightarrow neg p]].
+
+definition Zpred \def
+\lambda z. match z with
+[ OZ \Rightarrow neg O
+| (pos n) \Rightarrow
+ match n with
+ [ O \Rightarrow OZ
+ | (S p) \Rightarrow pos p]
+| (neg n) \Rightarrow neg (S n)].
+
+theorem Zpred_Zsucc: \forall z:Z. Zpred (Zsucc z) = z.
+intros.
+elim z
+[ reflexivity
+| reflexivity
+| elim n
+ [ reflexivity
+ | reflexivity
+ ]
+]
+qed.
+
+theorem Zsucc_Zpred: \forall z:Z. Zsucc (Zpred z) = z.
+intros.
+elim z
+[ reflexivity
+| elim n
+ [ reflexivity
+ | reflexivity
+ ]
+| reflexivity
+]
+qed.
+
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/chinese_reminder".
+
+include "nat/exp.ma".
+include "nat/gcd.ma".
+include "nat/permutation.ma".
+include "nat/congruence.ma".
+
+theorem and_congruent_congruent: \forall m,n,a,b:nat. O < n \to O < m \to
+gcd n m = (S O) \to ex nat (\lambda x. congruent x a m \land congruent x b n).
+intros.
+cut (\exists c,d.c*n - d*m = (S O) \lor d*m - c*n = (S O))
+[ elim Hcut.
+ elim H3.
+ elim H4
+ [
+ apply (ex_intro nat ? ((a+b*m)*a1*n-b*a2*m)).
+ split
+ [ (* congruent to a *)
+ cut (a1*n = a2*m + (S O))
+ [ rewrite > assoc_times.
+ rewrite > Hcut1.
+ rewrite < (sym_plus ? (a2*m)).
+ rewrite > distr_times_plus.
+ rewrite < times_n_SO.
+ rewrite > assoc_plus.
+ rewrite < assoc_times.
+ rewrite < times_plus_l.
+ rewrite > eq_minus_plus_plus_minus
+ [ auto
+ (*rewrite < times_minus_l.
+ rewrite > sym_plus.
+ apply (eq_times_plus_to_congruent ? ? ? ((b+(a+b*m)*a2)-b*a2))
+ [ assumption
+ | reflexivity
+ ]*)
+ | apply le_times_l.
+ apply (trans_le ? ((a+b*m)*a2))
+ [ apply le_times_l.
+ apply (trans_le ? (b*m));auto
+ (*[ rewrite > times_n_SO in \vdash (? % ?).
+ apply le_times_r.
+ assumption
+ | apply le_plus_n
+ ]*)
+ | apply le_plus_n
+ ]
+ ]
+ | apply minus_to_plus
+ [ apply lt_to_le.
+ change with (O + a2*m < a1*n).
+ apply lt_minus_to_plus.
+ auto
+ (*rewrite > H5.
+ unfold lt.
+ apply le_n*)
+ | assumption
+ ]
+ ]
+ | (* congruent to b *)
+ cut (a2*m = a1*n - (S O))
+ [ rewrite > (assoc_times b a2).
+ rewrite > Hcut1.
+ rewrite > distr_times_minus.
+ rewrite < assoc_times.
+ rewrite < eq_plus_minus_minus_minus
+ [ auto
+ (*rewrite < times_n_SO.
+ rewrite < times_minus_l.
+ rewrite < sym_plus.
+ apply (eq_times_plus_to_congruent ? ? ? ((a+b*m)*a1-b*a1))
+ [ assumption
+ | reflexivity
+ ]*)
+ | rewrite > assoc_times.
+ apply le_times_r.
+ (*auto genera un'esecuzione troppo lunga*)
+ apply (trans_le ? (a1*n - a2*m));auto
+ (*[ rewrite > H5.
+ apply le_n
+ | apply (le_minus_m ? (a2*m))
+ ]*)
+ | apply le_times_l.
+ apply le_times_l.
+ auto
+ (*apply (trans_le ? (b*m))
+ [ rewrite > times_n_SO in \vdash (? % ?).
+ apply le_times_r.
+ assumption
+ | apply le_plus_n
+ ]*)
+ ]
+ | apply sym_eq.
+ apply plus_to_minus.
+ rewrite > sym_plus.
+ apply minus_to_plus
+ [ apply lt_to_le.
+ change with (O + a2*m < a1*n).
+ apply lt_minus_to_plus.
+ auto
+ (*rewrite > H5.
+ unfold lt.
+ apply le_n*)
+ | assumption
+ ]
+ ]
+ ]
+ | (* and now the symmetric case; the price to pay for working
+ in nat instead than Z *)
+ apply (ex_intro nat ? ((b+a*n)*a2*m-a*a1*n)).
+ split
+ [(* congruent to a *)
+ cut (a1*n = a2*m - (S O))
+ [ rewrite > (assoc_times a a1).
+ rewrite > Hcut1.
+ rewrite > distr_times_minus.
+ rewrite < assoc_times.
+ rewrite < eq_plus_minus_minus_minus
+ [ auto
+ (*rewrite < times_n_SO.
+ rewrite < times_minus_l.
+ rewrite < sym_plus.
+ apply (eq_times_plus_to_congruent ? ? ? ((b+a*n)*a2-a*a2))
+ [ assumption
+ | reflexivity
+ ]*)
+ | rewrite > assoc_times.
+ apply le_times_r.
+ apply (trans_le ? (a2*m - a1*n));auto
+ (*[ rewrite > H5.
+ apply le_n
+ | apply (le_minus_m ? (a1*n))
+ ]*)
+ | rewrite > assoc_times.
+ rewrite > assoc_times.
+ apply le_times_l.
+ auto
+ (*apply (trans_le ? (a*n))
+ [ rewrite > times_n_SO in \vdash (? % ?).
+ apply le_times_r.
+ assumption.
+ | apply le_plus_n.
+ ]*)
+ ]
+ | apply sym_eq.
+ apply plus_to_minus.
+ rewrite > sym_plus.
+ apply minus_to_plus
+ [ apply lt_to_le.
+ change with (O + a1*n < a2*m).
+ apply lt_minus_to_plus.
+ auto
+ (*rewrite > H5.
+ unfold lt.
+ apply le_n*)
+ | assumption
+ ]
+ ]
+ | (* congruent to b *)
+ cut (a2*m = a1*n + (S O))
+ [ rewrite > assoc_times.
+ rewrite > Hcut1.
+ rewrite > (sym_plus (a1*n)).
+ rewrite > distr_times_plus.
+ rewrite < times_n_SO.
+ rewrite < assoc_times.
+ rewrite > assoc_plus.
+ rewrite < times_plus_l.
+ rewrite > eq_minus_plus_plus_minus
+ [ auto
+ (*rewrite < times_minus_l.
+ rewrite > sym_plus.
+ apply (eq_times_plus_to_congruent ? ? ? ((a+(b+a*n)*a1)-a*a1))
+ [ assumption
+ | reflexivity
+ ]*)
+ | apply le_times_l.
+ apply (trans_le ? ((b+a*n)*a1))
+ [ apply le_times_l.
+ auto
+ (*apply (trans_le ? (a*n))
+ [ rewrite > times_n_SO in \vdash (? % ?).
+ apply le_times_r.
+ assumption
+ | apply le_plus_n
+ ]*)
+ | apply le_plus_n
+ ]
+ ]
+ | apply minus_to_plus
+ [ apply lt_to_le.
+ change with (O + a1*n < a2*m).
+ apply lt_minus_to_plus.
+ auto
+ (*rewrite > H5.
+ unfold lt.
+ apply le_n*)
+ | assumption
+ ]
+ ]
+ ]
+ ]
+(* proof of the cut *)
+| (* qui auto non conclude il goal *)
+ rewrite < H2.
+ apply eq_minus_gcd
+]
+qed.
+
+theorem and_congruent_congruent_lt: \forall m,n,a,b:nat. O < n \to O < m \to
+gcd n m = (S O) \to
+ex nat (\lambda x. (congruent x a m \land congruent x b n) \land
+ (x < m*n)).
+intros.
+elim (and_congruent_congruent m n a b)
+[ elim H3.
+ apply (ex_intro ? ? (a1 \mod (m*n))).
+ split
+ [ split
+ [ apply (transitive_congruent m ? a1)
+ [ unfold congruent.
+ apply sym_eq.
+ change with (congruent a1 (a1 \mod (m*n)) m).
+ rewrite < sym_times.
+ auto
+ (*apply congruent_n_mod_times;assumption*)
+ | assumption
+ ]
+ | apply (transitive_congruent n ? a1)
+ [ unfold congruent.
+ apply sym_eq.
+ change with (congruent a1 (a1 \mod (m*n)) n).
+ auto
+ (*apply congruent_n_mod_times;assumption*)
+ | assumption
+ ]
+ ]
+ | apply lt_mod_m_m.
+ auto
+ (*rewrite > (times_n_O O).
+ apply lt_times;assumption*)
+ ]
+| assumption
+| assumption
+| assumption
+]
+qed.
+
+definition cr_pair : nat \to nat \to nat \to nat \to nat \def
+\lambda n,m,a,b.
+min (pred (n*m)) (\lambda x. andb (eqb (x \mod n) a) (eqb (x \mod m) b)).
+
+theorem cr_pair1: cr_pair (S (S O)) (S (S (S O))) O O = O.
+reflexivity.
+qed.
+
+theorem cr_pair2: cr_pair (S(S O)) (S(S(S O))) (S O) O = (S(S(S O))).
+auto.
+(*simplify.
+reflexivity.*)
+qed.
+
+theorem cr_pair3: cr_pair (S(S O)) (S(S(S O))) (S O) (S(S O)) = (S(S(S(S(S O))))).
+reflexivity.
+qed.
+
+theorem cr_pair4: cr_pair (S(S(S(S(S O))))) (S(S(S(S(S(S(S O))))))) (S(S(S O))) (S(S O)) =
+(S(S(S(S(S(S(S(S(S(S(S(S(S(S(S(S(S(S(S(S(S(S(S O))))))))))))))))))))))).
+reflexivity.
+qed.
+
+theorem mod_cr_pair : \forall m,n,a,b. a \lt m \to b \lt n \to
+gcd n m = (S O) \to
+(cr_pair m n a b) \mod m = a \land (cr_pair m n a b) \mod n = b.
+intros.
+cut (andb (eqb ((cr_pair m n a b) \mod m) a)
+ (eqb ((cr_pair m n a b) \mod n) b) = true)
+[ generalize in match Hcut.
+ apply andb_elim.
+ apply eqb_elim;intro
+ [ rewrite > H3.
+ simplify.
+ intro.
+ auto
+ (*split
+ [ reflexivity
+ | apply eqb_true_to_eq.
+ assumption
+ ]*)
+ | simplify.
+ intro.
+ (* l'invocazione di auto qui genera segmentation fault *)
+ apply False_ind.
+ (* l'invocazione di auto qui genera segmentation fault *)
+ apply not_eq_true_false.
+ (* l'invocazione di auto qui genera segmentation fault *)
+ apply sym_eq.
+ assumption
+ ]
+| apply (f_min_aux_true
+ (\lambda x. andb (eqb (x \mod m) a) (eqb (x \mod n) b)) (pred (m*n)) (pred (m*n))).
+ elim (and_congruent_congruent_lt m n a b)
+ [ apply (ex_intro ? ? a1).
+ split
+ [ split
+ [ auto
+ (*rewrite < minus_n_n.
+ apply le_O_n*)
+ | elim H3.
+ apply le_S_S_to_le.
+ apply (trans_le ? (m*n))
+ [ assumption
+ | apply (nat_case (m*n))
+ [ apply le_O_n
+ | intro.
+ auto
+ (*rewrite < pred_Sn.
+ apply le_n*)
+ ]
+ ]
+ ]
+ | elim H3.
+ elim H4.
+ apply andb_elim.
+ cut (a1 \mod m = a)
+ [ cut (a1 \mod n = b)
+ [ rewrite > (eq_to_eqb_true ? ? Hcut).
+ rewrite > (eq_to_eqb_true ? ? Hcut1).
+ (* l'invocazione di auto qui non chiude il goal *)
+ simplify.
+ reflexivity
+ | rewrite < (lt_to_eq_mod b n);
+ assumption
+ ]
+ | rewrite < (lt_to_eq_mod a m);assumption
+ ]
+ ]
+ | auto
+ (*apply (le_to_lt_to_lt ? b)
+ [ apply le_O_n
+ | assumption
+ ]*)
+ | auto
+ (*apply (le_to_lt_to_lt ? a)
+ [ apply le_O_n
+ | assumption
+ ]*)
+ | assumption
+ ]
+]
+qed.
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/compare".
+
+include "datatypes/bool.ma".
+include "datatypes/compare.ma".
+include "nat/orders.ma".
+
+let rec eqb n m \def
+match n with
+ [ O \Rightarrow
+ match m with
+ [ O \Rightarrow true
+ | (S q) \Rightarrow false]
+ | (S p) \Rightarrow
+ match m with
+ [ O \Rightarrow false
+ | (S q) \Rightarrow eqb p q]].
+
+theorem eqb_to_Prop: \forall n,m:nat.
+match (eqb n m) with
+[ true \Rightarrow n = m
+| false \Rightarrow n \neq m].
+intros.
+apply (nat_elim2
+(\lambda n,m:nat.match (eqb n m) with
+[ true \Rightarrow n = m
+| false \Rightarrow n \neq m]))
+[ intro.
+ elim n1;simplify;auto
+ (*[ simplify
+ reflexivity
+ | simplify.
+ apply not_eq_O_S
+ ]*)
+| intro.
+ simplify.
+ unfold Not.
+ intro.
+ apply (not_eq_O_S n1).
+ auto
+ (*apply sym_eq.
+ assumption*)
+| intros.
+ simplify.
+ generalize in match H.
+ elim ((eqb n1 m1));simplify
+ [ apply eq_f.
+ apply H1
+ | unfold Not.
+ intro.
+ apply H1.
+ auto
+ (*apply inj_S.
+ assumption*)
+ ]
+]
+qed.
+
+theorem eqb_elim : \forall n,m:nat.\forall P:bool \to Prop.
+(n=m \to (P true)) \to (n \neq m \to (P false)) \to (P (eqb n m)).
+intros.
+cut
+(match (eqb n m) with
+[ true \Rightarrow n = m
+| false \Rightarrow n \neq m] \to (P (eqb n m)))
+[ apply Hcut.
+ (* qui auto non conclude il goal*)
+ apply eqb_to_Prop
+| elim (eqb n m)
+ [ (*qui auto non conclude il goal*)
+ apply ((H H2))
+ | (*qui auto non conclude il goal*)
+ apply ((H1 H2))
+ ]
+]
+qed.
+
+theorem eqb_n_n: \forall n. eqb n n = true.
+intro.
+elim n;simplify;auto.
+(*[ simplify.reflexivity
+| simplify.assumption.
+]*)
+qed.
+
+theorem eqb_true_to_eq: \forall n,m:nat.
+eqb n m = true \to n = m.
+intros.
+change with
+match true with
+[ true \Rightarrow n = m
+| false \Rightarrow n \neq m].
+rewrite < H.
+(*qui auto non conclude il goal*)
+apply eqb_to_Prop.
+qed.
+
+theorem eqb_false_to_not_eq: \forall n,m:nat.
+eqb n m = false \to n \neq m.
+intros.
+change with
+match false with
+[ true \Rightarrow n = m
+| false \Rightarrow n \neq m].
+rewrite < H.
+(*qui auto non conclude il goal*)
+apply eqb_to_Prop.
+qed.
+
+theorem eq_to_eqb_true: \forall n,m:nat.
+n = m \to eqb n m = true.
+intros.
+auto.
+(*apply (eqb_elim n m)
+[ intros. reflexivity
+| intros.apply False_ind.apply (H1 H)
+]*)
+qed.
+
+theorem not_eq_to_eqb_false: \forall n,m:nat.
+\lnot (n = m) \to eqb n m = false.
+intros.apply (eqb_elim n m);intros
+[ apply False_ind.
+ apply (H H1)
+| reflexivity
+]
+qed.
+
+let rec leb n m \def
+match n with
+ [ O \Rightarrow true
+ | (S p) \Rightarrow
+ match m with
+ [ O \Rightarrow false
+ | (S q) \Rightarrow leb p q]].
+
+theorem leb_to_Prop: \forall n,m:nat.
+match (leb n m) with
+[ true \Rightarrow n \leq m
+| false \Rightarrow n \nleq m].
+intros.
+apply (nat_elim2
+(\lambda n,m:nat.match (leb n m) with
+[ true \Rightarrow n \leq m
+| false \Rightarrow n \nleq m]))
+[ simplify.
+ exact le_O_n
+| simplify.
+ exact not_le_Sn_O
+| intros 2.
+ simplify.
+ elim ((leb n1 m1));simplify
+ [ apply le_S_S.
+ (*qui auto non conclude il goal*)
+ apply H
+ | unfold Not.
+ intros.
+ apply H.
+ auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ ]
+]
+qed.
+
+theorem leb_elim: \forall n,m:nat. \forall P:bool \to Prop.
+(n \leq m \to (P true)) \to (n \nleq m \to (P false)) \to
+P (leb n m).
+intros.
+cut
+(match (leb n m) with
+[ true \Rightarrow n \leq m
+| false \Rightarrow n \nleq m] \to (P (leb n m)))
+[ apply Hcut.
+ (*qui auto non conclude il goal*)
+ apply leb_to_Prop
+| elim (leb n m)
+ [ (*qui auto non conclude il goal*)
+ apply ((H H2))
+ | (*qui auto non conclude il goal*)
+ apply ((H1 H2))
+ ]
+]
+qed.
+
+let rec nat_compare n m: compare \def
+match n with
+[ O \Rightarrow
+ match m with
+ [ O \Rightarrow EQ
+ | (S q) \Rightarrow LT ]
+| (S p) \Rightarrow
+ match m with
+ [ O \Rightarrow GT
+ | (S q) \Rightarrow nat_compare p q]].
+(**********)
+theorem nat_compare_n_n: \forall n:nat. nat_compare n n = EQ.
+intro.elim n
+[ auto
+ (*simplify.
+ reflexivity*)
+| simplify.
+ assumption
+]
+qed.
+
+theorem nat_compare_S_S: \forall n,m:nat.
+nat_compare n m = nat_compare (S n) (S m).
+intros.auto.
+(*simplify.reflexivity.*)
+qed.
+
+theorem S_pred: \forall n:nat.lt O n \to eq nat n (S (pred n)).
+intro.
+elim n;auto.
+(*[ apply False_ind.
+ exact (not_le_Sn_O O H)
+| apply eq_f.
+ apply pred_Sn
+]*)
+qed.
+
+theorem nat_compare_pred_pred:
+\forall n,m:nat.lt O n \to lt O m \to
+eq compare (nat_compare n m) (nat_compare (pred n) (pred m)).
+intros.
+apply (lt_O_n_elim n H).
+apply (lt_O_n_elim m H1).
+intros.
+auto.
+(*simplify.reflexivity.*)
+qed.
+
+theorem nat_compare_to_Prop: \forall n,m:nat.
+match (nat_compare n m) with
+ [ LT \Rightarrow n < m
+ | EQ \Rightarrow n=m
+ | GT \Rightarrow m < n ].
+intros.
+apply (nat_elim2 (\lambda n,m.match (nat_compare n m) with
+ [ LT \Rightarrow n < m
+ | EQ \Rightarrow n=m
+ | GT \Rightarrow m < n ]))
+[ intro.
+ elim n1;simplify;auto
+ (*[ reflexivity
+ | unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ ]*)
+| intro.
+ simplify.
+ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+| intros 2.
+ simplify.
+ elim ((nat_compare n1 m1));simplify
+ [ unfold lt.
+ apply le_S_S.
+ (*qui auto non chiude il goal*)
+ apply H
+ | apply eq_f.
+ (*qui auto non chiude il goal*)
+ apply H
+ | unfold lt.
+ apply le_S_S.
+ (*qui auto non chiude il goal*)
+ apply H
+ ]
+]
+qed.
+
+theorem nat_compare_n_m_m_n: \forall n,m:nat.
+nat_compare n m = compare_invert (nat_compare m n).
+intros.
+apply (nat_elim2 (\lambda n,m. nat_compare n m = compare_invert (nat_compare m n)));intros
+[ elim n1;auto(*;simplify;reflexivity*)
+| elim n1;auto(*;simplify;reflexivity*)
+| auto
+ (*simplify.elim H.reflexivity*)
+]
+qed.
+
+theorem nat_compare_elim : \forall n,m:nat. \forall P:compare \to Prop.
+(n < m \to P LT) \to (n=m \to P EQ) \to (m < n \to P GT) \to
+(P (nat_compare n m)).
+intros.
+cut (match (nat_compare n m) with
+[ LT \Rightarrow n < m
+| EQ \Rightarrow n=m
+| GT \Rightarrow m < n] \to
+(P (nat_compare n m)))
+[ apply Hcut.
+ (*auto non chiude il goal*)
+ apply nat_compare_to_Prop
+| elim ((nat_compare n m))
+ [ (*auto non chiude il goal*)
+ apply ((H H3))
+ | (*auto non chiude il goal*)
+ apply ((H1 H3))
+ | (*auto non chiude il goal*)
+ apply ((H2 H3))
+ ]
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / Matita is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/congruence".
+
+include "nat/relevant_equations.ma".
+include "nat/primes.ma".
+
+definition S_mod: nat \to nat \to nat \def
+\lambda n,m:nat. (S m) \mod n.
+
+definition congruent: nat \to nat \to nat \to Prop \def
+\lambda n,m,p:nat. mod n p = mod m p.
+
+interpretation "congruent" 'congruent n m p =
+ (cic:/matita/nat/congruence/congruent.con n m p).
+
+notation < "hvbox(n break \cong\sub p m)"
+ (*non associative*) with precedence 45
+for @{ 'congruent $n $m $p }.
+
+theorem congruent_n_n: \forall n,p:nat.congruent n n p.
+intros.
+unfold congruent.
+reflexivity.
+qed.
+
+theorem transitive_congruent: \forall p:nat. transitive nat
+(\lambda n,m. congruent n m p).
+intros.unfold transitive.
+unfold congruent.
+intros.
+whd.
+apply (trans_eq ? ? (y \mod p))
+[ (*qui auto non chiude il goal*)
+ apply H
+| (*qui auto non chiude il goal*)
+ apply H1
+]
+qed.
+
+theorem le_to_mod: \forall n,m:nat. n \lt m \to n = n \mod m.
+intros.
+auto.
+(*apply (div_mod_spec_to_eq2 n m O n (n/m) (n \mod m))
+[ constructor 1
+ [ assumption
+ | simplify.
+ reflexivity
+ ]
+| apply div_mod_spec_div_mod.
+ apply (le_to_lt_to_lt O n m)
+ [ apply le_O_n
+ | assumption
+ ]
+]*)
+qed.
+
+theorem mod_mod : \forall n,p:nat. O<p \to n \mod p = (n \mod p) \mod p.
+intros.
+auto.
+(*rewrite > (div_mod (n \mod p) p) in \vdash (? ? % ?)
+[ rewrite > (eq_div_O ? p)
+ [ reflexivity
+ | apply lt_mod_m_m.
+ assumption
+ ]
+| assumption
+]*)
+qed.
+
+theorem mod_times_mod : \forall n,m,p:nat. O<p \to O<m \to n \mod p = (n \mod (m*p)) \mod p.
+intros.
+apply (div_mod_spec_to_eq2 n p (n/p) (n \mod p)
+(n/(m*p)*m + (n \mod (m*p)/p)))
+[ auto.
+ (*apply div_mod_spec_div_mod.
+ assumption*)
+| constructor 1
+ [ auto
+ (*apply lt_mod_m_m.
+ assumption*)
+ | rewrite > times_plus_l.
+ rewrite > assoc_plus.
+ rewrite < div_mod
+ [ rewrite > assoc_times.
+ rewrite < div_mod;auto
+ (*[ reflexivity
+ | rewrite > (times_n_O O).
+ apply lt_times;assumption
+ ]*)
+ | assumption
+ ]
+ ]
+]
+qed.
+
+theorem congruent_n_mod_n :
+\forall n,p:nat. O < p \to congruent n (n \mod p) p.
+intros.
+unfold congruent.
+auto.
+(*apply mod_mod.
+assumption.*)
+qed.
+
+theorem congruent_n_mod_times :
+\forall n,m,p:nat. O < p \to O < m \to congruent n (n \mod (m*p)) p.
+intros.unfold congruent.
+apply mod_times_mod;assumption.
+qed.
+
+theorem eq_times_plus_to_congruent: \forall n,m,p,r:nat. O< p \to
+n = r*p+m \to congruent n m p.
+intros.
+unfold congruent.
+apply (div_mod_spec_to_eq2 n p (div n p) (mod n p) (r +(div m p)) (mod m p))
+[ auto
+ (*apply div_mod_spec_div_mod.
+ assumption*)
+| constructor 1
+ [ auto
+ (*apply lt_mod_m_m.
+ assumption*)
+ |
+(*cut (n = r * p + (m / p * p + m \mod p)).*)
+(*lapply (div_mod m p H).
+rewrite > sym_times.
+rewrite > distr_times_plus.
+(*rewrite > (sym_times p (m/p)).*)
+(*rewrite > sym_times.*)
+ rewrite > assoc_plus.
+ auto paramodulation.
+ rewrite < div_mod.
+ assumption.
+ assumption.
+*)
+ rewrite > sym_times.
+ rewrite > distr_times_plus.
+ rewrite > sym_times.
+ rewrite > (sym_times p).
+ rewrite > assoc_plus.
+ rewrite < div_mod;assumption.
+ ]
+]
+qed.
+
+theorem divides_to_congruent: \forall n,m,p:nat. O < p \to m \le n \to
+divides p (n - m) \to congruent n m p.
+intros.
+elim H2.
+apply (eq_times_plus_to_congruent n m p n2)
+[ assumption
+| rewrite < sym_plus.
+ apply minus_to_plus;auto
+ (*[ assumption
+ | rewrite > sym_times. assumption
+ ]*)
+]
+qed.
+
+theorem congruent_to_divides: \forall n,m,p:nat.
+O < p \to congruent n m p \to divides p (n - m).
+intros.
+unfold congruent in H1.
+apply (witness ? ? ((n / p)-(m / p))).
+rewrite > sym_times.
+rewrite > (div_mod n p) in \vdash (? ? % ?)
+[ rewrite > (div_mod m p) in \vdash (? ? % ?)
+ [ rewrite < (sym_plus (m \mod p)).
+ auto
+ (*rewrite < H1.
+ rewrite < (eq_minus_minus_minus_plus ? (n \mod p)).
+ rewrite < minus_plus_m_m.
+ apply sym_eq.
+ apply times_minus_l*)
+ | assumption
+ ]
+| assumption
+]
+qed.
+
+theorem mod_times: \forall n,m,p:nat.
+O < p \to mod (n*m) p = mod ((mod n p)*(mod m p)) p.
+intros.
+change with (congruent (n*m) ((mod n p)*(mod m p)) p).
+apply (eq_times_plus_to_congruent ? ? p
+((n / p)*p*(m / p) + (n / p)*(m \mod p) + (n \mod p)*(m / p)))
+[ assumption
+| apply (trans_eq ? ? (((n/p)*p+(n \mod p))*((m/p)*p+(m \mod p))))
+ [ apply eq_f2;auto(*;apply div_mod.assumption.*)
+ | apply (trans_eq ? ? (((n/p)*p)*((m/p)*p) + (n/p)*p*(m \mod p) +
+ (n \mod p)*((m / p)*p) + (n \mod p)*(m \mod p)))
+ [ apply times_plus_plus
+ | apply eq_f2
+ [ rewrite < assoc_times.
+ auto
+ (*rewrite > (assoc_times (n/p) p (m \mod p)).
+ rewrite > (sym_times p (m \mod p)).
+ rewrite < (assoc_times (n/p) (m \mod p) p).
+ rewrite < times_plus_l.
+ rewrite < (assoc_times (n \mod p)).
+ rewrite < times_plus_l.
+ apply eq_f2
+ [ apply eq_f2
+ [ reflexivity
+ | reflexivity
+ ]
+ | reflexivity
+ ]*)
+ | reflexivity
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem congruent_times: \forall n,m,n1,m1,p. O < p \to congruent n n1 p \to
+congruent m m1 p \to congruent (n*m) (n1*m1) p.
+unfold congruent.
+intros.
+rewrite > (mod_times n m p H).
+auto.
+(*rewrite > H1.
+rewrite > H2.
+apply sym_eq.
+apply mod_times.
+assumption.*)
+qed.
+
+theorem congruent_pi: \forall f:nat \to nat. \forall n,m,p:nat.O < p \to
+congruent (pi n f m) (pi n (\lambda m. mod (f m) p) m) p.
+intros.
+elim n;simplify
+[ auto
+ (*apply congruent_n_mod_n.
+ assumption*)
+| apply congruent_times
+ [ assumption
+ | auto
+ (*apply congruent_n_mod_n.
+ assumption*)
+ | (*NB: QUI AUTO NON RIESCE A CHIUDERE IL GOAL*)
+ assumption
+ ]
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* __ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/count".
+
+include "nat/relevant_equations.ma".
+include "nat/sigma_and_pi.ma".
+include "nat/permutation.ma".
+
+theorem sigma_f_g : \forall n,m:nat.\forall f,g:nat \to nat.
+sigma n (\lambda p.f p + g p) m = sigma n f m + sigma n g m.
+intros.
+elim n;simplify
+[ reflexivity
+| rewrite > H.
+ auto
+ (*rewrite > assoc_plus.
+ rewrite < (assoc_plus (g (S (n1+m)))).
+ rewrite > (sym_plus (g (S (n1+m)))).
+ rewrite > (assoc_plus (sigma n1 f m)).
+ rewrite < assoc_plus.
+ reflexivity*)
+]
+qed.
+
+theorem sigma_plus: \forall n,p,m:nat.\forall f:nat \to nat.
+sigma (S (p+n)) f m = sigma p (\lambda x.(f ((S n) + x))) m + sigma n f m.
+intros.
+elim p;simplify
+[ auto
+ (*rewrite < (sym_plus n m).
+ reflexivity*)
+| rewrite > assoc_plus in \vdash (? ? ? %).
+ rewrite < H.
+ auto
+ (*simplify.
+ rewrite < plus_n_Sm.
+ rewrite > (sym_plus n).
+ rewrite > assoc_plus.
+ rewrite < (sym_plus m).
+ rewrite < (assoc_plus n1).
+ reflexivity*)
+]
+qed.
+
+theorem sigma_plus1: \forall n,p,m:nat.\forall f:nat \to nat.
+sigma (p+(S n)) f m = sigma p (\lambda x.(f ((S n) + x))) m + sigma n f m.
+intros.
+elim p;simplify
+[ reflexivity
+| rewrite > assoc_plus in \vdash (? ? ? %).
+ rewrite < H.
+ rewrite < plus_n_Sm.
+ auto
+ (*rewrite < plus_n_Sm.simplify.
+ rewrite < (sym_plus n).
+ rewrite > assoc_plus.
+ rewrite < (sym_plus m).
+ rewrite < (assoc_plus n).
+ reflexivity*)
+]
+qed.
+
+theorem eq_sigma_sigma : \forall n,m:nat.\forall f:nat \to nat.
+sigma (pred ((S n)*(S m))) f O =
+sigma m (\lambda a.(sigma n (\lambda b.f (b*(S m) + a)) O)) O.
+intro.
+elim n;simplify
+[ rewrite < plus_n_O.
+ apply eq_sigma.
+ intros.
+ reflexivity
+| rewrite > sigma_f_g.
+ rewrite < plus_n_O.
+ rewrite < H.
+ auto
+
+ (*rewrite > (S_pred ((S n1)*(S m)))
+ [ apply sigma_plus1
+ | simplify.
+ unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ ]*)
+]
+qed.
+
+theorem eq_sigma_sigma1 : \forall n,m:nat.\forall f:nat \to nat.
+sigma (pred ((S n)*(S m))) f O =
+sigma n (\lambda a.(sigma m (\lambda b.f (b*(S n) + a)) O)) O.
+intros.
+rewrite > sym_times.
+apply eq_sigma_sigma.
+qed.
+
+theorem sigma_times: \forall n,m,p:nat.\forall f:nat \to nat.
+(sigma n f m)*p = sigma n (\lambda i.(f i) * p) m.
+intro.
+elim n;simplify
+[ reflexivity
+| rewrite < H.
+ apply times_plus_l
+]
+qed.
+
+definition bool_to_nat: bool \to nat \def
+\lambda b. match b with
+[ true \Rightarrow (S O)
+| false \Rightarrow O ].
+
+theorem bool_to_nat_andb: \forall a,b:bool.
+bool_to_nat (andb a b) = (bool_to_nat a)*(bool_to_nat b).
+intros.
+elim a;auto.
+(*[elim b
+ [ simplify.
+ reflexivity
+ | reflexivity
+ ]
+| reflexivity
+]*)
+qed.
+
+definition count : nat \to (nat \to bool) \to nat \def
+\lambda n.\lambda f. sigma (pred n) (\lambda n.(bool_to_nat (f n))) O.
+
+theorem count_times:\forall n,m:nat.
+\forall f,f1,f2:nat \to bool.
+\forall g:nat \to nat \to nat.
+\forall g1,g2: nat \to nat.
+(\forall a,b:nat. a < (S n) \to b < (S m) \to (g b a) < (S n)*(S m)) \to
+(\forall a,b:nat. a < (S n) \to b < (S m) \to (g1 (g b a)) = a) \to
+(\forall a,b:nat. a < (S n) \to b < (S m) \to (g2 (g b a)) = b) \to
+(\forall a,b:nat. a < (S n) \to b < (S m) \to f (g b a) = andb (f2 b) (f1 a)) \to
+(count ((S n)*(S m)) f) = (count (S n) f1)*(count (S m) f2).
+intros.unfold count.
+rewrite < eq_map_iter_i_sigma.
+rewrite > (permut_to_eq_map_iter_i plus assoc_plus sym_plus ? ? ?
+ (\lambda i.g (div i (S n)) (mod i (S n))))
+[ rewrite > eq_map_iter_i_sigma.
+ rewrite > eq_sigma_sigma1.
+ apply (trans_eq ? ?
+ (sigma n (\lambda a.
+ sigma m (\lambda b.(bool_to_nat (f2 b))*(bool_to_nat (f1 a))) O) O))
+ [ apply eq_sigma.intros.
+ apply eq_sigma.intros.
+ rewrite > (div_mod_spec_to_eq (i1*(S n) + i) (S n) ((i1*(S n) + i)/(S n))
+ ((i1*(S n) + i) \mod (S n)) i1 i)
+ [ rewrite > (div_mod_spec_to_eq2 (i1*(S n) + i) (S n) ((i1*(S n) + i)/(S n))
+ ((i1*(S n) + i) \mod (S n)) i1 i)
+ [ rewrite > H3;auto (*qui auto impiega parecchio tempo*)
+ (*[ apply bool_to_nat_andb
+ | unfold lt.
+ apply le_S_S.
+ assumption
+ | unfold lt.
+ apply le_S_S.
+ assumption
+ ]*)
+ | auto
+ (*apply div_mod_spec_div_mod.
+ unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | constructor 1;auto
+ (*[ unfold lt.
+ apply le_S_S.
+ assumption
+ | reflexivity
+ ]*)
+ ]
+ | auto
+ (*apply div_mod_spec_div_mod.
+ unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | constructor 1;auto
+ (*[ unfold lt.
+ apply le_S_S.
+ assumption
+ | reflexivity
+ ]*)
+ ]
+ | apply (trans_eq ? ?
+ (sigma n (\lambda n.((bool_to_nat (f1 n)) *
+ (sigma m (\lambda n.bool_to_nat (f2 n)) O))) O))
+ [ apply eq_sigma.
+ intros.
+ auto
+ (*rewrite > sym_times.
+ apply (trans_eq ? ?
+ (sigma m (\lambda n.(bool_to_nat (f2 n))*(bool_to_nat (f1 i))) O))
+ [ reflexivity.
+ | apply sym_eq.
+ apply sigma_times
+ ]*)
+ | auto
+ (*simplify.
+ apply sym_eq.
+ apply sigma_times*)
+ ]
+ ]
+| unfold permut.
+ split
+ [ intros.
+ rewrite < plus_n_O.
+ apply le_S_S_to_le.
+ rewrite < S_pred in \vdash (? ? %)
+ [ change with ((g (i/(S n)) (i \mod (S n))) \lt (S n)*(S m)).
+ apply H
+ [ auto
+ (*apply lt_mod_m_m.
+ unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | apply (lt_times_to_lt_l n).
+ apply (le_to_lt_to_lt ? i)
+ [ auto
+ (*rewrite > (div_mod i (S n)) in \vdash (? ? %)
+ [ rewrite > sym_plus.
+ apply le_plus_n
+ | unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ ]*)
+ | unfold lt.
+ rewrite > S_pred in \vdash (? ? %)
+ [ apply le_S_S.
+ auto
+ (*rewrite > plus_n_O in \vdash (? ? %).
+ rewrite > sym_times.
+ assumption*)
+ | auto
+ (*rewrite > (times_n_O O).
+ apply lt_times;
+ unfold lt;apply le_S_S;apply le_O_n*)
+ ]
+ ]
+ ]
+ | auto
+ (*rewrite > (times_n_O O).
+ apply lt_times;
+ unfold lt;apply le_S_S;apply le_O_n *)
+ ]
+ | rewrite < plus_n_O.
+ unfold injn.
+ intros.
+ cut (i < (S n)*(S m))
+ [ cut (j < (S n)*(S m))
+ [ cut ((i \mod (S n)) < (S n))
+ [ cut ((i/(S n)) < (S m))
+ [ cut ((j \mod (S n)) < (S n))
+ [ cut ((j/(S n)) < (S m))
+ [ rewrite > (div_mod i (S n))
+ [ rewrite > (div_mod j (S n))
+ [ rewrite < (H1 (i \mod (S n)) (i/(S n)) Hcut2 Hcut3).
+ rewrite < (H2 (i \mod (S n)) (i/(S n)) Hcut2 Hcut3) in \vdash (? ? (? % ?) ?).
+ rewrite < (H1 (j \mod (S n)) (j/(S n)) Hcut4 Hcut5).
+ rewrite < (H2 (j \mod (S n)) (j/(S n)) Hcut4 Hcut5) in \vdash (? ? ? (? % ?)).
+ auto
+ (*rewrite > H6.
+ reflexivity*)
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ | apply (lt_times_to_lt_l n).
+ apply (le_to_lt_to_lt ? j)
+ [ auto.
+ (*rewrite > (div_mod j (S n)) in \vdash (? ? %)
+ [ rewrite > sym_plus.
+ apply le_plus_n
+ | unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ ]*)
+ | rewrite < sym_times.
+ assumption
+ ]
+ ]
+ | auto
+ (*apply lt_mod_m_m.
+ unfold lt. apply le_S_S.
+ apply le_O_n*)
+ ]
+ | apply (lt_times_to_lt_l n).
+ apply (le_to_lt_to_lt ? i)
+ [ auto
+ (*rewrite > (div_mod i (S n)) in \vdash (? ? %)
+ [ rewrite > sym_plus.
+ apply le_plus_n
+ | unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ ]*)
+ | rewrite < sym_times.
+ assumption
+ ]
+ ]
+ | auto
+ (*apply lt_mod_m_m.
+ unfold lt. apply le_S_S.
+ apply le_O_n*)
+ ]
+ | unfold lt.
+ auto
+ (*rewrite > S_pred in \vdash (? ? %)
+ [ apply le_S_S.
+ assumption
+ | rewrite > (times_n_O O).
+ apply lt_times;
+ unfold lt; apply le_S_S;apply le_O_n
+ ]*)
+ ]
+ | unfold lt.
+ auto
+ (*rewrite > S_pred in \vdash (? ? %)
+ [ apply le_S_S.
+ assumption
+ | rewrite > (times_n_O O).
+ apply lt_times;
+ unfold lt; apply le_S_S;apply le_O_n
+ ]*)
+ ]
+ ]
+| intros.
+ apply False_ind.
+ apply (not_le_Sn_O m1 H4)
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / Matita is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/div_and_mod".
+
+include "datatypes/constructors.ma".
+include "nat/minus.ma".
+
+let rec mod_aux p m n: nat \def
+match (leb m n) with
+[ true \Rightarrow m
+| false \Rightarrow
+ match p with
+ [O \Rightarrow m
+ |(S q) \Rightarrow mod_aux q (m-(S n)) n]].
+
+definition mod : nat \to nat \to nat \def
+\lambda n,m.
+match m with
+[O \Rightarrow m
+| (S p) \Rightarrow mod_aux n n p].
+
+interpretation "natural remainder" 'module x y =
+ (cic:/matita/nat/div_and_mod/mod.con x y).
+
+let rec div_aux p m n : nat \def
+match (leb m n) with
+[ true \Rightarrow O
+| false \Rightarrow
+ match p with
+ [O \Rightarrow O
+ |(S q) \Rightarrow S (div_aux q (m-(S n)) n)]].
+
+definition div : nat \to nat \to nat \def
+\lambda n,m.
+match m with
+[O \Rightarrow S n
+| (S p) \Rightarrow div_aux n n p].
+
+interpretation "natural divide" 'divide x y =
+ (cic:/matita/nat/div_and_mod/div.con x y).
+
+theorem le_mod_aux_m_m:
+\forall p,n,m. n \leq p \to (mod_aux p n m) \leq m.
+intro.
+elim p
+[ apply (le_n_O_elim n H (\lambda n.(mod_aux O n m) \leq m)).
+ auto
+ (*simplify.
+ apply le_O_n*)
+| simplify.
+ apply (leb_elim n1 m);simplify;intro
+ [ assumption
+ | apply H.
+ cut (n1 \leq (S n) \to n1-(S m) \leq n)
+ [ auto
+ (*apply Hcut.
+ assumption*)
+ | elim n1;simplify;auto
+ (*[ apply le_O_n.
+ | apply (trans_le ? n2 n)
+ [ apply le_minus_m
+ | apply le_S_S_to_le.
+ assumption
+ ]
+ ]*)
+ ]
+ ]
+]
+qed.
+
+theorem lt_mod_m_m: \forall n,m. O < m \to (n \mod m) < m.
+intros 2.
+elim m
+[ apply False_ind.
+ apply (not_le_Sn_O O H)
+| simplify.
+ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_mod_aux_m_m.
+ apply le_n*)
+]
+qed.
+
+theorem div_aux_mod_aux: \forall p,n,m:nat.
+(n=(div_aux p n m)*(S m) + (mod_aux p n m)).
+intro.
+elim p;simplify
+[ elim (leb n m);auto
+ (*simplify;apply refl_eq.*)
+| apply (leb_elim n1 m);simplify;intro
+ [ apply refl_eq
+ | rewrite > assoc_plus.
+ elim (H (n1-(S m)) m).
+ change with (n1=(S m)+(n1-(S m))).
+ rewrite < sym_plus.
+ auto
+ (*apply plus_minus_m_m.
+ change with (m < n1).
+ apply not_le_to_lt.
+ exact H1*)
+ ]
+]
+qed.
+
+theorem div_mod: \forall n,m:nat. O < m \to n=(n / m)*m+(n \mod m).
+intros 2.
+elim m
+[ elim (not_le_Sn_O O H)
+| simplify.
+ apply div_aux_mod_aux
+]
+qed.
+
+inductive div_mod_spec (n,m,q,r:nat) : Prop \def
+div_mod_spec_intro: r < m \to n=q*m+r \to (div_mod_spec n m q r).
+
+(*
+definition div_mod_spec : nat \to nat \to nat \to nat \to Prop \def
+\lambda n,m,q,r:nat.r < m \land n=q*m+r).
+*)
+
+theorem div_mod_spec_to_not_eq_O: \forall n,m,q,r.(div_mod_spec n m q r) \to m \neq O.
+intros 4.
+unfold Not.
+intros.
+elim H.
+absurd (le (S r) O);auto.
+(*[ rewrite < H1.
+ assumption
+| exact (not_le_Sn_O r).
+]*)
+qed.
+
+theorem div_mod_spec_div_mod:
+\forall n,m. O < m \to (div_mod_spec n m (n / m) (n \mod m)).
+intros.
+auto.
+(*apply div_mod_spec_intro
+[ apply lt_mod_m_m.
+ assumption
+| apply div_mod.
+ assumption
+]*)
+qed.
+
+theorem div_mod_spec_to_eq :\forall a,b,q,r,q1,r1.
+(div_mod_spec a b q r) \to (div_mod_spec a b q1 r1) \to
+(eq nat q q1).
+intros.
+elim H.
+elim H1.
+apply (nat_compare_elim q q1)
+[ intro.
+ apply False_ind.
+ cut (eq nat ((q1-q)*b+r1) r)
+ [ cut (b \leq (q1-q)*b+r1)
+ [ cut (b \leq r)
+ [ apply (lt_to_not_le r b H2 Hcut2)
+ | elim Hcut.
+ assumption
+ ]
+ | apply (trans_le ? ((q1-q)*b));auto
+ (*[ apply le_times_n.
+ apply le_SO_minus.
+ exact H6
+ | rewrite < sym_plus.
+ apply le_plus_n
+ ]*)
+ ]
+ | rewrite < sym_times.
+ rewrite > distr_times_minus.
+ rewrite > plus_minus;auto
+ (*[ rewrite > sym_times.
+ rewrite < H5.
+ rewrite < sym_times.
+ apply plus_to_minus.
+ apply H3
+ | apply le_times_r.
+ apply lt_to_le.
+ apply H6
+ ]*)
+ ]
+| (* eq case *)
+ auto
+ (*intros.
+ assumption*)
+| (* the following case is symmetric *)
+ intro.
+ apply False_ind.
+ cut (eq nat ((q-q1)*b+r) r1)
+ [ cut (b \leq (q-q1)*b+r)
+ [ cut (b \leq r1)
+ [ apply (lt_to_not_le r1 b H4 Hcut2)
+ | elim Hcut.
+ assumption
+ ]
+ | apply (trans_le ? ((q-q1)*b));auto
+ (*[ apply le_times_n.
+ apply le_SO_minus.
+ exact H6
+ | rewrite < sym_plus.
+ apply le_plus_n
+ ]*)
+ ]
+ | rewrite < sym_times.
+ rewrite > distr_times_minus.
+ rewrite > plus_minus;auto
+ (*[ rewrite > sym_times.
+ rewrite < H3.
+ rewrite < sym_times.
+ apply plus_to_minus.
+ apply H5
+ | apply le_times_r.
+ apply lt_to_le.
+ apply H6
+ ]*)
+ ]
+]
+qed.
+
+theorem div_mod_spec_to_eq2 :\forall a,b,q,r,q1,r1.
+(div_mod_spec a b q r) \to (div_mod_spec a b q1 r1) \to
+(eq nat r r1).
+intros.
+elim H.
+elim H1.
+apply (inj_plus_r (q*b)).
+rewrite < H3.
+rewrite > (div_mod_spec_to_eq a b q r q1 r1 H H1).
+assumption.
+qed.
+
+theorem div_mod_spec_times : \forall n,m:nat.div_mod_spec ((S n)*m) (S n) m O.
+intros.
+auto.
+(*constructor 1
+[ unfold lt.
+ apply le_S_S.
+ apply le_O_n
+| rewrite < plus_n_O.
+ rewrite < sym_times.
+ reflexivity
+]*)
+qed.
+
+
+(*il corpo del seguente teorema non e' stato strutturato *)
+(* some properties of div and mod *)
+theorem div_times: \forall n,m:nat. ((S n)*m) / (S n) = m.
+intros.
+apply (div_mod_spec_to_eq ((S n)*m) (S n) ? ? ? O).
+goal 15. (* ?11 is closed with the following tactics *)
+apply div_mod_spec_div_mod.auto.auto.
+(*unfold lt.apply le_S_S.apply le_O_n.
+apply div_mod_spec_times.*)
+qed.
+
+theorem div_n_n: \forall n:nat. O < n \to n / n = S O.
+intros.
+apply (div_mod_spec_to_eq n n (n / n) (n \mod n) (S O) O);auto.
+(*[ apply div_mod_spec_div_mod.
+ assumption
+| constructor 1
+ [ assumption
+ | rewrite < plus_n_O.
+ simplify.
+ rewrite < plus_n_O.
+ reflexivity
+ ]
+] *)
+qed.
+
+theorem eq_div_O: \forall n,m. n < m \to n / m = O.
+intros.
+apply (div_mod_spec_to_eq n m (n/m) (n \mod m) O n);auto.
+(*[ apply div_mod_spec_div_mod.
+ apply (le_to_lt_to_lt O n m)
+ [ apply le_O_n
+ | assumption
+ ]
+| constructor 1
+ [ assumption
+ | reflexivity
+ ]
+]*)
+qed.
+
+theorem mod_n_n: \forall n:nat. O < n \to n \mod n = O.
+intros.
+apply (div_mod_spec_to_eq2 n n (n / n) (n \mod n) (S O) O);auto.
+(*[ apply div_mod_spec_div_mod.
+ assumption
+| constructor 1
+ [ assumption.
+ | rewrite < plus_n_O.
+ simplify.
+ rewrite < plus_n_O.
+ reflexivity
+ ]
+]*)
+qed.
+
+theorem mod_S: \forall n,m:nat. O < m \to S (n \mod m) < m \to
+((S n) \mod m) = S (n \mod m).
+intros.
+apply (div_mod_spec_to_eq2 (S n) m ((S n) / m) ((S n) \mod m) (n / m) (S (n \mod m)))
+[ auto
+ (*apply div_mod_spec_div_mod.
+ assumption*)
+| constructor 1
+ [ assumption
+ | rewrite < plus_n_Sm.
+ auto
+ (*apply eq_f.
+ apply div_mod.
+ assumption*)
+ ]
+]
+qed.
+
+theorem mod_O_n: \forall n:nat.O \mod n = O.
+intro.
+elim n;auto.
+ (*simplify;reflexivity*)
+
+qed.
+
+theorem lt_to_eq_mod:\forall n,m:nat. n < m \to n \mod m = n.
+intros.
+apply (div_mod_spec_to_eq2 n m (n/m) (n \mod m) O n);auto.
+(*[ apply div_mod_spec_div_mod.
+ apply (le_to_lt_to_lt O n m)
+ [ apply le_O_n
+ | assumption
+ ]
+| constructor 1.
+ [ assumption
+ | reflexivity
+ ]
+]*)
+qed.
+
+(* injectivity *)
+theorem injective_times_r: \forall n:nat.injective nat nat (\lambda m:nat.(S n)*m).
+change with (\forall n,p,q:nat.(S n)*p = (S n)*q \to p=q).
+intros.
+rewrite < (div_times n).
+auto.
+(*rewrite < (div_times n q).
+apply eq_f2
+[ assumption
+| reflexivity
+]*)
+qed.
+
+variant inj_times_r : \forall n,p,q:nat.(S n)*p = (S n)*q \to p=q \def
+injective_times_r.
+
+theorem lt_O_to_injective_times_r: \forall n:nat. O < n \to injective nat nat (\lambda m:nat.n*m).
+simplify.
+intros 4.
+apply (lt_O_n_elim n H).
+intros.
+auto.
+(*apply (inj_times_r m).
+assumption.*)
+qed.
+
+variant inj_times_r1:\forall n. O < n \to \forall p,q:nat.n*p = n*q \to p=q
+\def lt_O_to_injective_times_r.
+
+theorem injective_times_l: \forall n:nat.injective nat nat (\lambda m:nat.m*(S n)).
+simplify.
+intros.
+auto.
+(*apply (inj_times_r n x y).
+rewrite < sym_times.
+rewrite < (sym_times y).
+assumption.*)
+qed.
+
+variant inj_times_l : \forall n,p,q:nat. p*(S n) = q*(S n) \to p=q \def
+injective_times_l.
+
+theorem lt_O_to_injective_times_l: \forall n:nat. O < n \to injective nat nat (\lambda m:nat.m*n).
+simplify.
+intros 4.
+apply (lt_O_n_elim n H).
+intros.
+auto.
+(*apply (inj_times_l m).
+assumption.*)
+qed.
+
+variant inj_times_l1:\forall n. O < n \to \forall p,q:nat.p*n = q*n \to p=q
+\def lt_O_to_injective_times_l.
+
+(* n_divides computes the pair (div,mod) *)
+
+(* p is just an upper bound, acc is an accumulator *)
+let rec n_divides_aux p n m acc \def
+ match n \mod m with
+ [ O \Rightarrow
+ match p with
+ [ O \Rightarrow pair nat nat acc n
+ | (S p) \Rightarrow n_divides_aux p (n / m) m (S acc)]
+ | (S a) \Rightarrow pair nat nat acc n].
+
+(* n_divides n m = <q,r> if m divides n q times, with remainder r *)
+definition n_divides \def \lambda n,m:nat.n_divides_aux n n m O.
--- /dev/null
+(**************************************************************************)
+(* __ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/euler_theorem".
+
+include "nat/map_iter_p.ma".
+include "nat/totient.ma".
+
+(* a reformulation of totient using card insted of count *)
+lemma totient_card: \forall n.
+totient n = card n (\lambda i.eqb (gcd i n) (S O)).
+intro.
+apply (nat_case n)
+[ reflexivity
+| intro.
+ apply (nat_case m)
+ [ reflexivity
+ | intro.
+ apply count_card1;auto
+ (*[ reflexivity
+ | auto.rewrite > gcd_n_n.
+ reflexivity
+ ]*)
+ ]
+]
+qed.
+
+theorem gcd_pi_p: \forall n,k. O < k \to k \le n \to
+gcd n (pi_p (\lambda i.eqb (gcd i n) (S O)) k) = (S O).
+intros 3.
+elim H
+[ rewrite > pi_p_S.
+ cut (eqb (gcd (S O) n) (S O) = true)
+ [ rewrite > Hcut.
+ auto
+ (*change with ((gcd n (S O)) = (S O)).
+ auto*)
+ | auto
+ (*apply eq_to_eqb_true.auto*)
+ ]
+| rewrite > pi_p_S.
+ apply eqb_elim
+ [ intro.
+ change with
+ ((gcd n ((S n1)*(pi_p (\lambda i.eqb (gcd i n) (S O)) n1))) = (S O)).
+ apply eq_gcd_times_SO
+ [ auto
+ (*unfold.
+ apply le_S.
+ assumption*)
+ | apply lt_O_pi_p.
+ | auto
+ (*rewrite > sym_gcd.
+ assumption.*)
+ | apply H2.
+ auto
+ (*apply (trans_le ? (S n1))
+ [ apply le_n_Sn
+ | assumption
+ ]*)
+ ]
+ | intro.
+ change with
+ (gcd n (pi_p (\lambda i.eqb (gcd i n) (S O)) n1) = (S O)).
+ apply H2.
+ auto
+ (*apply (trans_le ? (S n1))
+ [ apply le_n_Sn
+ | assumption
+ ]*)
+ ]
+]
+qed.
+
+theorem congruent_map_iter_p_times:\forall f:nat \to nat. \forall a,n:nat.
+O < a \to
+congruent
+(map_iter_p n (\lambda i.eqb (gcd i a) (S O)) (\lambda x.f x) (S O) times)
+(map_iter_p n (\lambda i.eqb (gcd i a) (S O))
+ (\lambda x.f x \mod a) (S O) times) a.
+intros.
+elim n
+[ auto
+ (*rewrite > map_iter_p_O.
+ apply (congruent_n_n ? a)*)
+| apply (eqb_elim (gcd (S n1) a) (S O))
+ [ intro.
+ rewrite > map_iter_p_S_true
+ [ rewrite > map_iter_p_S_true
+ [ apply congruent_times
+ [ assumption
+ | auto
+ (*apply congruent_n_mod_n.
+ assumption*)
+ | (*NB qui auto non chiude il goal*)
+ assumption
+ ]
+ | auto
+ (*apply eq_to_eqb_true.
+ assumption*)
+ ]
+ | auto
+ (*apply eq_to_eqb_true.
+ assumption*)
+ ]
+ | intro.
+ rewrite > map_iter_p_S_false
+ [ rewrite > map_iter_p_S_false
+ [ (*BN qui auto non chiude il goal*)
+ assumption
+ | auto
+ (*apply not_eq_to_eqb_false.
+ assumption*)
+ ]
+ | auto
+ (*apply not_eq_to_eqb_false.
+ assumption*)
+ ]
+ ]
+]
+qed.
+
+theorem permut_p_mod: \forall a,n. S O < n \to O < a \to gcd a n = (S O) \to
+permut_p (\lambda x:nat.a*x \mod n) (\lambda i:nat.eqb (gcd i n) (S O)) n.
+intros.
+lapply (lt_S_to_lt ? ? H) as H3.
+unfold permut_p.
+simplify.
+intros.
+split
+[ split
+ [ auto
+ (*apply lt_to_le.
+ apply lt_mod_m_m.
+ assumption*)
+ | rewrite > sym_gcd.
+ rewrite > gcd_mod
+ [ apply eq_to_eqb_true.
+ rewrite > sym_gcd.
+ apply eq_gcd_times_SO
+ [ assumption
+ | apply (gcd_SO_to_lt_O i n H).
+ auto
+ (*apply eqb_true_to_eq.
+ assumption*)
+ | auto
+ (*rewrite > sym_gcd.
+ assumption*)
+ | auto
+ (*rewrite > sym_gcd.
+ apply eqb_true_to_eq.
+ assumption*)
+ ]
+ | assumption
+ ]
+ ]
+| intros.
+ lapply (gcd_SO_to_lt_n ? ? H H4 (eqb_true_to_eq ? ? H5)) as H9.
+ lapply (gcd_SO_to_lt_n ? ? H H7 (eqb_true_to_eq ? ? H6)) as H10.
+ lapply (gcd_SO_to_lt_O ? ? H (eqb_true_to_eq ? ? H5)) as H11.
+ lapply (gcd_SO_to_lt_O ? ? H (eqb_true_to_eq ? ? H6)) as H12.
+ unfold Not.
+ intro.
+ apply H8.
+ apply (nat_compare_elim i j)
+ [ intro.
+ absurd (j < n)
+ [ assumption
+ | apply le_to_not_lt.
+ apply (trans_le ? (j -i))
+ [ apply divides_to_le
+ [(*fattorizzare*)
+ auto (*qui auto e' efficace, ma impiega un tempo piuttosto alto a chiudere il goal corrente*)
+ (*apply (lt_plus_to_lt_l i).
+ simplify.
+ rewrite < (plus_minus_m_m)
+ [ assumption
+ | apply lt_to_le.
+ assumption
+ ]*)
+ | apply (gcd_SO_to_divides_times_to_divides a)
+ [ assumption
+ | auto
+ (*rewrite > sym_gcd.
+ assumption*)
+ | apply mod_O_to_divides
+ [ assumption
+ | rewrite > distr_times_minus.
+ auto
+ ]
+ ]
+ ]
+ | auto
+ ]
+ ]
+ | auto
+ (*intro.
+ assumption*)
+ | intro.
+ absurd (i < n)
+ [ assumption
+ | apply le_to_not_lt.
+ apply (trans_le ? (i -j))
+ [ apply divides_to_le
+ [(*fattorizzare*)
+ auto (*qui auto e' efficace, ma impiega un tempo piuttosto alto per concludere il goal attuale*)
+ (*apply (lt_plus_to_lt_l j).
+ simplify.
+ rewrite < (plus_minus_m_m)
+ [ assumption
+ | apply lt_to_le.
+ assumption
+ ]*)
+ | apply (gcd_SO_to_divides_times_to_divides a)
+ [ assumption
+ | auto
+ (*rewrite > sym_gcd.
+ assumption*)
+ | apply mod_O_to_divides
+ [ assumption
+ | rewrite > distr_times_minus.
+ auto
+ ]
+ ]
+ ]
+ | auto
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem congruent_exp_totient_SO: \forall n,a:nat. (S O) < n \to
+gcd a n = (S O) \to congruent (exp a (totient n)) (S O) n.
+intros.
+cut (O < a)
+[ apply divides_to_congruent
+ [ auto
+ (*apply (trans_lt ? (S O)).
+ apply lt_O_S.
+ assumption*)
+ | auto
+ (*change with (O < exp a (totient n)).
+ apply lt_O_exp.
+ assumption*)
+ | apply (gcd_SO_to_divides_times_to_divides (pi_p (\lambda i.eqb (gcd i n) (S O)) n))
+ [ auto
+ (*apply (trans_lt ? (S O)).
+ apply lt_O_S.
+ assumption*)
+ | auto
+ (*apply gcd_pi_p
+ [ apply (trans_lt ? (S O)).
+ apply lt_O_S.
+ assumption
+ | apply le_n
+ ]*)
+ | rewrite < sym_times.
+ rewrite > times_minus_l.
+ rewrite > (sym_times (S O)).
+ rewrite < times_n_SO.
+ rewrite > totient_card.
+ rewrite > a_times_pi_p.
+ apply congruent_to_divides
+ [ auto
+ (*apply (trans_lt ? (S O)).
+ apply lt_O_S.
+ assumption*)
+ | apply (transitive_congruent n ?
+ (map_iter_p n (\lambda i.eqb (gcd i n) (S O)) (\lambda x.a*x \mod n) (S O) times))
+ [ auto
+ (*apply (congruent_map_iter_p_times ? n n).
+ apply (trans_lt ? (S O))
+ [ apply lt_O_S
+ | assumption
+ ]*)
+ | unfold pi_p.
+ cut ( (map_iter_p n (\lambda i:nat.eqb (gcd i n) (S O)) (\lambda n:nat.n) (S O) times)
+ = (map_iter_p n (\lambda i:nat.eqb (gcd i n) (S O)) (\lambda x:nat.a*x\mod n) (S O) times))
+ [ rewrite < Hcut1.
+ apply congruent_n_n
+ | apply (eq_map_iter_p_permut ? ? ? ? ? (λm.m))
+ [ apply assoc_times
+ | apply sym_times
+ | apply (permut_p_mod ? ? H Hcut H1)
+ | simplify.
+ apply not_eq_to_eqb_false.
+ unfold.
+ intro.
+ auto
+ (*apply (lt_to_not_eq (S O) n)
+ [ assumption
+ | apply sym_eq.
+ assumption
+ ]*)
+ ]
+ ]
+ ]
+ ]
+ ]
+ ]
+| elim (le_to_or_lt_eq O a (le_O_n a));auto
+ (*[ assumption
+ | auto.absurd (gcd a n = S O)
+ [ assumption
+ | rewrite < H2.
+ simplify.
+ unfold.intro.
+ apply (lt_to_not_eq (S O) n)
+ [ assumption
+ | apply sym_eq.
+ assumption
+ ]
+ ]
+ ]*)
+]
+qed.
+
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/exp".
+
+include "nat/div_and_mod.ma".
+
+let rec exp n m on m\def
+ match m with
+ [ O \Rightarrow (S O)
+ | (S p) \Rightarrow (times n (exp n p)) ].
+
+interpretation "natural exponent" 'exp a b = (cic:/matita/nat/exp/exp.con a b).
+
+theorem exp_plus_times : \forall n,p,q:nat.
+n \sup (p + q) = (n \sup p) * (n \sup q).
+intros.
+elim p;simplify;auto.
+(*[ rewrite < plus_n_O.
+ reflexivity
+| rewrite > H.
+ symmetry.
+ apply assoc_times
+]*)
+qed.
+
+theorem exp_n_O : \forall n:nat. S O = n \sup O.
+intro.
+auto.
+(*simplify.
+reflexivity.*)
+qed.
+
+theorem exp_n_SO : \forall n:nat. n = n \sup (S O).
+intro.
+auto.
+(*simplify.
+rewrite < times_n_SO.
+reflexivity.*)
+qed.
+
+theorem exp_exp_times : \forall n,p,q:nat.
+(n \sup p) \sup q = n \sup (p * q).
+intros.
+elim q;simplify
+[ auto.
+ (*rewrite < times_n_O.
+ simplify.
+ reflexivity*)
+| rewrite > H.
+ rewrite < exp_plus_times.
+ auto
+ (*rewrite < times_n_Sm.
+ reflexivity*)
+]
+qed.
+
+theorem lt_O_exp: \forall n,m:nat. O < n \to O < n \sup m.
+intros.
+elim m;simplify;auto.
+ (*unfold lt
+[ apply le_n
+| rewrite > times_n_SO.
+ apply le_times;assumption
+]*)
+qed.
+
+theorem lt_m_exp_nm: \forall n,m:nat. (S O) < n \to m < n \sup m.
+intros.
+elim m;simplify;unfold lt;
+[ apply le_n.
+| apply (trans_le ? ((S(S O))*(S n1)))
+ [ simplify.
+ rewrite < plus_n_Sm.
+ apply le_S_S.
+ auto
+ (*apply le_S_S.
+ rewrite < sym_plus.
+ apply le_plus_n*)
+ | auto
+ (*apply le_times;assumption*)
+ ]
+]
+qed.
+
+theorem exp_to_eq_O: \forall n,m:nat. (S O) < n
+\to n \sup m = (S O) \to m = O.
+intros.
+apply antisym_le
+[ apply le_S_S_to_le.
+ rewrite < H1.
+ auto
+ (*change with (m < n \sup m).
+ apply lt_m_exp_nm.
+ assumption*)
+| apply le_O_n
+]
+qed.
+
+theorem injective_exp_r: \forall n:nat. (S O) < n \to
+injective nat nat (\lambda m:nat. n \sup m).
+simplify.
+intros 4.
+apply (nat_elim2 (\lambda x,y.n \sup x = n \sup y \to x = y))
+[ intros.
+ auto
+ (*apply sym_eq.
+ apply (exp_to_eq_O n)
+ [ assumption
+ | rewrite < H1.
+ reflexivity
+ ]*)
+| intros.
+ apply (exp_to_eq_O n);assumption
+| intros.
+ apply eq_f.
+ apply H1.
+ (* esprimere inj_times senza S *)
+ cut (\forall a,b:nat.O < n \to n*a=n*b \to a=b)
+ [ apply Hcut
+ [ auto
+ (*simplify.
+ unfold lt.
+ apply le_S_S_to_le.
+ apply le_S.
+ assumption*)
+ | (*NB qui auto non chiude il goal, chiuso invece chiamando solo la tattica assumption*)
+ assumption
+ ]
+ | intros 2.
+ apply (nat_case n);intros;auto
+ (*[ apply False_ind.
+ apply (not_le_Sn_O O H3)
+ | apply (inj_times_r m1).
+ assumption
+ ]*)
+ ]
+]
+qed.
+
+variant inj_exp_r: \forall p:nat. (S O) < p \to \forall n,m:nat.
+p \sup n = p \sup m \to n = m \def
+injective_exp_r.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / Matita is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/factorial".
+
+include "nat/le_arith.ma".
+
+let rec fact n \def
+ match n with
+ [ O \Rightarrow (S O)
+ | (S m) \Rightarrow (S m)*(fact m)].
+
+interpretation "factorial" 'fact n = (cic:/matita/nat/factorial/fact.con n).
+
+theorem le_SO_fact : \forall n. (S O) \le n!.
+intro.
+elim n
+[ auto
+ (*simplify.
+ apply le_n*)
+| change with ((S O) \le (S n1)*n1!).
+ auto
+ (*apply (trans_le ? ((S n1)*(S O)))
+ [ simplify.
+ apply le_S_S.
+ apply le_O_n
+ | apply le_times_r.
+ assumption
+ ]*)
+]
+qed.
+
+theorem le_SSO_fact : \forall n. (S O) < n \to (S(S O)) \le n!.
+intro.
+apply (nat_case n)
+[ intro.
+ auto
+ (*apply False_ind.
+ apply (not_le_Sn_O (S O) H).*)
+| intros.
+ change with ((S (S O)) \le (S m)*m!).
+ apply (trans_le ? ((S(S O))*(S O)));auto
+ (*[ apply le_n
+ | apply le_times
+ [ exact H
+ | apply le_SO_fact
+ ]
+ ]*)
+]
+qed.
+
+theorem le_n_fact_n: \forall n. n \le n!.
+intro.
+elim n
+[ apply le_O_n
+| change with (S n1 \le (S n1)*n1!).
+ apply (trans_le ? ((S n1)*(S O)));auto
+ (*[ rewrite < times_n_SO.
+ apply le_n
+ | apply le_times.
+ apply le_n.
+ apply le_SO_fact
+ ]*)
+]
+qed.
+
+theorem lt_n_fact_n: \forall n. (S(S O)) < n \to n < n!.
+intro.
+apply (nat_case n)
+[ intro.
+ auto
+ (*apply False_ind.
+ apply (not_le_Sn_O (S(S O)) H)*)
+| intros.
+ change with ((S m) < (S m)*m!).
+ apply (lt_to_le_to_lt ? ((S m)*(S (S O))))
+ [ rewrite < sym_times.
+ simplify.
+ unfold lt.
+ apply le_S_S.
+ auto
+ (*rewrite < plus_n_O.
+ apply le_plus_n*)
+ | apply le_times_r.
+ auto
+ (*apply le_SSO_fact.
+ simplify.
+ unfold lt.
+ apply le_S_S_to_le.
+ exact H*)
+ ]
+]
+qed.
+
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / Matita is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/factorization".
+
+include "nat/ord.ma".
+include "nat/gcd.ma".
+include "nat/nth_prime.ma".
+
+(* the following factorization algorithm looks for the largest prime
+ factor. *)
+definition max_prime_factor \def \lambda n:nat.
+(max n (\lambda p:nat.eqb (n \mod (nth_prime p)) O)).
+
+(* max_prime_factor is indeed a factor *)
+theorem divides_max_prime_factor_n:
+ \forall n:nat. (S O) < n
+ \to nth_prime (max_prime_factor n) \divides n.
+intros.
+apply divides_b_true_to_divides
+[ apply lt_O_nth_prime_n
+| apply (f_max_true (\lambda p:nat.eqb (n \mod (nth_prime p)) O) n);
+ cut (\exists i. nth_prime i = smallest_factor n)
+ [ elim Hcut.
+ apply (ex_intro nat ? a).
+ split
+ [ apply (trans_le a (nth_prime a))
+ [ auto
+ (*apply le_n_fn.
+ exact lt_nth_prime_n_nth_prime_Sn*)
+ | rewrite > H1.
+ apply le_smallest_factor_n
+ ]
+ | rewrite > H1.
+ (*CSC: simplify here does something nasty! *)
+ change with (divides_b (smallest_factor n) n = true).
+ apply divides_to_divides_b_true
+ [ auto
+ (*apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | apply lt_SO_smallest_factor.
+ assumption
+ ]*)
+ | auto
+ (*letin x \def le.
+ auto new*)
+ (*
+ apply divides_smallest_factor_n;
+ apply (trans_lt ? (S O));
+ [ unfold lt; apply le_n;
+ | assumption; ] *)
+ ]
+ ]
+ | auto
+ (*
+ apply prime_to_nth_prime;
+ apply prime_smallest_factor_n;
+ assumption; *)
+ ]
+]
+qed.
+
+theorem divides_to_max_prime_factor : \forall n,m. (S O) < n \to O < m \to n \divides m \to
+max_prime_factor n \le max_prime_factor m.
+intros.
+unfold max_prime_factor.
+apply f_m_to_le_max
+[ auto
+ (*apply (trans_le ? n)
+ [ apply le_max_n
+ | apply divides_to_le;assumption
+ ]*)
+| change with (divides_b (nth_prime (max_prime_factor n)) m = true).
+ apply divides_to_divides_b_true
+ [ auto
+ (*cut (prime (nth_prime (max_prime_factor n)))
+ [ apply lt_O_nth_prime_n
+ | apply prime_nth_prime
+ ]*)
+ | auto
+ (*cut (nth_prime (max_prime_factor n) \divides n)
+ [ auto
+ | auto
+ ] *)
+ (*
+ [ apply (transitive_divides ? n);
+ [ apply divides_max_prime_factor_n.
+ assumption.
+ | assumption.
+ ]
+ | apply divides_b_true_to_divides;
+ [ apply lt_O_nth_prime_n.
+ | apply divides_to_divides_b_true;
+ [ apply lt_O_nth_prime_n.
+ | apply divides_max_prime_factor_n.
+ assumption.
+ ]
+ ]
+ ]
+ *)
+ ]
+]
+qed.
+
+theorem p_ord_to_lt_max_prime_factor: \forall n,p,q,r. O < n \to
+p = max_prime_factor n \to
+(pair nat nat q r) = p_ord n (nth_prime p) \to
+(S O) < r \to max_prime_factor r < p.
+intros.
+rewrite > H1.
+cut (max_prime_factor r \lt max_prime_factor n \lor
+ max_prime_factor r = max_prime_factor n)
+[ elim Hcut
+ [ assumption
+ | absurd (nth_prime (max_prime_factor n) \divides r)
+ [ rewrite < H4.
+ auto
+ (*apply divides_max_prime_factor_n.
+ assumption*)
+ | unfold Not.
+ intro.
+ cut (r \mod (nth_prime (max_prime_factor n)) \neq O)
+ [ auto
+ (*unfold Not in Hcut1.
+ auto new*)
+ (*
+ apply Hcut1.apply divides_to_mod_O;
+ [ apply lt_O_nth_prime_n.
+ | assumption.
+ ]
+ *)
+ | letin z \def le.
+ cut(pair nat nat q r=p_ord_aux n n (nth_prime (max_prime_factor n)));
+ [ 2: rewrite < H1.
+ assumption
+ | letin x \def le.
+ auto width = 4 new
+ ]
+ (* CERCA COME MAI le_n non lo applica se lo trova come Const e non Rel *)
+ ]
+ (*
+ apply (p_ord_aux_to_not_mod_O n n ? q r);
+ [ apply lt_SO_nth_prime_n.
+ | assumption.
+ | apply le_n.
+ | rewrite < H1.assumption.
+ ]
+ ].
+ *)
+ ]
+ ]
+| apply (le_to_or_lt_eq (max_prime_factor r) (max_prime_factor n)).
+ apply divides_to_max_prime_factor
+ [ assumption
+ | assumption
+ | apply (witness r n ((nth_prime p) \sup q)).
+ rewrite < sym_times.
+ apply (p_ord_aux_to_exp n n ? q r)
+ [ apply lt_O_nth_prime_n
+ | assumption
+ ]
+ ]
+]
+qed.
+
+theorem p_ord_to_lt_max_prime_factor1: \forall n,p,q,r. O < n \to
+max_prime_factor n \le p \to
+(pair nat nat q r) = p_ord n (nth_prime p) \to
+(S O) < r \to max_prime_factor r < p.
+intros.
+cut (max_prime_factor n < p \lor max_prime_factor n = p)
+[ elim Hcut
+ [ apply (le_to_lt_to_lt ? (max_prime_factor n))
+ [ apply divides_to_max_prime_factor
+ [ assumption
+ | assumption
+ | apply (witness r n ((nth_prime p) \sup q)).
+ rewrite > sym_times.
+ (*qui auto non chiude il goal*)
+ apply (p_ord_aux_to_exp n n)
+ [ apply lt_O_nth_prime_n.
+ | assumption
+ ]
+ ]
+ | assumption
+ ]
+ | apply (p_ord_to_lt_max_prime_factor n ? q);auto
+ (*[ assumption
+ | apply sym_eq.
+ assumption
+ | assumption
+ | assumption
+ ]*)
+ ]
+| apply (le_to_or_lt_eq ? p H1)
+]
+qed.
+
+(* datatypes and functions *)
+
+inductive nat_fact : Set \def
+ nf_last : nat \to nat_fact
+ | nf_cons : nat \to nat_fact \to nat_fact.
+
+inductive nat_fact_all : Set \def
+ nfa_zero : nat_fact_all
+ | nfa_one : nat_fact_all
+ | nfa_proper : nat_fact \to nat_fact_all.
+
+let rec factorize_aux p n acc \def
+ match p with
+ [ O \Rightarrow acc
+ | (S p1) \Rightarrow
+ match p_ord n (nth_prime p1) with
+ [ (pair q r) \Rightarrow
+ factorize_aux p1 r (nf_cons q acc)]].
+
+definition factorize : nat \to nat_fact_all \def \lambda n:nat.
+ match n with
+ [ O \Rightarrow nfa_zero
+ | (S n1) \Rightarrow
+ match n1 with
+ [ O \Rightarrow nfa_one
+ | (S n2) \Rightarrow
+ let p \def (max (S(S n2)) (\lambda p:nat.eqb ((S(S n2)) \mod (nth_prime p)) O)) in
+ match p_ord (S(S n2)) (nth_prime p) with
+ [ (pair q r) \Rightarrow
+ nfa_proper (factorize_aux p r (nf_last (pred q)))]]].
+
+let rec defactorize_aux f i \def
+ match f with
+ [ (nf_last n) \Rightarrow (nth_prime i) \sup (S n)
+ | (nf_cons n g) \Rightarrow
+ (nth_prime i) \sup n *(defactorize_aux g (S i))].
+
+definition defactorize : nat_fact_all \to nat \def
+\lambda f : nat_fact_all.
+match f with
+[ nfa_zero \Rightarrow O
+| nfa_one \Rightarrow (S O)
+| (nfa_proper g) \Rightarrow defactorize_aux g O].
+
+theorem lt_O_defactorize_aux:
+ \forall f:nat_fact.
+ \forall i:nat.
+ O < defactorize_aux f i.
+intro.
+elim f
+[1,2:
+ simplify;
+ unfold lt;
+ rewrite > times_n_SO;auto
+ (*apply le_times
+ [ change with (O < nth_prime i).
+ apply lt_O_nth_prime_n
+ |2,3:
+ change with (O < exp (nth_prime i) n);
+ apply lt_O_exp;
+ apply lt_O_nth_prime_n
+ | change with (O < defactorize_aux n1 (S i)).
+ apply H
+ ] *)
+]
+qed.
+
+theorem lt_SO_defactorize_aux: \forall f:nat_fact.\forall i:nat.
+S O < defactorize_aux f i.
+intro.
+elim f
+[ simplify.
+ unfold lt.
+ rewrite > times_n_SO.
+ auto
+ (*apply le_times
+ [ change with (S O < nth_prime i).
+ apply lt_SO_nth_prime_n
+ | change with (O < exp (nth_prime i) n).
+ apply lt_O_exp.
+ apply lt_O_nth_prime_n
+ ]*)
+| simplify.
+ unfold lt.
+ rewrite > times_n_SO.
+ rewrite > sym_times.
+ auto
+ (*apply le_times
+ [ change with (O < exp (nth_prime i) n).
+ apply lt_O_exp.
+ apply lt_O_nth_prime_n
+ | change with (S O < defactorize_aux n1 (S i)).
+ apply H
+ ]*)
+]
+qed.
+
+theorem defactorize_aux_factorize_aux :
+\forall p,n:nat.\forall acc:nat_fact.O < n \to
+((n=(S O) \land p=O) \lor max_prime_factor n < p) \to
+defactorize_aux (factorize_aux p n acc) O = n*(defactorize_aux acc p).
+intro.
+elim p
+[ simplify.
+ elim H1
+ [ elim H2.
+ auto
+ (*rewrite > H3.
+ rewrite > sym_times.
+ apply times_n_SO*)
+ | apply False_ind.
+ apply (not_le_Sn_O (max_prime_factor n) H2)
+ ]
+| simplify.
+ (* generalizing the goal: I guess there exists a better way *)
+ cut (\forall q,r.(pair nat nat q r) = (p_ord_aux n1 n1 (nth_prime n)) \to
+ defactorize_aux match (p_ord_aux n1 n1 (nth_prime n)) with
+ [(pair q r) \Rightarrow (factorize_aux n r (nf_cons q acc))] O =
+ n1*defactorize_aux acc (S n))
+ [ (*invocando auto in questo punto, dopo circa 7 minuti l'esecuzione non era ancora terminata
+ ne' con un errore ne' chiudendo il goal
+ *)
+ apply (Hcut (fst ? ? (p_ord_aux n1 n1 (nth_prime n)))
+ (snd ? ? (p_ord_aux n1 n1 (nth_prime n)))).
+ auto
+ (*apply sym_eq.apply eq_pair_fst_snd*)
+ | intros.
+ rewrite < H3.
+ simplify.
+ cut (n1 = r * (nth_prime n) \sup q)
+ [ rewrite > H
+ [ simplify.
+ auto
+ (*rewrite < assoc_times.
+ rewrite < Hcut.
+ reflexivity.*)
+ | auto
+ (*cut (O < r \lor O = r)
+ [ elim Hcut1
+ [ assumption
+ | absurd (n1 = O)
+ [ rewrite > Hcut.
+ rewrite < H4.
+ reflexivity
+ | unfold Not.
+ intro.
+ apply (not_le_Sn_O O).
+ rewrite < H5 in \vdash (? ? %).
+ assumption
+ ]
+ ]
+ | apply le_to_or_lt_eq.
+ apply le_O_n
+ ]*)
+ | cut ((S O) < r \lor (S O) \nlt r)
+ [ elim Hcut1
+ [ right.
+ apply (p_ord_to_lt_max_prime_factor1 n1 ? q r)
+ [ assumption
+ | elim H2
+ [ elim H5.
+ apply False_ind.
+ apply (not_eq_O_S n).
+ auto
+ (*apply sym_eq.
+ assumption*)
+ | auto
+ (*apply le_S_S_to_le.
+ exact H5*)
+ ]
+ | assumption
+ | assumption
+ ]
+ | cut (r=(S O))
+ [ apply (nat_case n)
+ [ auto
+ (*left.
+ split
+ [ assumption
+ | reflexivity
+ ]*)
+ | intro.
+ right.
+ rewrite > Hcut2.
+ auto
+ (*simplify.
+ unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ | cut (r < (S O) ∨ r=(S O))
+ [ elim Hcut2
+ [ absurd (O=r)
+ [ auto
+ (*apply le_n_O_to_eq.
+ apply le_S_S_to_le.
+ exact H5*)
+ | unfold Not.
+ intro.
+ auto
+ (*cut (O=n1)
+ [ apply (not_le_Sn_O O).
+ rewrite > Hcut3 in ⊢ (? ? %).
+ assumption
+ | rewrite > Hcut.
+ rewrite < H6.
+ reflexivity
+ ]*)
+ ]
+ | assumption
+ ]
+ | auto
+ (*apply (le_to_or_lt_eq r (S O)).
+ apply not_lt_to_le.
+ assumption*)
+ ]
+ ]
+ ]
+ | apply (decidable_lt (S O) r)
+ ]
+ ]
+ | rewrite > sym_times.
+ apply (p_ord_aux_to_exp n1 n1)
+ [ apply lt_O_nth_prime_n
+ | assumption
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem defactorize_factorize: \forall n:nat.defactorize (factorize n) = n.
+intro.
+apply (nat_case n)
+[ reflexivity
+| intro.
+ apply (nat_case m)
+ [ reflexivity
+ | intro.(*CSC: simplify here does something really nasty *)
+ change with
+ (let p \def (max (S(S m1)) (\lambda p:nat.eqb ((S(S m1)) \mod (nth_prime p)) O)) in
+ defactorize (match p_ord (S(S m1)) (nth_prime p) with
+ [ (pair q r) \Rightarrow
+ nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))).
+ intro.
+ (* generalizing the goal; find a better way *)
+ cut (\forall q,r.(pair nat nat q r) = (p_ord (S(S m1)) (nth_prime p)) \to
+ defactorize (match p_ord (S(S m1)) (nth_prime p) with
+ [ (pair q r) \Rightarrow
+ nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1)))
+ [ (*invocando auto qui, dopo circa 300 secondi non si ottiene alcun risultato*)
+ apply (Hcut (fst ? ? (p_ord (S(S m1)) (nth_prime p)))
+ (snd ? ? (p_ord (S(S m1)) (nth_prime p)))).
+ auto
+ (*apply sym_eq.
+ apply eq_pair_fst_snd*)
+ | intros.
+ rewrite < H.
+ simplify.
+ cut ((S(S m1)) = (nth_prime p) \sup q *r)
+ [ cut (O<r)
+ [ rewrite > defactorize_aux_factorize_aux
+ [ (*CSC: simplify here does something really nasty *)
+ change with (r*(nth_prime p) \sup (S (pred q)) = (S(S m1))).
+ cut ((S (pred q)) = q)
+ [ (*invocando auto qui, dopo circa 300 secondi non si ottiene ancora alcun risultato*)
+ rewrite > Hcut2.
+ auto
+ (*rewrite > sym_times.
+ apply sym_eq.
+ apply (p_ord_aux_to_exp (S(S m1)))
+ [ apply lt_O_nth_prime_n
+ | assumption
+ ]*)
+ | (* O < q *)
+ apply sym_eq.
+ apply S_pred.
+ cut (O < q \lor O = q)
+ [ elim Hcut2
+ [ assumption
+ | absurd (nth_prime p \divides S (S m1))
+ [ apply (divides_max_prime_factor_n (S (S m1))).
+ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_S_S.
+ apply le_O_n.*)
+ | cut ((S(S m1)) = r)
+ [ rewrite > Hcut3 in \vdash (? (? ? %)).
+ (*CSC: simplify here does something really nasty *)
+ change with (nth_prime p \divides r \to False).
+ intro.
+ apply (p_ord_aux_to_not_mod_O (S(S m1)) (S(S m1)) (nth_prime p) q r) [ apply lt_SO_nth_prime_n
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | apply le_n
+ | assumption
+ | (*invocando auto qui, dopo circa 300 secondi non si ottiene ancora alcun risultato*)
+ apply divides_to_mod_O
+ [ apply lt_O_nth_prime_n
+ | assumption
+ ]
+ ]
+ | rewrite > times_n_SO in \vdash (? ? ? %).
+ rewrite < sym_times.
+ rewrite > (exp_n_O (nth_prime p)).
+ rewrite > H1 in \vdash (? ? ? (? (? ? %) ?)).
+ assumption
+ ]
+ ]
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ apply le_O_n*)
+ ]
+ ]
+ | assumption
+ | (* e adesso l'ultimo goal. TASSI: che ora non e' piu' l'ultimo :P *)
+ cut ((S O) < r \lor S O \nlt r)
+ [ elim Hcut2
+ [ right.
+ apply (p_ord_to_lt_max_prime_factor1 (S(S m1)) ? q r);auto
+ (*[ unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ | apply le_n
+ | assumption
+ | assumption
+ ]*)
+ | cut (r=(S O))
+ [ apply (nat_case p)
+ [ auto
+ (*left.
+ split
+ [ assumption
+ | reflexivity
+ ]*)
+ | intro.
+ right.
+ rewrite > Hcut3.
+ auto
+ (*simplify.
+ unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ | cut (r \lt (S O) \or r=(S O))
+ [ elim Hcut3
+ [ absurd (O=r);auto
+ (*[ apply le_n_O_to_eq.
+ apply le_S_S_to_le.
+ exact H2
+ | unfold Not.
+ intro.
+ apply (not_le_Sn_O O).
+ rewrite > H3 in \vdash (? ? %).
+ assumption
+ ]*)
+ | assumption
+ ]
+ | auto
+ (*apply (le_to_or_lt_eq r (S O)).
+ apply not_lt_to_le.
+ assumption*)
+ ]
+ ]
+ ]
+ | apply (decidable_lt (S O) r)
+ ]
+ ]
+ | (* O < r *)
+ cut (O < r \lor O = r)
+ [ elim Hcut1
+ [ assumption
+ | apply False_ind.
+ apply (not_eq_O_S (S m1)).
+ rewrite > Hcut.
+ rewrite < H1.
+ auto
+ (*rewrite < times_n_O.
+ reflexivity*)
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ apply le_O_n*)
+ ]
+ ]
+ | (* prova del cut *)
+ goal 20.
+ apply (p_ord_aux_to_exp (S(S m1)));auto
+ (*[ apply lt_O_nth_prime_n
+ | assumption
+ ]*)
+ (* fine prova cut *)
+ ]
+ ]
+ ]
+]
+qed.
+
+let rec max_p f \def
+match f with
+[ (nf_last n) \Rightarrow O
+| (nf_cons n g) \Rightarrow S (max_p g)].
+
+let rec max_p_exponent f \def
+match f with
+[ (nf_last n) \Rightarrow n
+| (nf_cons n g) \Rightarrow max_p_exponent g].
+
+theorem divides_max_p_defactorize: \forall f:nat_fact.\forall i:nat.
+nth_prime ((max_p f)+i) \divides defactorize_aux f i.
+intro.
+elim f
+[ simplify.
+ auto
+ (*apply (witness ? ? ((nth_prime i) \sup n)).
+ reflexivity*)
+| change with
+ (nth_prime (S(max_p n1)+i) \divides
+ (nth_prime i) \sup n *(defactorize_aux n1 (S i))).
+ elim (H (S i)).
+ rewrite > H1.
+ rewrite < sym_times.
+ rewrite > assoc_times.
+ auto
+ (*rewrite < plus_n_Sm.
+ apply (witness ? ? (n2* (nth_prime i) \sup n)).
+ reflexivity*)
+]
+qed.
+
+theorem divides_exp_to_divides:
+\forall p,n,m:nat. prime p \to
+p \divides n \sup m \to p \divides n.
+intros 3.
+elim m
+[ simplify in H1.
+ auto
+ (*apply (transitive_divides p (S O))
+ [ assumption
+ | apply divides_SO_n
+ ]*)
+| cut (p \divides n \lor p \divides n \sup n1)
+ [ elim Hcut
+ [ assumption
+ | auto
+ (*apply H;assumption*)
+ ]
+ | auto
+ (*apply divides_times_to_divides
+ [ assumption
+ | exact H2
+ ]*)
+ ]
+]
+qed.
+
+theorem divides_exp_to_eq:
+\forall p,q,m:nat. prime p \to prime q \to
+p \divides q \sup m \to p = q.
+intros.
+unfold prime in H1.
+elim H1.
+apply H4
+[ apply (divides_exp_to_divides p q m);assumption
+| (*invocando auto in questo punto, dopo piu' di 8 minuti la computazione non
+ * era ancora terminata.
+ *)
+ unfold prime in H.
+ (*invocando auto anche in questo punto, dopo piu' di 10 minuti la computazione
+ * non era ancora terminata.
+ *)
+ elim H.
+ assumption
+]
+qed.
+
+theorem not_divides_defactorize_aux: \forall f:nat_fact. \forall i,j:nat.
+i < j \to nth_prime i \ndivides defactorize_aux f j.
+intro.
+elim f
+[ change with
+ (nth_prime i \divides (nth_prime j) \sup (S n) \to False).
+ intro.
+ absurd ((nth_prime i) = (nth_prime j))
+ [ apply (divides_exp_to_eq ? ? (S n));auto
+ (*[ apply prime_nth_prime
+ | apply prime_nth_prime
+ | assumption
+ ]*)
+ | unfold Not.
+ intro.
+ cut (i = j)
+ [ apply (not_le_Sn_n i).
+ rewrite > Hcut in \vdash (? ? %).
+ assumption
+ | apply (injective_nth_prime ? ? H2)
+ ]
+ ]
+| unfold Not.
+ simplify.
+ intro.
+ cut (nth_prime i \divides (nth_prime j) \sup n
+ \lor nth_prime i \divides defactorize_aux n1 (S j))
+ [ elim Hcut
+ [ absurd ((nth_prime i) = (nth_prime j))
+ [ apply (divides_exp_to_eq ? ? n);auto
+ (*[ apply prime_nth_prime
+ | apply prime_nth_prime
+ | assumption
+ ]*)
+ | unfold Not.
+ intro.
+ cut (i = j)
+ [ apply (not_le_Sn_n i).
+ rewrite > Hcut1 in \vdash (? ? %).
+ assumption
+ | apply (injective_nth_prime ? ? H4)
+ ]
+ ]
+ | apply (H i (S j))
+ [ auto
+ (*apply (trans_lt ? j)
+ [ assumption
+ | unfold lt.
+ apply le_n
+ ]*)
+ | assumption
+ ]
+ ]
+ | auto
+ (*apply divides_times_to_divides.
+ apply prime_nth_prime.
+ assumption*)
+ ]
+]
+qed.
+
+lemma not_eq_nf_last_nf_cons: \forall g:nat_fact.\forall n,m,i:nat.
+\lnot (defactorize_aux (nf_last n) i= defactorize_aux (nf_cons m g) i).
+intros.
+change with
+(exp (nth_prime i) (S n) = defactorize_aux (nf_cons m g) i \to False).
+intro.
+cut (S(max_p g)+i= i)
+[ apply (not_le_Sn_n i).
+ rewrite < Hcut in \vdash (? ? %). (*chiamando auto qui da uno strano errore "di tipo"*)
+ simplify.
+ auto
+ (*apply le_S_S.
+ apply le_plus_n*)
+| apply injective_nth_prime.
+ apply (divides_exp_to_eq ? ? (S n))
+ [ apply prime_nth_prime
+ | apply prime_nth_prime
+ | rewrite > H.
+ change with (divides (nth_prime ((max_p (nf_cons m g))+i))
+ (defactorize_aux (nf_cons m g) i)).
+ apply divides_max_p_defactorize
+ ]
+]
+qed.
+
+lemma not_eq_nf_cons_O_nf_cons: \forall f,g:nat_fact.\forall n,i:nat.
+\lnot (defactorize_aux (nf_cons O f) i= defactorize_aux (nf_cons (S n) g) i).
+intros.
+simplify.
+unfold Not.
+rewrite < plus_n_O.
+intro.
+apply (not_divides_defactorize_aux f i (S i) ?)
+[ auto
+ (*unfold lt.
+ apply le_n*)
+| auto
+ (*rewrite > H.
+ rewrite > assoc_times.
+ apply (witness ? ? ((exp (nth_prime i) n)*(defactorize_aux g (S i)))).
+ reflexivity*)
+]
+qed.
+
+theorem eq_defactorize_aux_to_eq: \forall f,g:nat_fact.\forall i:nat.
+defactorize_aux f i = defactorize_aux g i \to f = g.
+intro.
+elim f
+[ generalize in match H.
+ elim g
+ [ apply eq_f.
+ apply inj_S.
+ apply (inj_exp_r (nth_prime i))
+ [ apply lt_SO_nth_prime_n
+ | (*qui auto non conclude il goal attivo*)
+ assumption
+ ]
+ | apply False_ind.
+ (*auto chiamato qui NON conclude il goal attivo*)
+ apply (not_eq_nf_last_nf_cons n2 n n1 i H2)
+ ]
+| generalize in match H1.
+ elim g
+ [ apply False_ind.
+ apply (not_eq_nf_last_nf_cons n1 n2 n i).
+ auto
+ (*apply sym_eq.
+ assumption*)
+ | simplify in H3.
+ generalize in match H3.
+ apply (nat_elim2 (\lambda n,n2.
+ ((nth_prime i) \sup n)*(defactorize_aux n1 (S i)) =
+ ((nth_prime i) \sup n2)*(defactorize_aux n3 (S i)) \to
+ nf_cons n n1 = nf_cons n2 n3))
+ [ intro.
+ elim n4
+ [ auto
+ (*apply eq_f.
+ apply (H n3 (S i))
+ simplify in H4.
+ rewrite > plus_n_O.
+ rewrite > (plus_n_O (defactorize_aux n3 (S i))).
+ assumption*)
+ | apply False_ind.
+ apply (not_eq_nf_cons_O_nf_cons n1 n3 n5 i).
+ (*auto chiamato qui NON chiude il goal attivo*)
+ assumption
+ ]
+ | intros.
+ apply False_ind.
+ apply (not_eq_nf_cons_O_nf_cons n3 n1 n4 i).
+ apply sym_eq.
+ (*auto chiamato qui non chiude il goal*)
+ assumption
+ | intros.
+ cut (nf_cons n4 n1 = nf_cons m n3)
+ [ cut (n4=m)
+ [ cut (n1=n3)
+ [ auto
+ (*rewrite > Hcut1.
+ rewrite > Hcut2.
+ reflexivity*)
+ | change with
+ (match nf_cons n4 n1 with
+ [ (nf_last m) \Rightarrow n1
+ | (nf_cons m g) \Rightarrow g ] = n3).
+ rewrite > Hcut.
+ auto
+ (*simplify.
+ reflexivity*)
+ ]
+ | change with
+ (match nf_cons n4 n1 with
+ [ (nf_last m) \Rightarrow m
+ | (nf_cons m g) \Rightarrow m ] = m).
+ (*invocando auto qui, dopo circa 8 minuti la computazione non era ancora terminata*)
+ rewrite > Hcut.
+ auto
+ (*simplify.
+ reflexivity*)
+ ]
+ | apply H4.
+ simplify in H5.
+ apply (inj_times_r1 (nth_prime i))
+ [ apply lt_O_nth_prime_n
+ | rewrite < assoc_times.
+ rewrite < assoc_times.
+ assumption
+ ]
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem injective_defactorize_aux: \forall i:nat.
+injective nat_fact nat (\lambda f.defactorize_aux f i).
+simplify.
+intros.
+apply (eq_defactorize_aux_to_eq x y i H).
+qed.
+
+theorem injective_defactorize:
+injective nat_fact_all nat defactorize.
+unfold injective.
+change with (\forall f,g.defactorize f = defactorize g \to f=g).
+intro.
+elim f
+[ generalize in match H.
+ elim g
+ [ (* zero - zero *)
+ reflexivity
+ | (* zero - one *)
+ simplify in H1.
+ apply False_ind.
+ apply (not_eq_O_S O H1)
+ | (* zero - proper *)
+ simplify in H1.
+ apply False_ind.
+ apply (not_le_Sn_n O).
+ rewrite > H1 in \vdash (? ? %).
+ auto
+ (*change with (O < defactorize_aux n O).
+ apply lt_O_defactorize_aux*)
+ ]
+| generalize in match H.
+ elim g
+ [ (* one - zero *)
+ simplify in H1.
+ apply False_ind.
+ auto
+ (*apply (not_eq_O_S O).
+ apply sym_eq.
+ assumption*)
+ | (* one - one *)
+ reflexivity
+ | (* one - proper *)
+ simplify in H1.
+ apply False_ind.
+ apply (not_le_Sn_n (S O)).
+ rewrite > H1 in \vdash (? ? %).
+ auto
+ (*change with ((S O) < defactorize_aux n O).
+ apply lt_SO_defactorize_aux*)
+ ]
+| generalize in match H.
+ elim g
+ [ (* proper - zero *)
+ simplify in H1.
+ apply False_ind.
+ apply (not_le_Sn_n O).
+ rewrite < H1 in \vdash (? ? %).
+ auto
+ (*change with (O < defactorize_aux n O).
+ apply lt_O_defactorize_aux.*)
+ | (* proper - one *)
+ simplify in H1.
+ apply False_ind.
+ apply (not_le_Sn_n (S O)).
+ rewrite < H1 in \vdash (? ? %).
+ auto
+ (*change with ((S O) < defactorize_aux n O).
+ apply lt_SO_defactorize_aux.*)
+ | (* proper - proper *)
+ apply eq_f.
+ apply (injective_defactorize_aux O).
+ (*invocata qui la tattica auto NON chiude il goal, chiuso invece
+ *da exact H1
+ *)
+ exact H1
+ ]
+]
+qed.
+
+theorem factorize_defactorize:
+\forall f,g: nat_fact_all. factorize (defactorize f) = f.
+intros.
+auto.
+(*apply injective_defactorize.
+apply defactorize_factorize.
+*)
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/fermat_little_theorem".
+
+include "nat/exp.ma".
+include "nat/gcd.ma".
+include "nat/permutation.ma".
+include "nat/congruence.ma".
+
+theorem permut_S_mod: \forall n:nat. permut (S_mod (S n)) n.
+intro.
+unfold permut.
+split
+[ intros.
+ unfold S_mod.
+ auto
+ (*apply le_S_S_to_le.
+ change with ((S i) \mod (S n) < S n).
+ apply lt_mod_m_m.
+ unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+| unfold injn.
+ intros.
+ apply inj_S.
+ rewrite < (lt_to_eq_mod i (S n))
+ [ rewrite < (lt_to_eq_mod j (S n))
+ [ cut (i < n \lor i = n)
+ [ cut (j < n \lor j = n)
+ [ elim Hcut
+ [ elim Hcut1
+ [ (* i < n, j< n *)
+ rewrite < mod_S
+ [ rewrite < mod_S
+ [ (*qui auto non chiude il goal, chiuso invece dalla tattica
+ * apply H2
+ *)
+ apply H2
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | rewrite > lt_to_eq_mod;
+ auto(*unfold lt;apply le_S_S;assumption*)
+ ]
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | rewrite > lt_to_eq_mod
+ [ auto
+ (*unfold lt.
+ apply le_S_S.
+ assumption*)
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ assumption*)
+ ]
+ ]
+ | (* i < n, j=n *)
+ unfold S_mod in H2.
+ simplify.
+ apply False_ind.
+ apply (not_eq_O_S (i \mod (S n))).
+ apply sym_eq.
+ rewrite < (mod_n_n (S n))
+ [ rewrite < H4 in \vdash (? ? ? (? %?)).
+ rewrite < mod_S
+ [ assumption
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | rewrite > lt_to_eq_mod;
+ auto;(*unfold lt;apply le_S_S;assumption*)
+ ]
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ ]
+ | (* i = n, j < n *)
+ elim Hcut1
+ [ apply False_ind.
+ apply (not_eq_O_S (j \mod (S n))).
+ rewrite < (mod_n_n (S n))
+ [ rewrite < H3 in \vdash (? ? (? %?) ?).
+ rewrite < mod_S
+ [ assumption
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | rewrite > lt_to_eq_mod;
+ auto(*unfold lt;apply le_S_S;assumption*)
+ ]
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ |(* i = n, j= n*)
+ auto
+ (*rewrite > H3.
+ rewrite > H4.
+ reflexivity*)
+ ]
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ assumption*)
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ assumption*)
+ ]
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ assumption*)
+ ]
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ assumption*)
+ ]
+]
+qed.
+
+(*
+theorem eq_fact_pi: \forall n,m:nat. n < m \to n! = pi n (S_mod m).
+intro.elim n.
+simplify.reflexivity.
+change with (S n1)*n1!=(S_mod m n1)*(pi n1 (S_mod m)).
+unfold S_mod in \vdash (? ? ? (? % ?)).
+rewrite > lt_to_eq_mod.
+apply eq_f.apply H.apply (trans_lt ? (S n1)).
+simplify. apply le_n.assumption.assumption.
+qed.
+*)
+
+theorem prime_to_not_divides_fact: \forall p:nat. prime p \to \forall n:nat.
+n \lt p \to \not divides p n!.
+intros 3.
+elim n
+[ unfold Not.
+ intros.
+ apply (lt_to_not_le (S O) p)
+ [ unfold prime in H.
+ elim H.
+ assumption
+ | auto
+ (*apply divides_to_le.
+ unfold lt.
+ apply le_n.
+ assumption*)
+ ]
+| change with (divides p ((S n1)*n1!) \to False).
+ intro.
+ cut (divides p (S n1) \lor divides p n1!)
+ [ elim Hcut
+ [ apply (lt_to_not_le (S n1) p)
+ [ assumption
+ | auto
+ (*apply divides_to_le
+ [ unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ | assumption
+ ]*)
+ ]
+ | auto
+ (*apply H1
+ [ apply (trans_lt ? (S n1))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]
+ | assumption
+ ]*)
+ ]
+ | auto
+ (*apply divides_times_to_divides;
+ assumption*)
+ ]
+]
+qed.
+
+theorem permut_mod: \forall p,a:nat. prime p \to
+\lnot divides p a\to permut (\lambda n.(mod (a*n) p)) (pred p).
+unfold permut.
+intros.
+split
+[ intros.
+ apply le_S_S_to_le.
+ apply (trans_le ? p)
+ [ change with (mod (a*i) p < p).
+ apply lt_mod_m_m.
+ unfold prime in H.
+ elim H.
+ auto
+ (*unfold lt.
+ apply (trans_le ? (S (S O)))
+ [ apply le_n_Sn
+ | assumption
+ ]*)
+ | rewrite < S_pred
+ [ apply le_n
+ | unfold prime in H.
+ elim H.
+ auto
+ (*apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]*)
+ ]
+ ]
+| unfold injn.
+ intros.
+ apply (nat_compare_elim i j)
+ [ (* i < j *)
+ intro.
+ absurd (j-i \lt p)
+ [ unfold lt.
+ rewrite > (S_pred p)
+ [ auto
+ (*apply le_S_S.
+ apply le_plus_to_minus.
+ apply (trans_le ? (pred p))
+ [ assumption
+ | rewrite > sym_plus.
+ apply le_plus_n
+ ]*)
+ | auto
+ (*unfold prime in H.
+ elim H.
+ apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]*)
+ ]
+ | apply (le_to_not_lt p (j-i)).
+ apply divides_to_le
+ [ auto
+ (*unfold lt.
+ apply le_SO_minus.
+ assumption*)
+ | cut (divides p a \lor divides p (j-i))
+ [ elim Hcut
+ [ apply False_ind.
+ auto
+ (*apply H1.
+ assumption*)
+ | assumption
+ ]
+ | apply divides_times_to_divides
+ [ assumption
+ | rewrite > distr_times_minus.
+ apply eq_mod_to_divides
+ [ auto
+ (*unfold prime in H.
+ elim H.
+ apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]*)
+ | auto
+ (*apply sym_eq.
+ apply H4*)
+ ]
+ ]
+ ]
+ ]
+ ]
+ |(* i = j *)
+ auto
+ (*intro.
+ assumption*)
+ | (* j < i *)
+ intro.
+ absurd (i-j \lt p)
+ [ unfold lt.
+ rewrite > (S_pred p)
+ [ auto
+ (*apply le_S_S.
+ apply le_plus_to_minus.
+ apply (trans_le ? (pred p))
+ [ assumption
+ | rewrite > sym_plus.
+ apply le_plus_n
+ ]*)
+ | auto
+ (*unfold prime in H.
+ elim H.
+ apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]*)
+ ]
+ | apply (le_to_not_lt p (i-j)).
+ apply divides_to_le
+ [ auto
+ (*unfold lt.
+ apply le_SO_minus.
+ assumption*)
+ | cut (divides p a \lor divides p (i-j))
+ [ elim Hcut
+ [ apply False_ind.
+ auto
+ (*apply H1.
+ assumption*)
+ | assumption
+ ]
+ | apply divides_times_to_divides
+ [ assumption
+ | rewrite > distr_times_minus.
+ apply eq_mod_to_divides
+ [ auto
+ (*unfold prime in H.
+ elim H.
+ apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]*)
+ | apply H4
+ ]
+ ]
+ ]
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem congruent_exp_pred_SO: \forall p,a:nat. prime p \to \lnot divides p a \to
+congruent (exp a (pred p)) (S O) p.
+intros.
+cut (O < a)
+[ cut (O < p)
+ [ cut (O < pred p)
+ [ apply divides_to_congruent
+ [ assumption
+ | auto
+ (*change with (O < exp a (pred p)).
+ apply lt_O_exp.
+ assumption*)
+ | cut (divides p (exp a (pred p)-(S O)) \lor divides p (pred p)!)
+ [ elim Hcut3
+ [ assumption
+ | apply False_ind.
+ apply (prime_to_not_divides_fact p H (pred p))
+ [ unfold lt.
+ auto
+ (*rewrite < (S_pred ? Hcut1).
+ apply le_n*)
+ | assumption
+ ]
+ ]
+ | apply divides_times_to_divides
+ [ assumption
+ | rewrite > times_minus_l.
+ rewrite > (sym_times (S O)).
+ rewrite < times_n_SO.
+ rewrite > (S_pred (pred p) Hcut2).
+ rewrite > eq_fact_pi.
+ rewrite > exp_pi_l.
+ apply congruent_to_divides
+ [ assumption
+ | apply (transitive_congruent p ?
+ (pi (pred (pred p)) (\lambda m. a*m \mod p) (S O)))
+ [ apply (congruent_pi (\lambda m. a*m)).
+ assumption
+ | cut (pi (pred(pred p)) (\lambda m.m) (S O)
+ = pi (pred(pred p)) (\lambda m.a*m \mod p) (S O))
+ [ rewrite > Hcut3.
+ apply congruent_n_n
+ | rewrite < eq_map_iter_i_pi.
+ rewrite < eq_map_iter_i_pi.
+ apply (permut_to_eq_map_iter_i ? ? ? ? ? (λm.m))
+ [ apply assoc_times
+ | (*NB qui auto non chiude il goal, chiuso invece dalla sola
+ ( tattica apply sys_times
+ *)
+ apply sym_times
+ | rewrite < plus_n_Sm.
+ rewrite < plus_n_O.
+ rewrite < (S_pred ? Hcut2).
+ auto
+ (*apply permut_mod
+ [ assumption
+ | assumption
+ ]*)
+ | intros.
+ cut (m=O)
+ [ auto
+ (*rewrite > Hcut3.
+ rewrite < times_n_O.
+ apply mod_O_n.*)
+ | auto
+ (*apply sym_eq.
+ apply le_n_O_to_eq.
+ apply le_S_S_to_le.
+ assumption*)
+ ]
+ ]
+ ]
+ ]
+ ]
+ ]
+ ]
+ ]
+ | unfold lt.
+ apply le_S_S_to_le.
+ rewrite < (S_pred ? Hcut1).
+ unfold prime in H.
+ elim H.
+ assumption
+ ]
+ | unfold prime in H.
+ elim H.
+ auto
+ (*apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]*)
+ ]
+| cut (O < a \lor O = a)
+ [ elim Hcut
+ [ assumption
+ | apply False_ind.
+ apply H1.
+ rewrite < H2.
+ auto
+ (*apply (witness ? ? O).
+ apply times_n_O*)
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ apply le_O_n*)
+ ]
+]
+qed.
+
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / Matita is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/gcd".
+
+include "nat/primes.ma".
+
+let rec gcd_aux p m n: nat \def
+match divides_b n m with
+[ true \Rightarrow n
+| false \Rightarrow
+ match p with
+ [O \Rightarrow n
+ |(S q) \Rightarrow gcd_aux q n (m \mod n)]].
+
+definition gcd : nat \to nat \to nat \def
+\lambda n,m:nat.
+ match leb n m with
+ [ true \Rightarrow
+ match n with
+ [ O \Rightarrow m
+ | (S p) \Rightarrow gcd_aux (S p) m (S p) ]
+ | false \Rightarrow
+ match m with
+ [ O \Rightarrow n
+ | (S p) \Rightarrow gcd_aux (S p) n (S p) ]].
+
+theorem divides_mod: \forall p,m,n:nat. O < n \to p \divides m \to p \divides n \to
+p \divides (m \mod n).
+intros.elim H1.elim H2.
+apply (witness ? ? (n2 - n1*(m / n))).
+(*apply witness[|*)
+ rewrite > distr_times_minus.
+ rewrite < H3 in \vdash (? ? ? (? % ?)).
+ rewrite < assoc_times.
+ rewrite < H4 in \vdash (? ? ? (? ? (? % ?))).
+ apply sym_eq.
+ apply plus_to_minus.
+ rewrite > sym_times.
+ auto.
+ (*letin x \def div.
+ rewrite < (div_mod ? ? H).
+ reflexivity.*)
+(*]*)
+qed.
+
+theorem divides_mod_to_divides: \forall p,m,n:nat. O < n \to
+p \divides (m \mod n) \to p \divides n \to p \divides m.
+intros.
+elim H1.
+elim H2.
+apply (witness p m ((n1*(m / n))+n2)).
+rewrite > distr_times_plus.
+rewrite < H3.
+rewrite < assoc_times.
+rewrite < H4.
+rewrite < sym_times.
+auto.
+(*apply div_mod.
+assumption.*)
+qed.
+
+
+theorem divides_gcd_aux_mn: \forall p,m,n. O < n \to n \le m \to n \le p \to
+gcd_aux p m n \divides m \land gcd_aux p m n \divides n.
+intro.
+elim p
+[ absurd (O < n);auto
+ (*[ assumption
+ | apply le_to_not_lt.
+ assumption
+ ]*)
+| cut ((n1 \divides m) \lor (n1 \ndivides m))
+ [ simplify.
+ elim Hcut
+ [ rewrite > divides_to_divides_b_true
+ [ simplify.
+ auto
+ (*split
+ [ assumption
+ | apply (witness n1 n1 (S O)).
+ apply times_n_SO
+ ]*)
+ | assumption
+ | assumption
+ ]
+ | rewrite > not_divides_to_divides_b_false
+ [ simplify.
+ cut (gcd_aux n n1 (m \mod n1) \divides n1 \land
+ gcd_aux n n1 (m \mod n1) \divides mod m n1)
+ [ elim Hcut1.
+ auto
+ (*split
+ [ apply (divides_mod_to_divides ? ? n1);assumption
+ | assumption
+ ]*)
+ | apply H
+ [ cut (O \lt m \mod n1 \lor O = mod m n1)
+ [ elim Hcut1
+ [ assumption
+ | apply False_ind.
+ auto
+ (*apply H4.
+ apply mod_O_to_divides
+ [ assumption
+ | apply sym_eq.
+ assumption
+ ]*)
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ apply le_O_n*)
+ ]
+ | auto
+ (*apply lt_to_le.
+ apply lt_mod_m_m.
+ assumption*)
+ | apply le_S_S_to_le.
+ apply (trans_le ? n1);auto
+ (*[ auto.change with (m \mod n1 < n1).
+ apply lt_mod_m_m.
+ assumption
+ | assumption
+ ]*)
+ ]
+ ]
+ | assumption
+ | assumption
+ ]
+ ]
+ | auto
+ (*apply (decidable_divides n1 m).
+ assumption*)
+ ]
+]
+qed.
+
+theorem divides_gcd_nm: \forall n,m.
+gcd n m \divides m \land gcd n m \divides n.
+intros.
+(*CSC: simplify simplifies too much because of a redex in gcd *)
+change with
+(match leb n m with
+ [ true \Rightarrow
+ match n with
+ [ O \Rightarrow m
+ | (S p) \Rightarrow gcd_aux (S p) m (S p) ]
+ | false \Rightarrow
+ match m with
+ [ O \Rightarrow n
+ | (S p) \Rightarrow gcd_aux (S p) n (S p) ] ] \divides m
+\land
+match leb n m with
+ [ true \Rightarrow
+ match n with
+ [ O \Rightarrow m
+ | (S p) \Rightarrow gcd_aux (S p) m (S p) ]
+ | false \Rightarrow
+ match m with
+ [ O \Rightarrow n
+ | (S p) \Rightarrow gcd_aux (S p) n (S p) ] ] \divides n).
+apply (leb_elim n m)
+[ apply (nat_case1 n)
+ [ simplify.
+ intros.
+ auto
+ (*split
+ [ apply (witness m m (S O)).
+ apply times_n_SO
+ | apply (witness m O O).
+ apply times_n_O
+ ]*)
+ | intros.
+ change with
+ (gcd_aux (S m1) m (S m1) \divides m
+ \land
+ gcd_aux (S m1) m (S m1) \divides (S m1)).
+ auto
+ (*apply divides_gcd_aux_mn
+ [ unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ | assumption
+ | apply le_n
+ ]*)
+ ]
+| simplify.
+ intro.
+ apply (nat_case1 m)
+ [ simplify.
+ intros.
+ auto
+ (*split
+ [ apply (witness n O O).
+ apply times_n_O
+ | apply (witness n n (S O)).
+ apply times_n_SO
+ ]*)
+ | intros.
+ change with
+ (gcd_aux (S m1) n (S m1) \divides (S m1)
+ \land
+ gcd_aux (S m1) n (S m1) \divides n).
+ cut (gcd_aux (S m1) n (S m1) \divides n
+ \land
+ gcd_aux (S m1) n (S m1) \divides S m1)
+ [ elim Hcut.
+ auto
+ (*split;assumption*)
+ | apply divides_gcd_aux_mn
+ [ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | apply not_lt_to_le.
+ unfold Not.
+ unfold lt.
+ intro.
+ apply H.
+ rewrite > H1.
+ auto
+ (*apply (trans_le ? (S n))
+ [ apply le_n_Sn
+ | assumption
+ ]*)
+ | apply le_n
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem divides_gcd_n: \forall n,m. gcd n m \divides n.
+intros.
+exact (proj2 ? ? (divides_gcd_nm n m)). (*auto non termina la dimostrazione*)
+qed.
+
+theorem divides_gcd_m: \forall n,m. gcd n m \divides m.
+intros.
+exact (proj1 ? ? (divides_gcd_nm n m)). (*auto non termina la dimostrazione*)
+qed.
+
+theorem divides_gcd_aux: \forall p,m,n,d. O < n \to n \le m \to n \le p \to
+d \divides m \to d \divides n \to d \divides gcd_aux p m n.
+intro.
+elim p
+[ absurd (O < n);auto
+ (*[ assumption
+ | apply le_to_not_lt.
+ assumption
+ ]*)
+| simplify.
+ cut (n1 \divides m \lor n1 \ndivides m)
+ [ elim Hcut.
+ rewrite > divides_to_divides_b_true;auto.
+ (*[ simplify.
+ assumption.
+ | assumption.
+ | assumption.
+ ]*)
+ rewrite > not_divides_to_divides_b_false
+ [ simplify.
+ apply H
+ [ cut (O \lt m \mod n1 \lor O = m \mod n1)
+ [ elim Hcut1
+ [ assumption
+ |
+ absurd (n1 \divides m);auto
+ (*[ apply mod_O_to_divides
+ [ assumption.
+ | apply sym_eq.assumption.
+ ]
+ | assumption
+ ]*)
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ apply le_O_n*)
+ ]
+ | auto
+ (*apply lt_to_le.
+ apply lt_mod_m_m.
+ assumption*)
+ | apply le_S_S_to_le.
+ auto
+ (*apply (trans_le ? n1)
+ [ change with (m \mod n1 < n1).
+ apply lt_mod_m_m.
+ assumption
+ | assumption
+ ]*)
+ | assumption
+ | auto
+ (*apply divides_mod;
+ assumption*)
+ ]
+ | assumption
+ | assumption
+ ]
+ | auto
+ (*apply (decidable_divides n1 m).
+ assumption*)
+ ]
+]qed.
+
+theorem divides_d_gcd: \forall m,n,d.
+d \divides m \to d \divides n \to d \divides gcd n m.
+intros.
+(*CSC: here simplify simplifies too much because of a redex in gcd *)
+change with
+(d \divides
+match leb n m with
+ [ true \Rightarrow
+ match n with
+ [ O \Rightarrow m
+ | (S p) \Rightarrow gcd_aux (S p) m (S p) ]
+ | false \Rightarrow
+ match m with
+ [ O \Rightarrow n
+ | (S p) \Rightarrow gcd_aux (S p) n (S p) ]]).
+apply (leb_elim n m)
+[ apply (nat_case1 n)
+ [ simplify.
+ intros.
+ assumption
+ | intros.
+ change with (d \divides gcd_aux (S m1) m (S m1)).
+ apply divides_gcd_aux
+ [ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n.*)
+ | assumption.
+ | apply le_n. (*chiude il goal anche con auto*)
+ | assumption.
+ | rewrite < H2.
+ assumption
+ ]
+ ]
+| apply (nat_case1 m)
+ [ simplify.
+ intros.
+ assumption
+ | intros.
+ change with (d \divides gcd_aux (S m1) n (S m1)).
+ apply divides_gcd_aux
+ [ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | auto
+ (*apply lt_to_le.
+ apply not_le_to_lt.
+ assumption*)
+ | apply le_n (*chiude il goal anche con auto*)
+ | assumption
+ | rewrite < H2.
+ assumption
+ ]
+ ]
+]
+qed.
+
+theorem eq_minus_gcd_aux: \forall p,m,n.O < n \to n \le m \to n \le p \to
+\exists a,b. a*n - b*m = gcd_aux p m n \lor b*m - a*n = gcd_aux p m n.
+intro.
+elim p
+[ absurd (O < n);auto
+ (*[ assumption
+ | apply le_to_not_lt
+ assumption.
+ ]*)
+| cut (O < m)
+ [ cut (n1 \divides m \lor n1 \ndivides m)
+ [ simplify.
+ elim Hcut1
+ [ rewrite > divides_to_divides_b_true
+ [ simplify.
+ apply (ex_intro ? ? (S O)).
+ apply (ex_intro ? ? O).
+ auto
+ (*left.
+ simplify.
+ rewrite < plus_n_O.
+ apply sym_eq.apply minus_n_O*)
+ | assumption
+ | assumption
+ ]
+ | rewrite > not_divides_to_divides_b_false
+ [ change with
+ (\exists a,b.
+ a*n1 - b*m = gcd_aux n n1 (m \mod n1)
+ \lor
+ b*m - a*n1 = gcd_aux n n1 (m \mod n1)).
+ cut
+ (\exists a,b.
+ a*(m \mod n1) - b*n1= gcd_aux n n1 (m \mod n1)
+ \lor
+ b*n1 - a*(m \mod n1) = gcd_aux n n1 (m \mod n1))
+ [ elim Hcut2.
+ elim H5.
+ elim H6
+ [ (* first case *)
+ rewrite < H7.
+ apply (ex_intro ? ? (a1+a*(m / n1))).
+ apply (ex_intro ? ? a).
+ right.
+ rewrite < sym_plus.
+ rewrite < (sym_times n1).
+ rewrite > distr_times_plus.
+ rewrite > (sym_times n1).
+ rewrite > (sym_times n1).
+ rewrite > (div_mod m n1) in \vdash (? ? (? % ?) ?);auto
+ (*[ rewrite > assoc_times.
+ rewrite < sym_plus.
+ rewrite > distr_times_plus.
+ rewrite < eq_minus_minus_minus_plus.
+ rewrite < sym_plus.auto.
+ rewrite < plus_minus
+ [ rewrite < minus_n_n.
+ reflexivity
+ | apply le_n
+ ]
+ | assumption
+ ]*)
+ | (* second case *)
+ rewrite < H7.
+ apply (ex_intro ? ? (a1+a*(m / n1))).
+ apply (ex_intro ? ? a).
+ left.
+ (* clear Hcut2.clear H5.clear H6.clear H. *)
+ rewrite > sym_times.
+ rewrite > distr_times_plus.
+ rewrite > sym_times.
+ rewrite > (sym_times n1).
+ rewrite > (div_mod m n1) in \vdash (? ? (? ? %) ?)
+ [ rewrite > distr_times_plus.
+ rewrite > assoc_times.
+ rewrite < eq_minus_minus_minus_plus.
+ auto
+ (*rewrite < sym_plus.
+ rewrite < plus_minus
+ [ rewrite < minus_n_n.
+ reflexivity
+ | apply le_n
+ ]*)
+ | assumption
+ ]
+ ]
+ | apply (H n1 (m \mod n1))
+ [ cut (O \lt m \mod n1 \lor O = m \mod n1)
+ [ elim Hcut2
+ [ assumption
+ | absurd (n1 \divides m);auto
+ (*[ apply mod_O_to_divides
+ [ assumption
+ | symmetry.
+ assumption
+ ]
+ | assumption
+ ]*)
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ apply le_O_n*)
+ ]
+ | auto
+ (*apply lt_to_le.
+ apply lt_mod_m_m.
+ assumption*)
+ | apply le_S_S_to_le.
+ auto
+ (*apply (trans_le ? n1)
+ [ change with (m \mod n1 < n1).
+ apply lt_mod_m_m.
+ assumption
+ | assumption
+ ]*)
+ ]
+ ]
+ | assumption
+ | assumption
+ ]
+ ]
+ | auto
+ (*apply (decidable_divides n1 m).
+ assumption*)
+ ]
+ | auto
+ (*apply (lt_to_le_to_lt ? n1);assumption *)
+ ]
+]
+qed.
+
+theorem eq_minus_gcd:
+ \forall m,n.\exists a,b.a*n - b*m = (gcd n m) \lor b*m - a*n = (gcd n m).
+intros.
+unfold gcd.
+apply (leb_elim n m)
+[ apply (nat_case1 n)
+ [ simplify.
+ intros.
+ apply (ex_intro ? ? O).
+ apply (ex_intro ? ? (S O)).
+ auto
+ (*right.simplify.
+ rewrite < plus_n_O.
+ apply sym_eq.
+ apply minus_n_O*)
+ | intros.
+ change with
+ (\exists a,b.
+ a*(S m1) - b*m = (gcd_aux (S m1) m (S m1))
+ \lor b*m - a*(S m1) = (gcd_aux (S m1) m (S m1))).
+ auto
+ (*apply eq_minus_gcd_aux
+ [ unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ | assumption
+ | apply le_n
+ ]*)
+ ]
+| apply (nat_case1 m)
+ [ simplify.
+ intros.
+ apply (ex_intro ? ? (S O)).
+ apply (ex_intro ? ? O).
+ auto
+ (*left.simplify.
+ rewrite < plus_n_O.
+ apply sym_eq.
+ apply minus_n_O*)
+ | intros.
+ change with
+ (\exists a,b.
+ a*n - b*(S m1) = (gcd_aux (S m1) n (S m1))
+ \lor b*(S m1) - a*n = (gcd_aux (S m1) n (S m1))).
+ cut
+ (\exists a,b.
+ a*(S m1) - b*n = (gcd_aux (S m1) n (S m1))
+ \lor
+ b*n - a*(S m1) = (gcd_aux (S m1) n (S m1)))
+ [ elim Hcut.
+ elim H2.
+ elim H3;apply (ex_intro ? ? a1);auto
+ (*[ apply (ex_intro ? ? a1).
+ apply (ex_intro ? ? a).
+ right.
+ assumption
+ | apply (ex_intro ? ? a1).
+ apply (ex_intro ? ? a).
+ left.
+ assumption
+ ]*)
+ | apply eq_minus_gcd_aux;auto
+ (*[ unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ | auto.apply lt_to_le.
+ apply not_le_to_lt.
+ assumption
+ | apply le_n.
+ ]*)
+ ]
+ ]
+]
+qed.
+
+(* some properties of gcd *)
+
+theorem gcd_O_n: \forall n:nat. gcd O n = n.
+auto.
+(*intro.simplify.reflexivity.*)
+qed.
+
+theorem gcd_O_to_eq_O:\forall m,n:nat. (gcd m n) = O \to
+m = O \land n = O.
+intros.
+cut (O \divides n \land O \divides m)
+[ elim Hcut.
+ auto
+ (*elim H2.
+ split
+ [ assumption
+ | elim H1.
+ assumption
+ ]*)
+| rewrite < H.
+ apply divides_gcd_nm
+]
+qed.
+
+theorem lt_O_gcd:\forall m,n:nat. O < n \to O < gcd m n.
+intros.
+auto.
+(*apply (nat_case1 (gcd m n))
+[ intros.
+ generalize in match (gcd_O_to_eq_O m n H1).
+ intros.
+ elim H2.
+ rewrite < H4 in \vdash (? ? %).
+ assumption
+| intros.
+ unfold lt.
+ apply le_S_S.
+ apply le_O_n
+]*)
+qed.
+
+theorem gcd_n_n: \forall n.gcd n n = n.
+intro.
+auto.
+(*elim n
+[ reflexivity
+| apply le_to_le_to_eq
+ [ apply divides_to_le
+ [ apply lt_O_S
+ | apply divides_gcd_n
+ ]
+ | apply divides_to_le
+ [ apply lt_O_gcd.apply lt_O_S
+ | apply divides_d_gcd
+ [ apply divides_n_n
+ | apply divides_n_n
+ ]
+ ]
+ ]
+]*)
+qed.
+
+theorem gcd_SO_to_lt_O: \forall i,n. (S O) < n \to gcd i n = (S O) \to
+O < i.
+intros.
+elim (le_to_or_lt_eq ? ? (le_O_n i))
+[ assumption
+| absurd ((gcd i n) = (S O))
+ [ assumption
+ | rewrite < H2.
+ simplify.
+ unfold.
+ intro.
+ apply (lt_to_not_eq (S O) n H).
+ auto
+ (*apply sym_eq.
+ assumption*)
+ ]
+]
+qed.
+
+theorem gcd_SO_to_lt_n: \forall i,n. (S O) < n \to i \le n \to gcd i n = (S O) \to
+i < n.
+intros.
+elim (le_to_or_lt_eq ? ? H1)
+ [assumption
+ |absurd ((gcd i n) = (S O))
+ [assumption
+ |rewrite > H3.
+ rewrite > gcd_n_n.
+ unfold.intro.
+ apply (lt_to_not_eq (S O) n H).
+ apply sym_eq.assumption
+ ]
+ ]
+qed.
+
+theorem gcd_n_times_nm: \forall n,m. O < m \to gcd n (n*m) = n.
+intro.apply (nat_case n)
+ [intros.reflexivity
+ |intros.
+ apply le_to_le_to_eq
+ [apply divides_to_le
+ [apply lt_O_S|apply divides_gcd_n]
+ |apply divides_to_le
+ [apply lt_O_gcd.rewrite > (times_n_O O).
+ apply lt_times[apply lt_O_S|assumption]
+ |apply divides_d_gcd
+ [apply (witness ? ? m1).reflexivity
+ |apply divides_n_n
+ ]
+ ]
+ ]
+ ]
+qed.
+
+theorem symmetric_gcd: symmetric nat gcd.
+(*CSC: bug here: unfold symmetric does not work *)
+change with
+(\forall n,m:nat. gcd n m = gcd m n).
+intros.
+auto.
+(*cut (O < (gcd n m) \lor O = (gcd n m))
+[ elim Hcut
+ [ cut (O < (gcd m n) \lor O = (gcd m n))
+ [ elim Hcut1
+ [ apply antisym_le
+ [ apply divides_to_le
+ [ assumption
+ | apply divides_d_gcd
+ [ apply divides_gcd_n
+ | apply divides_gcd_m
+ ]
+ ]
+ | apply divides_to_le
+ [ assumption
+ | apply divides_d_gcd
+ [ apply divides_gcd_n
+ | apply divides_gcd_m
+ ]
+ ]
+ ]
+ | rewrite < H1.
+ cut (m=O \land n=O)
+ [ elim Hcut2.
+ rewrite > H2.
+ rewrite > H3.
+ reflexivity
+ | apply gcd_O_to_eq_O.
+ apply sym_eq.
+ assumption
+ ]
+ ]
+ | apply le_to_or_lt_eq.
+ apply le_O_n
+ ]
+ | rewrite < H.
+ cut (n=O \land m=O)
+ [ elim Hcut1.
+ rewrite > H1.
+ rewrite > H2.
+ reflexivity
+ | apply gcd_O_to_eq_O.apply sym_eq.
+ assumption
+ ]
+ ]
+| apply le_to_or_lt_eq.
+ apply le_O_n
+]*)
+qed.
+
+variant sym_gcd: \forall n,m:nat. gcd n m = gcd m n \def
+symmetric_gcd.
+
+theorem le_gcd_times: \forall m,n,p:nat. O< p \to gcd m n \le gcd m (n*p).
+intros.
+apply (nat_case n)
+[ apply le_n
+| intro.
+ apply divides_to_le
+ [ apply lt_O_gcd.
+ auto
+ (*rewrite > (times_n_O O).
+ apply lt_times
+ [ auto.unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ | assumption
+ ]*)
+ | apply divides_d_gcd;auto
+ (*[ apply (transitive_divides ? (S m1))
+ [ apply divides_gcd_m
+ | apply (witness ? ? p).
+ reflexivity
+ ]
+ | apply divides_gcd_n
+ ]*)
+ ]
+]
+qed.
+
+theorem gcd_times_SO_to_gcd_SO: \forall m,n,p:nat. O < n \to O < p \to
+gcd m (n*p) = (S O) \to gcd m n = (S O).
+intros.
+apply antisymmetric_le
+[ rewrite < H2.
+ auto
+ (*apply le_gcd_times.
+ assumption*)
+| auto
+ (*change with (O < gcd m n).
+ apply lt_O_gcd.
+ assumption*)
+]
+qed.
+
+(* for the "converse" of the previous result see the end of this development *)
+
+theorem eq_gcd_SO_to_not_divides: \forall n,m. (S O) < n \to
+(gcd n m) = (S O) \to \lnot (divides n m).
+intros.unfold.intro.
+elim H2.
+generalize in match H1.
+rewrite > H3.
+intro.
+cut (O < n2)
+[ elim (gcd_times_SO_to_gcd_SO n n n2 ? ? H4)
+ [ cut (gcd n (n*n2) = n);auto
+ (*[ auto.apply (lt_to_not_eq (S O) n)
+ [ assumption
+ | rewrite < H4.
+ assumption
+ ]
+ | apply gcd_n_times_nm.
+ assumption
+ ]*)
+ | auto
+ (*apply (trans_lt ? (S O))
+ [ apply le_n
+ | assumption
+ ]*)
+ | assumption
+ ]
+| elim (le_to_or_lt_eq O n2 (le_O_n n2))
+ [ assumption
+ | apply False_ind.
+ apply (le_to_not_lt n (S O))
+ [ rewrite < H4.
+ apply divides_to_le;auto
+ (*[ rewrite > H4.
+ apply lt_O_S
+ | apply divides_d_gcd
+ [ apply (witness ? ? n2).
+ reflexivity
+ | apply divides_n_n
+ ]
+ ]*)
+ | assumption
+ ]
+ ]
+]
+qed.
+
+theorem gcd_SO_n: \forall n:nat. gcd (S O) n = (S O).
+intro.
+auto.
+(*apply antisym_le
+[ apply divides_to_le
+ [ unfold lt.
+ apply le_n
+ | apply divides_gcd_n
+ ]
+| cut (O < gcd (S O) n \lor O = gcd (S O) n)
+ [ elim Hcut
+ [ assumption
+ | apply False_ind.
+ apply (not_eq_O_S O).
+ cut ((S O)=O \land n=O)
+ [ elim Hcut1.
+ apply sym_eq.
+ assumption
+ | apply gcd_O_to_eq_O.
+ apply sym_eq.
+ assumption
+ ]
+ ]
+ | apply le_to_or_lt_eq.
+ apply le_O_n
+ ]
+]*)
+qed.
+
+theorem divides_gcd_mod: \forall m,n:nat. O < n \to
+divides (gcd m n) (gcd n (m \mod n)).
+intros.
+auto.
+(*apply divides_d_gcd
+[ apply divides_mod
+ [ assumption
+ | apply divides_gcd_n
+ | apply divides_gcd_m
+ ]
+| apply divides_gcd_m.
+]*)
+qed.
+
+theorem divides_mod_gcd: \forall m,n:nat. O < n \to
+divides (gcd n (m \mod n)) (gcd m n) .
+intros.
+auto.
+(*apply divides_d_gcd
+[ apply divides_gcd_n
+| apply (divides_mod_to_divides ? ? n)
+ [ assumption
+ | apply divides_gcd_m
+ | apply divides_gcd_n
+ ]
+]*)
+qed.
+
+theorem gcd_mod: \forall m,n:nat. O < n \to
+(gcd n (m \mod n)) = (gcd m n) .
+intros.
+auto.
+(*apply antisymmetric_divides
+[ apply divides_mod_gcd.
+ assumption
+| apply divides_gcd_mod.
+ assumption
+]*)
+qed.
+
+(* gcd and primes *)
+
+theorem prime_to_gcd_SO: \forall n,m:nat. prime n \to n \ndivides m \to
+gcd n m = (S O).
+intros.unfold prime in H.
+elim H.
+apply antisym_le
+[ apply not_lt_to_le.unfold Not.unfold lt.
+ intro.
+ apply H1.
+ rewrite < (H3 (gcd n m));auto
+ (*[ apply divides_gcd_m
+ | apply divides_gcd_n
+ | assumption
+ ]*)
+| cut (O < gcd n m \lor O = gcd n m)
+ [ elim Hcut
+ [ assumption
+ | apply False_ind.
+ apply (not_le_Sn_O (S O)).
+ cut (n=O \land m=O)
+ [ elim Hcut1.
+ rewrite < H5 in \vdash (? ? %).
+ assumption
+ | auto
+ (*apply gcd_O_to_eq_O.
+ apply sym_eq.
+ assumption*)
+ ]
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ apply le_O_n*)
+ ]
+]
+qed.
+
+theorem divides_times_to_divides: \forall n,p,q:nat.prime n \to n \divides p*q \to
+n \divides p \lor n \divides q.
+intros.
+cut (n \divides p \lor n \ndivides p)
+[elim Hcut
+ [left.assumption
+ |right.
+ cut (\exists a,b. a*n - b*p = (S O) \lor b*p - a*n = (S O))
+ [elim Hcut1.elim H3.elim H4
+ [(* first case *)
+ rewrite > (times_n_SO q).rewrite < H5.
+ rewrite > distr_times_minus.
+ rewrite > (sym_times q (a1*p)).
+ rewrite > (assoc_times a1).
+ elim H1.
+ (*
+ rewrite > H6.
+ applyS (witness n (n*(q*a-a1*n2)) (q*a-a1*n2))
+ reflexivity. *)
+ applyS (witness n ? ? (refl_eq ? ?)) (* timeout=50 *)
+ (*
+ rewrite < (sym_times n).rewrite < assoc_times.
+ rewrite > (sym_times q).rewrite > assoc_times.
+ rewrite < (assoc_times a1).rewrite < (sym_times n).
+ rewrite > (assoc_times n).
+ rewrite < distr_times_minus.
+ apply (witness ? ? (q*a-a1*n2)).reflexivity
+ *)
+ |(* second case *)
+ rewrite > (times_n_SO q).rewrite < H5.
+ rewrite > distr_times_minus.
+ rewrite > (sym_times q (a1*p)).
+ rewrite > (assoc_times a1).
+ elim H1.
+ auto
+ (*rewrite > H6.
+ rewrite < sym_times.rewrite > assoc_times.
+ rewrite < (assoc_times q).
+ rewrite < (sym_times n).
+ rewrite < distr_times_minus.
+ apply (witness ? ? (n2*a1-q*a)).
+ reflexivity*)
+ ](* end second case *)
+ | rewrite < (prime_to_gcd_SO n p);auto
+ (* [ apply eq_minus_gcd
+ | assumption
+ | assumption
+ ]*)
+ ]
+ ]
+ | apply (decidable_divides n p).
+ apply (trans_lt ? (S O))
+ [ auto
+ (*unfold lt.
+ apply le_n*)
+ | unfold prime in H.
+ elim H.
+ assumption
+ ]
+ ]
+qed.
+
+theorem eq_gcd_times_SO: \forall m,n,p:nat. O < n \to O < p \to
+gcd m n = (S O) \to gcd m p = (S O) \to gcd m (n*p) = (S O).
+intros.
+apply antisymmetric_le
+[ apply not_lt_to_le.
+ unfold Not.intro.
+ cut (divides (smallest_factor (gcd m (n*p))) n \lor
+ divides (smallest_factor (gcd m (n*p))) p)
+ [ elim Hcut
+ [ apply (not_le_Sn_n (S O)).
+ change with ((S O) < (S O)).
+ rewrite < H2 in \vdash (? ? %).
+ apply (lt_to_le_to_lt ? (smallest_factor (gcd m (n*p))))
+ [ auto
+ (*apply lt_SO_smallest_factor.
+ assumption*)
+ | apply divides_to_le;auto
+ (*[ rewrite > H2.
+ unfold lt.
+ apply le_n
+ | apply divides_d_gcd
+ [ assumption
+ | apply (transitive_divides ? (gcd m (n*p)))
+ [ apply divides_smallest_factor_n.
+ apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]
+ | apply divides_gcd_n
+ ]
+ ]
+ ]*)
+ ]
+ | apply (not_le_Sn_n (S O)).
+ change with ((S O) < (S O)).
+ rewrite < H3 in \vdash (? ? %).
+ apply (lt_to_le_to_lt ? (smallest_factor (gcd m (n*p))))
+ [ apply lt_SO_smallest_factor.
+ assumption
+ | apply divides_to_le;auto
+ (*[ rewrite > H3.
+ unfold lt.
+ apply le_n
+ | apply divides_d_gcd
+ [ assumption
+ | apply (transitive_divides ? (gcd m (n*p)))
+ [ apply divides_smallest_factor_n.
+ apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]
+ | apply divides_gcd_n
+ ]
+ ]
+ ]*)
+ ]
+ ]
+ | apply divides_times_to_divides;auto
+ (*[ apply prime_smallest_factor_n.
+ assumption
+ | auto.apply (transitive_divides ? (gcd m (n*p)))
+ [ apply divides_smallest_factor_n.
+ apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]
+ | apply divides_gcd_m
+ ]
+ ]*)
+ ]
+| auto
+ (*change with (O < gcd m (n*p)).
+ apply lt_O_gcd.
+ rewrite > (times_n_O O).
+ apply lt_times;assumption.*)
+]
+qed.
+
+theorem gcd_SO_to_divides_times_to_divides: \forall m,n,p:nat. O < n \to
+gcd n m = (S O) \to n \divides (m*p) \to n \divides p.
+intros.
+cut (n \divides p \lor n \ndivides p)
+[ elim Hcut
+ [ assumption
+ | cut (\exists a,b. a*n - b*m = (S O) \lor b*m - a*n = (S O))
+ [ elim Hcut1.elim H4.elim H5
+ [(* first case *)
+ rewrite > (times_n_SO p).rewrite < H6.
+ rewrite > distr_times_minus.
+ rewrite > (sym_times p (a1*m)).
+ rewrite > (assoc_times a1).
+ elim H2.
+ applyS (witness n ? ? (refl_eq ? ?)) (* timeout=50 *).
+ |(* second case *)
+ rewrite > (times_n_SO p).rewrite < H6.
+ rewrite > distr_times_minus.
+ rewrite > (sym_times p (a1*m)).
+ rewrite > (assoc_times a1).
+ elim H2.
+ applyS (witness n ? ? (refl_eq ? ?)).
+ ](* end second case *)
+ | rewrite < H1.
+ apply eq_minus_gcd
+ ]
+ ]
+| auto
+ (*apply (decidable_divides n p).
+ assumption.*)
+]
+qed.
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/le_arith".
+
+include "nat/times.ma".
+include "nat/orders.ma".
+
+(* plus *)
+theorem monotonic_le_plus_r:
+\forall n:nat.monotonic nat le (\lambda m.n + m).
+simplify.intros.
+elim n;simplify
+[ assumption
+| auto
+ (*apply le_S_S.assumption*)
+]
+qed.
+
+theorem le_plus_r: \forall p,n,m:nat. n \le m \to p + n \le p + m
+\def monotonic_le_plus_r.
+
+theorem monotonic_le_plus_l:
+\forall m:nat.monotonic nat le (\lambda n.n + m).
+simplify.intros.
+ (*rewrite < sym_plus.
+ rewrite < (sym_plus m).*)
+ applyS le_plus_r.
+ assumption.
+qed.
+
+(* IN ALTERNATIVA:
+
+theorem monotonic_le_plus_l:
+\forall m:nat.monotonic nat le (\lambda n.n + m).
+simplify.intros.
+ rewrite < sym_plus.
+ rewrite < (sym_plus m).
+ auto.
+qed.
+*)
+theorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p
+\def monotonic_le_plus_l.
+
+theorem le_plus: \forall n1,n2,m1,m2:nat. n1 \le n2 \to m1 \le m2
+\to n1 + m1 \le n2 + m2.
+intros.
+auto.
+(*apply (trans_le ? (n2 + m1)).
+apply le_plus_l.assumption.
+apply le_plus_r.assumption.*)
+qed.
+
+theorem le_plus_n :\forall n,m:nat. m \le n + m.
+intros.
+change with (O+m \le n+m).
+auto.
+(*apply le_plus_l.
+ apply le_O_n.*)
+qed.
+
+theorem eq_plus_to_le: \forall n,m,p:nat.n=m+p \to m \le n.
+intros.
+rewrite > H.
+rewrite < sym_plus.
+apply le_plus_n. (* a questo punto funziona anche: auto.*)
+qed.
+
+(* times *)
+theorem monotonic_le_times_r:
+\forall n:nat.monotonic nat le (\lambda m. n * m).
+simplify.intros.elim n;simplify
+[ apply le_O_n.
+| auto.
+(*apply le_plus;
+ assumption. *) (* chiudo entrambi i goal attivi in questo modo*)
+]
+qed.
+
+theorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m
+\def monotonic_le_times_r.
+
+theorem monotonic_le_times_l:
+\forall m:nat.monotonic nat le (\lambda n.n*m).
+simplify.intros.
+(*rewrite < sym_times.
+ rewrite < (sym_times m).
+*)
+applyS le_times_r.
+assumption.
+qed.
+
+(* IN ALTERNATIVA:
+theorem monotonic_le_times_l:
+\forall m:nat.monotonic nat le (\lambda n.n*m).
+simplify.intros.
+rewrite < sym_times.
+rewrite < (sym_times m).
+auto.
+qed.
+*)
+
+theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p
+\def monotonic_le_times_l.
+
+theorem le_times: \forall n1,n2,m1,m2:nat. n1 \le n2 \to m1 \le m2
+\to n1*m1 \le n2*m2.
+intros.
+auto.
+(*apply (trans_le ? (n2*m1)).
+apply le_times_l.assumption.
+apply le_times_r.assumption.*)
+qed.
+
+theorem le_times_n: \forall n,m:nat.(S O) \le n \to m \le n*m.
+intros.elim H;simplify
+[ auto
+ (*elim (plus_n_O ?).
+ apply le_n....*)
+| auto
+ (*rewrite < sym_plus.
+ apply le_plus_n.*)
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/lt_arith".
+
+include "nat/div_and_mod.ma".
+
+(* plus *)
+theorem monotonic_lt_plus_r:
+\forall n:nat.monotonic nat lt (\lambda m.n+m).
+simplify.intros.
+elim n;simplify
+[ assumption
+| auto.
+ (*unfold lt.
+ apply le_S_S.
+ assumption.*)
+]
+qed.
+
+variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def
+monotonic_lt_plus_r.
+
+theorem monotonic_lt_plus_l:
+\forall n:nat.monotonic nat lt (\lambda m.m+n).
+simplify.
+intros.
+(*rewrite < sym_plus.
+ rewrite < (sym_plus n).*)
+applyS lt_plus_r.assumption.
+qed.
+(* IN ALTERNATIVA: mantengo le 2 rewrite, e concludo con auto. *)
+
+variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def
+monotonic_lt_plus_l.
+
+theorem lt_plus: \forall n,m,p,q:nat. n < m \to p < q \to n + p < m + q.
+intros.
+auto.
+(*apply (trans_lt ? (n + q)).
+apply lt_plus_r.assumption.
+apply lt_plus_l.assumption.*)
+qed.
+
+theorem lt_plus_to_lt_l :\forall n,p,q:nat. p+n < q+n \to p<q.
+intro.elim n
+[ rewrite > plus_n_O.
+ rewrite > (plus_n_O q).
+ assumption.
+| apply H.
+ unfold lt.
+ apply le_S_S_to_le.
+ rewrite > plus_n_Sm.
+ rewrite > (plus_n_Sm q).
+ exact H1.
+]
+qed.
+
+theorem lt_plus_to_lt_r :\forall n,p,q:nat. n+p < n+q \to p<q.
+intros.
+apply (lt_plus_to_lt_l n).
+rewrite > sym_plus.
+rewrite > (sym_plus q).
+assumption.
+qed.
+
+(* times and zero *)
+theorem lt_O_times_S_S: \forall n,m:nat.O < (S n)*(S m).
+intros.
+auto.
+(*simplify.
+unfold lt.
+apply le_S_S.
+apply le_O_n.*)
+qed.
+
+(* times *)
+theorem monotonic_lt_times_r:
+\forall n:nat.monotonic nat lt (\lambda m.(S n)*m).
+simplify.
+intros.elim n
+[ auto
+ (*simplify.
+ rewrite < plus_n_O.
+ rewrite < plus_n_O.
+ assumption.*)
+| apply lt_plus;assumption (* chiudo con assumption entrambi i sottogoal attivi*)
+]
+qed.
+
+theorem lt_times_r: \forall n,p,q:nat. p < q \to (S n) * p < (S n) * q
+\def monotonic_lt_times_r.
+
+theorem monotonic_lt_times_l:
+\forall m:nat.monotonic nat lt (\lambda n.n * (S m)).
+simplify.
+intros.
+(*rewrite < sym_times.
+ rewrite < (sym_times (S m)).*)
+applyS lt_times_r.assumption.
+qed.
+
+(* IN ALTERNATIVA: mantengo le 2 rewrite, e concludo con auto. *)
+
+
+variant lt_times_l: \forall n,p,q:nat. p<q \to p*(S n) < q*(S n)
+\def monotonic_lt_times_l.
+
+theorem lt_times:\forall n,m,p,q:nat. n<m \to p<q \to n*p < m*q.
+intro.
+elim n
+[ apply (lt_O_n_elim m H).
+ intro.
+ cut (lt O q)
+ [ apply (lt_O_n_elim q Hcut).
+ intro.
+ auto
+ (*change with (O < (S m1)*(S m2)).
+ apply lt_O_times_S_S.*)
+ | apply (ltn_to_ltO p q H1). (*funziona anche auto*)
+ ]
+| apply (trans_lt ? ((S n1)*q))
+ [ auto
+ (*apply lt_times_r.
+ assumption.*)
+ | cut (lt O q)
+ [ apply (lt_O_n_elim q Hcut).
+ intro.
+ auto
+ (*apply lt_times_l.
+ assumption.*)
+ |apply (ltn_to_ltO p q H2). (*funziona anche auto*)
+ ]
+ ]
+]
+qed.
+
+theorem lt_times_to_lt_l:
+\forall n,p,q:nat. p*(S n) < q*(S n) \to p < q.
+intros.
+cut (p < q \lor p \nlt q)
+[ elim Hcut
+ [ assumption
+ | absurd (p * (S n) < q * (S n))
+ [ assumption
+ | apply le_to_not_lt.
+ auto
+ (*apply le_times_l.
+ apply not_lt_to_le.
+ assumption.*)
+ ]
+ ]
+|exact (decidable_lt p q).
+]
+qed.
+
+theorem lt_times_to_lt_r:
+\forall n,p,q:nat. (S n)*p < (S n)*q \to lt p q.
+intros.
+apply (lt_times_to_lt_l n).
+rewrite < sym_times.
+rewrite < (sym_times (S n)).
+assumption.
+qed.
+
+theorem nat_compare_times_l : \forall n,p,q:nat.
+nat_compare p q = nat_compare ((S n) * p) ((S n) * q).
+intros.
+apply nat_compare_elim
+[ intro.
+ apply nat_compare_elim
+ [ intro.
+ reflexivity
+ | intro.
+ absurd (p=q)
+ [ auto.
+ (*apply (inj_times_r n).
+ assumption*)
+ | auto
+ (*apply lt_to_not_eq.
+ assumption*)
+ ]
+ | intro.
+ absurd (q<p)
+ [ auto.
+ (*apply (lt_times_to_lt_r n).
+ assumption.*)
+ | auto
+ (*apply le_to_not_lt.
+ apply lt_to_le.
+ assumption*)
+ ]
+ ].
+| intro.
+ auto
+ (*rewrite < H.
+ rewrite > nat_compare_n_n.
+ reflexivity*)
+| intro.
+ apply nat_compare_elim
+ [ intro.
+ absurd (p<q)
+ [ auto
+ (*apply (lt_times_to_lt_r n).
+ assumption.*)
+ | auto
+ (*apply le_to_not_lt.
+ apply lt_to_le.
+ assumption.*)
+ ]
+ | intro.
+ absurd (q=p)
+ [ auto.
+ (*symmetry.
+ apply (inj_times_r n).
+ assumption*)
+ | auto
+ (*apply lt_to_not_eq.
+ assumption*)
+ ]
+ | intro.
+ reflexivity
+ ]
+].
+qed.
+
+(* div *)
+
+theorem eq_mod_O_to_lt_O_div: \forall n,m:nat. O < m \to O < n\to n \mod m = O \to O < n / m.
+intros 4.
+apply (lt_O_n_elim m H).
+intros.
+apply (lt_times_to_lt_r m1).
+rewrite < times_n_O.
+rewrite > (plus_n_O ((S m1)*(n / (S m1)))).
+rewrite < H2.
+rewrite < sym_times.
+rewrite < div_mod
+[ rewrite > H2.
+ assumption.
+| auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n.*)
+]
+qed.
+
+theorem lt_div_n_m_n: \forall n,m:nat. (S O) < m \to O < n \to n / m \lt n.
+intros.
+apply (nat_case1 (n / m))
+[ intro.
+ assumption
+| intros.
+ rewrite < H2.
+ rewrite > (div_mod n m) in \vdash (? ? %)
+ [ apply (lt_to_le_to_lt ? ((n / m)*m))
+ [ apply (lt_to_le_to_lt ? ((n / m)*(S (S O))))
+ [ rewrite < sym_times.
+ rewrite > H2.
+ simplify.
+ unfold lt.
+ auto.
+ (*rewrite < plus_n_O.
+ rewrite < plus_n_Sm.
+ apply le_S_S.
+ apply le_S_S.
+ apply le_plus_n*)
+ | auto
+ (*apply le_times_r.
+ assumption*)
+ ]
+ | auto
+ (*rewrite < sym_plus.
+ apply le_plus_n*)
+ ]
+ | auto
+ (*apply (trans_lt ? (S O)).
+ unfold lt.
+ apply le_n.
+ assumption*)
+ ]
+]
+qed.
+
+(* general properties of functions *)
+theorem monotonic_to_injective: \forall f:nat\to nat.
+monotonic nat lt f \to injective nat nat f.
+unfold injective.intros.
+apply (nat_compare_elim x y)
+[ intro.
+ apply False_ind.
+ apply (not_le_Sn_n (f x)).
+ rewrite > H1 in \vdash (? ? %).
+ change with (f x < f y).
+ auto
+ (*apply H.
+ apply H2*)
+| intros.
+ assumption
+| intro.
+ apply False_ind.
+ apply (not_le_Sn_n (f y)).
+ rewrite < H1 in \vdash (? ? %).
+ change with (f y < f x).
+ auto
+ (*apply H.
+ apply H2*)
+]
+qed.
+
+theorem increasing_to_injective: \forall f:nat\to nat.
+increasing f \to injective nat nat f.
+intros.
+auto.
+(*apply monotonic_to_injective.
+apply increasing_to_monotonic.
+assumption.*)
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/map_iter_p.ma".
+
+include "nat/permutation.ma".
+include "nat/count.ma".
+
+let rec map_iter_p n p (g:nat \to nat) (a:nat) f \def
+ match n with
+ [ O \Rightarrow a
+ | (S k) \Rightarrow
+ match p (S k) with
+ [true \Rightarrow f (g (S k)) (map_iter_p k p g a f)
+ |false \Rightarrow map_iter_p k p g a f]
+ ].
+
+theorem eq_map_iter_p: \forall g1,g2:nat \to nat.
+\forall p:nat \to bool.
+\forall f:nat \to nat \to nat. \forall a,n:nat.
+(\forall m:nat. m \le n \to g1 m = g2 m) \to
+map_iter_p n p g1 a f = map_iter_p n p g2 a f.
+intros 6.
+elim n
+[ auto
+ (*simplify.
+ reflexivity*)
+| simplify.
+ elim (p (S n1))
+ [ simplify.
+ apply eq_f2
+ [ auto
+ (*apply H1.
+ apply le_n*)
+ | simplify.
+ apply H.
+ intros.
+ auto
+ (*apply H1.
+ apply le_S.
+ assumption*)
+ ]
+ | simplify.
+ apply H.
+ intros.
+ auto
+ (*apply H1.
+ apply le_S.
+ assumption*)
+ ]
+]
+qed.
+
+(* useful since simplify simpifies too much *)
+
+theorem map_iter_p_O: \forall p.\forall g.\forall f. \forall a:nat.
+map_iter_p O p g a f = a.
+intros.
+auto.
+(*simplify.reflexivity.*)
+qed.
+
+theorem map_iter_p_S_true: \forall p.\forall g.\forall f. \forall a,n:nat.
+p (S n) = true \to
+map_iter_p (S n) p g a f = f (g (S n)) (map_iter_p n p g a f).
+intros.simplify.rewrite > H.reflexivity.
+qed.
+
+theorem map_iter_p_S_false: \forall p.\forall g.\forall f. \forall a,n:nat.
+p (S n) = false \to
+map_iter_p (S n) p g a f = map_iter_p n p g a f.
+intros.simplify.rewrite > H.reflexivity.
+qed.
+
+(* map_iter examples *)
+definition pi_p \def \lambda p. \lambda n.
+map_iter_p n p (\lambda n.n) (S O) times.
+
+lemma pi_p_S: \forall n.\forall p.
+pi_p p (S n) =
+ match p (S n) with
+ [true \Rightarrow (S n)*(pi_p p n)
+ |false \Rightarrow (pi_p p n)
+ ].
+intros.reflexivity.
+qed.
+
+lemma lt_O_pi_p: \forall n.\forall p.
+O < pi_p p n.
+intros.
+elim n
+[ auto
+ (*simplify.
+ apply le_n*)
+| rewrite > pi_p_S.
+ elim p (S n1)
+ [ change with (O < (S n1)*(pi_p p n1)).
+ auto
+ (*rewrite >(times_n_O n1).
+ apply lt_times
+ [ apply le_n
+ | assumption
+ ]*)
+ | assumption
+ ]
+]
+qed.
+
+let rec card n p \def
+ match n with
+ [O \Rightarrow O
+ |(S m) \Rightarrow
+ (bool_to_nat (p (S m))) + (card m p)].
+
+lemma count_card: \forall p.\forall n.
+p O = false \to count (S n) p = card n p.
+intros.
+elim n
+[ simplify.
+ auto
+ (*rewrite > H.
+ reflexivity*)
+| simplify.
+ rewrite < plus_n_O.
+ apply eq_f.
+ (*qui auto non chiude un goal chiuso invece dal solo assumption*)
+ assumption
+]
+qed.
+
+lemma count_card1: \forall p.\forall n.
+p O = false \to p n = false \to count n p = card n p.
+intros 3.
+apply (nat_case n)
+[ intro.
+ simplify.
+ auto
+ (*rewrite > H.
+ reflexivity*)
+| intros.rewrite > (count_card ? ? H).
+ simplify.
+ auto
+ (*rewrite > H1.
+ reflexivity*)
+]
+qed.
+
+lemma a_times_pi_p: \forall p. \forall a,n.
+exp a (card n p) * pi_p p n = map_iter_p n p (\lambda n.a*n) (S O) times.
+intros.
+elim n
+[ auto
+ (*simplify.
+ reflexivity*)
+| simplify.
+ apply (bool_elim ? (p (S n1)))
+ [ intro.
+ change with
+ (a*exp a (card n1 p) * ((S n1) * (pi_p p n1)) =
+ a*(S n1)*map_iter_p n1 p (\lambda n.a*n) (S O) times).
+ rewrite < H.
+ auto
+ | intro.
+ (*la chiamata di auto in questo punto dopo circa 8 minuti non aveva
+ * ancora generato un risultato
+ *)
+ assumption
+ ]
+]
+qed.
+
+definition permut_p \def \lambda f. \lambda p:nat\to bool. \lambda n.
+\forall i. i \le n \to p i = true \to ((f i \le n \land p (f i) = true)
+\land (\forall j. p j = true \to j \le n \to i \neq j \to (f i \neq f j))).
+
+definition extentional_eq_n \def \lambda f,g:nat \to nat.\lambda n.
+\forall x. x \le n \to f x = g x.
+
+lemma extentional_eq_n_to_permut_p: \forall f,g. \forall p. \forall n.
+extentional_eq_n f g n\to permut_p f p n \to permut_p g p n.
+intros.
+unfold permut_p.
+intros.
+elim (H1 i H2 H3).
+split
+[ elim H4.
+ split
+ [ rewrite < (H i H2).
+ assumption
+ | rewrite < (H i H2).
+ assumption
+ ]
+| intros.
+ unfold.
+ intro.
+ apply (H5 j H6 H7 H8).
+ rewrite > (H i H2).
+ rewrite > (H j H7).
+ assumption
+]
+qed.
+
+theorem permut_p_compose: \forall f,g.\forall p.\forall n.
+permut_p f p n \to permut_p g p n \to permut_p (compose ? ? ? g f) p n.
+intros.
+unfold permut_p.
+intros.
+elim (H i H2 H3).
+elim H4.
+elim (H1 (f i) H6 H7).
+elim H8.
+split
+[ split
+ [ unfold compose.
+ assumption
+ | unfold compose.
+ auto
+ (*rewrite < H11.
+ reflexivity*)
+ ]
+| intros.
+ unfold compose.
+ apply (H9 (f j))
+ [ elim (H j H13 H12).
+ auto
+ (*elim H15.
+ rewrite < H18.
+ reflexivity*)
+ | elim (H j H13 H12).
+ elim H15.
+ assumption.
+ | apply (H5 j H12 H13 H14)
+ ]
+]
+qed.
+
+
+theorem permut_p_S_to_permut_p: \forall f.\forall p.\forall n.
+permut_p f p (S n) \to (f (S n)) = (S n) \to permut_p f p n.
+intros.
+unfold permut_p.
+intros.
+split
+[ elim (H i (le_S i n H2) H3).
+ split
+ [ elim H4.
+ elim (le_to_or_lt_eq (f i) (S n))
+ [ auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ | absurd (f i = (S n))
+ [ assumption
+ | rewrite < H1.
+ apply H5
+ [ auto
+ (*rewrite < H8.
+ assumption*)
+ | apply le_n
+ | unfold.intro.rewrite > H8 in H2.
+ apply (not_le_Sn_n n).rewrite < H9.
+ assumption
+ ]
+ ]
+ | assumption
+ ]
+ | auto
+ (*elim H4.
+ assumption*)
+ ]
+| intros.
+ elim (H i (le_S i n H2) H3).
+ auto
+ (*apply H8
+ [ assumption
+ | apply le_S.
+ assumption
+ | assumption
+ ]*)
+]
+qed.
+
+lemma permut_p_transpose: \forall p.\forall i,j,n.
+i \le n \to j \le n \to p i = p j \to
+permut_p (transpose i j) p n.
+unfold permut_p.
+intros.
+split
+[ split
+ [ unfold transpose.
+ apply (eqb_elim i1 i)
+ [ intro.
+ apply (eqb_elim i1 j)
+ [ simplify.intro.assumption
+ | simplify.intro.assumption
+ ]
+ | intro.
+ apply (eqb_elim i1 j)
+ [ simplify.intro.assumption
+ | simplify.intro.assumption
+ ]
+ ]
+ | unfold transpose.
+ apply (eqb_elim i1 i)
+ [ intro.
+ apply (eqb_elim i1 j)
+ [ simplify.intro.
+ auto
+ (*rewrite < H6.
+ assumption*)
+ | simplify.intro.
+ auto
+ (*rewrite < H2.
+ rewrite < H5.
+ assumption*)
+ ]
+ | intro.
+ apply (eqb_elim i1 j)
+ [ simplify.intro.
+ auto
+ (*rewrite > H2.
+ rewrite < H6.
+ assumption*)
+ | simplify.intro.
+ assumption
+ ]
+ ]
+ ]
+| intros.
+ unfold Not.
+ intro.
+ auto
+ (*apply H7.
+ apply (injective_transpose ? ? ? ? H8)*)
+]
+qed.
+
+
+theorem eq_map_iter_p_k: \forall f,g.\forall p.\forall a,k,n:nat.
+p (S n-k) = true \to (\forall i. (S n)-k < i \to i \le (S n) \to (p i) = false) \to
+map_iter_p (S n) p g a f = map_iter_p (S n-k) p g a f.
+intros 5.
+elim k 3
+[ auto
+ (*rewrite < minus_n_O.
+ reflexivity*)
+| apply (nat_case n1)
+ [ intros.
+ rewrite > map_iter_p_S_false
+ [ reflexivity
+ | auto
+ (*apply H2
+ [ simplify.
+ apply lt_O_S.
+ | apply le_n
+ ]*)
+ ]
+ | intros.
+ rewrite > map_iter_p_S_false
+ [ rewrite > (H m H1)
+ [ reflexivity
+ | intros.
+ apply (H2 i H3).
+ auto
+ (*apply le_S.
+ assumption*)
+ ]
+ | auto
+ (*apply H2
+ [ auto
+ | apply le_n
+ ]*)
+ ]
+ ]
+]
+qed.
+
+
+
+theorem eq_map_iter_p_a: \forall p.\forall f.\forall g. \forall a,n:nat.
+(\forall i.i \le n \to p i = false) \to map_iter_p n p g a f = a.
+intros 5.
+elim n
+[ auto
+ (*simplify.
+ reflexivity*)
+| rewrite > map_iter_p_S_false
+ [ apply H.
+ intros.
+ auto
+ (*apply H1.
+ apply le_S.
+ assumption*)
+ | auto
+ (*apply H1.
+ apply le_n*)
+ ]
+]
+qed.
+
+theorem eq_map_iter_p_transpose: \forall p.\forall f.associative nat f \to
+symmetric2 nat nat f \to \forall g. \forall a,k,n:nat. k < n \to
+(p (S n) = true) \to (p (n-k)) = true \to (\forall i. n-k < i \to i \le n \to (p i) = false)
+\to map_iter_p (S n) p g a f = map_iter_p (S n) p (\lambda m. g (transpose (n-k) (S n) m)) a f.
+intros 8.
+apply (nat_case n)
+[ intro.
+ absurd (k < O)
+ [ assumption
+ | auto
+ (*apply le_to_not_lt.
+ apply le_O_n*)
+ ]
+| intros.
+ rewrite > (map_iter_p_S_true ? ? ? ? ? H3).
+ rewrite > (map_iter_p_S_true ? ? ? ? ? H3).
+ rewrite > (eq_map_iter_p_k ? ? ? ? ? ? H4 H5).
+ rewrite > (eq_map_iter_p_k ? ? ? ? ? ? H4 H5).
+ generalize in match H4.
+ rewrite > minus_Sn_m
+ [ intro.
+ rewrite > (map_iter_p_S_true ? ? ? ? ? H6).
+ rewrite > (map_iter_p_S_true ? ? ? ? ? H6).
+ rewrite > transpose_i_j_j.
+ rewrite > transpose_i_j_i.
+ cut (map_iter_p (m-k) p g a f =
+ map_iter_p (m-k) p (\lambda x.g (transpose (S(m-k)) (S(S m)) x)) a f)
+ [ rewrite < Hcut.
+ rewrite < H.
+ rewrite < H1 in \vdash (? ? (? % ?) ?).
+ auto
+ (*rewrite > H.
+ reflexivity*)
+ | apply eq_map_iter_p.
+ intros.
+ unfold transpose.
+ cut (eqb m1 (S (m-k)) =false)
+ [ cut (eqb m1 (S (S m)) =false)
+ [ rewrite > Hcut.
+ rewrite > Hcut1.
+ reflexivity
+ | apply not_eq_to_eqb_false.
+ apply lt_to_not_eq.
+ apply (le_to_lt_to_lt ? m)
+ [ auto
+ (*apply (trans_le ? (m-k))
+ [ assumption
+ | auto
+ ]*)
+ | auto
+ (*apply le_S.
+ apply le_n*)
+ ]
+ ]
+ | apply not_eq_to_eqb_false.
+ apply lt_to_not_eq.
+ auto
+ (*unfold.
+ auto*)
+ ]
+ ]
+ | auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ ]
+]
+qed.
+
+theorem eq_map_iter_p_transpose1: \forall p.\forall f.associative nat f \to
+symmetric2 nat nat f \to \forall g. \forall a,k1,k2,n:nat. O < k1 \to k1 < k2 \to k2 \le n \to
+(p k1) = true \to (p k2) = true \to (\forall i. k1 < i \to i < k2 \to (p i) = false)
+\to map_iter_p n p g a f = map_iter_p n p (\lambda m. g (transpose k1 k2 m)) a f.
+intros 10.
+elim n 2
+[ absurd (k2 \le O)
+ [ assumption
+ | auto
+ (*apply lt_to_not_le.
+ apply (trans_lt ? k1 ? H2 H3)*)
+ ]
+| apply (eqb_elim (S n1) k2)
+ [ intro.
+ rewrite < H4.
+ intros.
+ cut (k1 = n1 - (n1 -k1))
+ [ rewrite > Hcut.
+ apply (eq_map_iter_p_transpose p f H H1 g a (n1-k1))
+ [ cut (k1 \le n1);auto
+ | assumption
+ | auto
+ (*rewrite < Hcut.
+ assumption*)
+ | rewrite < Hcut.
+ intros.
+ apply (H9 i H10).
+ auto
+ (*unfold.
+ auto*)
+ ]
+ | apply sym_eq.
+ auto
+ (*apply plus_to_minus.
+ auto*)
+ ]
+ | intros.
+ cut ((S n1) \neq k1)
+ [ apply (bool_elim ? (p (S n1)))
+ [ intro.
+ rewrite > map_iter_p_S_true
+ [ rewrite > map_iter_p_S_true
+ [ cut ((transpose k1 k2 (S n1)) = (S n1))
+ [ rewrite > Hcut1.
+ apply eq_f.
+ apply (H3 H5)
+ [ elim (le_to_or_lt_eq ? ? H6)
+ [ auto
+ | absurd (S n1=k2)
+ [ auto
+ (*apply sym_eq.
+ assumption*)
+ | assumption
+ ]
+ ]
+ | assumption
+ | assumption
+ | assumption
+ ]
+ | unfold transpose.
+ rewrite > (not_eq_to_eqb_false ? ? Hcut).
+ rewrite > (not_eq_to_eqb_false ? ? H4).
+ reflexivity
+ ]
+ | assumption
+ ]
+ | assumption
+ ]
+ | intro.
+ rewrite > map_iter_p_S_false
+ [ rewrite > map_iter_p_S_false
+ [ apply (H3 H5)
+ [ elim (le_to_or_lt_eq ? ? H6)
+ [ auto
+ | (*l'invocazione di auto qui genera segmentation fault*)
+ absurd (S n1=k2)
+ [ auto
+ (*apply sym_eq.
+ assumption*)
+ | assumption
+ ]
+ ]
+ | assumption
+ | assumption
+ | assumption
+ ]
+ | assumption
+ ]
+ | assumption
+ ]
+ ]
+ | unfold.
+ intro.
+ absurd (k1 < k2)
+ [ assumption
+ | apply le_to_not_lt.
+ rewrite < H10.
+ assumption
+ ]
+ ]
+ ]
+]
+qed.
+
+lemma decidable_n:\forall p.\forall n.
+(\forall m. m \le n \to (p m) = false) \lor
+(\exists m. m \le n \land (p m) = true \land
+\forall i. m < i \to i \le n \to (p i) = false).
+intros.
+elim n
+[ apply (bool_elim ? (p O))
+ [ intro.
+ right.
+ apply (ex_intro ? ? O).
+ split
+ [ auto
+ (*split
+ [ apply le_n
+ | assumption
+ ]*)
+ | intros.
+ absurd (O<i)
+ [ assumption
+ | auto
+ (*apply le_to_not_lt.
+ assumption*)
+ ]
+ ]
+ | intro.
+ left.
+ intros.
+ apply (le_n_O_elim m H1).
+ assumption
+ ]
+| apply (bool_elim ? (p (S n1)))
+ [ intro.
+ right.
+ apply (ex_intro ? ? (S n1)).
+ split
+ [ split
+ [ apply le_n
+ | assumption
+ ]
+ | intros.
+ absurd (S n1<i)
+ [ assumption
+ | auto
+ (*apply le_to_not_lt.
+ assumption*)
+ ]
+ ]
+ | elim H
+ [ left.
+ intros.
+ elim (le_to_or_lt_eq m (S n1) H3)
+ [ auto
+ (*apply H1.
+ apply le_S_S_to_le.
+ assumption*)
+ | auto
+ (*rewrite > H4.
+ assumption*)
+ ]
+ | right.
+ elim H1.
+ elim H3.
+ elim H4.
+ apply (ex_intro ? ? a).
+ split
+ [ auto
+ (*split
+ [ apply le_S.
+ assumption
+ | assumption
+ ]*)
+ | intros.
+ elim (le_to_or_lt_eq i (S n1) H9)
+ [ auto
+ (*apply H5
+ [ assumption
+ | apply le_S_S_to_le.
+ assumption
+ ]*)
+ | auto
+ (*rewrite > H10.
+ assumption*)
+ ]
+ ]
+ ]
+ ]
+]
+qed.
+
+lemma decidable_n1:\forall p.\forall n,j. j \le n \to (p j)=true \to
+(\forall m. j < m \to m \le n \to (p m) = false) \lor
+(\exists m. j < m \land m \le n \land (p m) = true \land
+\forall i. m < i \to i \le n \to (p i) = false).
+intros.
+elim (decidable_n p n)
+[ absurd ((p j)=true)
+ [ assumption
+ | unfold.
+ intro.
+ apply not_eq_true_false.
+ auto
+ (*rewrite < H3.
+ apply H2.
+ assumption*)
+ ]
+| elim H2.
+ clear H2.
+ apply (nat_compare_elim j a)
+ [ intro.
+ right.
+ apply (ex_intro ? ? a).
+ elim H3.clear H3.
+ elim H4.clear H4.
+ split
+ [ auto
+ (*split
+ [ split
+ [ assumption
+ | assumption
+ ]
+ | assumption
+ ]*)
+ | assumption
+ ]
+ | intro.
+ rewrite > H2.
+ left.
+ elim H3 2.
+ (*qui la tattica auto non conclude il goal, concluso invece con l'invocazione
+ *della sola tattica assumption
+ *)
+ assumption
+ | intro.
+ absurd (p j = true)
+ [ assumption
+ | unfold.
+ intro.
+ apply not_eq_true_false.
+ rewrite < H4.
+ elim H3.
+ auto
+ (*clear H3.
+ apply (H6 j H2).assumption*)
+ ]
+ ]
+]
+qed.
+
+lemma decidable_n2:\forall p.\forall n,j. j \le n \to (p j)=true \to
+(\forall m. j < m \to m \le n \to (p m) = false) \lor
+(\exists m. j < m \land m \le n \land (p m) = true \land
+\forall i. j < i \to i < m \to (p i) = false).
+intros 3.
+elim n
+[ left.
+ apply (le_n_O_elim j H).
+ intros.
+ absurd (m \le O)
+ [ assumption
+ | auto
+ (*apply lt_to_not_le.
+ assumption*)
+ ]
+| elim (le_to_or_lt_eq ? ? H1)
+ [ cut (j \le n1)
+ [ elim (H Hcut H2)
+ [ apply (bool_elim ? (p (S n1)))
+ [ intro.
+ right.
+ apply (ex_intro ? ? (S n1)).
+ split
+ [ auto
+ (*split
+ [ split
+ [ assumption
+ | apply le_n
+ ]
+ | assumption
+ ]*)
+ | intros.
+ auto
+ (*apply (H4 i H6).
+ apply le_S_S_to_le.
+ assumption*)
+ ]
+ | intro.
+ left.
+ intros.
+ elim (le_to_or_lt_eq ? ? H7)
+ [ auto
+ (*apply H4
+ [ assumption
+ | apply le_S_S_to_le.
+ assumption
+ ]*)
+ | auto
+ (*rewrite > H8.
+ assumption*)
+ ]
+ ]
+ | right.
+ elim H4.clear H4.
+ elim H5.clear H5.
+ elim H4.clear H4.
+ elim H5.clear H5.
+ apply (ex_intro ? ? a).
+ split
+ [ split
+ [ auto
+ (*split
+ [ assumption
+ | apply le_S.
+ assumption
+ ]*)
+ | assumption
+ ]
+ | (*qui auto non chiude il goal, chiuso invece mediante l'invocazione
+ *della sola tattica assumption
+ *)
+ assumption
+ ]
+ ]
+ | auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ ]
+ | left.
+ intros.
+ absurd (j < m)
+ [ assumption
+ | apply le_to_not_lt.
+ rewrite > H3.
+ assumption
+ ]
+ ]
+]
+qed.
+
+(* tutti d spostare *)
+theorem lt_minus_to_lt_plus:
+\forall n,m,p. n - m < p \to n < m + p.
+intros 2.
+apply (nat_elim2 ? ? ? ? n m)
+[ simplify.
+ intros.
+ auto
+| intros 2.
+ auto
+ (*rewrite < minus_n_O.
+ intro.
+ assumption*)
+| intros.
+ simplify.
+ cut (n1 < m1+p)
+ [ auto
+ | apply H.
+ apply H1
+ ]
+]
+qed.
+
+theorem lt_plus_to_lt_minus:
+\forall n,m,p. m \le n \to n < m + p \to n - m < p.
+intros 2.
+apply (nat_elim2 ? ? ? ? n m)
+[ simplify.
+ intros 3.
+ apply (le_n_O_elim ? H).
+ auto
+ (*simplify.
+ intros.
+ assumption*)
+| simplify.
+ intros.
+ assumption
+| intros.
+ simplify.
+ apply H
+ [ auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ | apply le_S_S_to_le.
+ apply H2
+ ]
+]
+qed.
+
+theorem minus_m_minus_mn: \forall n,m. n\le m \to n=m-(m-n).
+intros.
+apply sym_eq.
+auto.
+(*apply plus_to_minus.
+auto.*)
+qed.
+
+theorem eq_map_iter_p_transpose2: \forall p.\forall f.associative nat f \to
+symmetric2 nat nat f \to \forall g. \forall a,k,n:nat. O < k \to k \le n \to
+(p (S n) = true) \to (p k) = true
+\to map_iter_p (S n) p g a f = map_iter_p (S n) p (\lambda m. g (transpose k (S n) m)) a f.
+intros 10.
+cut (k = (S n)-(S n -k))
+[ generalize in match H3.
+ clear H3.
+ generalize in match g.
+ generalize in match H2.
+ clear H2.
+ rewrite > Hcut.
+ (*generalize in match Hcut.clear Hcut.*)
+ (* generalize in match H3.clear H3.*)
+ (* something wrong here
+ rewrite > Hcut in \vdash (?\rarr ?\rarr %). *)
+ apply (nat_elim1 (S n - k)).
+ intros.
+ elim (decidable_n2 p n (S n -m) H4 H6)
+ [ apply (eq_map_iter_p_transpose1 p f H H1 f1 a)
+ [ assumption
+ | auto
+ (*unfold.
+ auto*)
+ | apply le_n
+ | assumption
+ | assumption
+ | intros.
+ auto
+ (*apply H7
+ [ assumption
+ | apply le_S_S_to_le.
+ assumption
+ ]*)
+ ]
+ | elim H7.
+ clear H7.
+ elim H8.clear H8.
+ elim H7.clear H7.
+ elim H8.clear H8.
+ apply (trans_eq ? ?
+ (map_iter_p (S n) p (\lambda i.f1 (transpose a1 (S n) (transpose (S n -m) a1 i))) a f))
+ [ apply (trans_eq ? ?
+ (map_iter_p (S n) p (\lambda i.f1 (transpose a1 (S n) i)) a f))
+ [ cut (a1 = (S n -(S n -a1)))
+ [ rewrite > Hcut1.
+ apply H2
+ [ apply lt_plus_to_lt_minus
+ [ auto
+ (*apply le_S.
+ assumption*)
+ | rewrite < sym_plus.
+ auto
+ (*apply lt_minus_to_lt_plus.
+ assumption*)
+ ]
+ | rewrite < Hcut1.
+ auto
+ (*apply (trans_lt ? (S n -m));
+ assumption*)
+ | rewrite < Hcut1.
+ assumption
+ | assumption
+ | auto
+ (*rewrite < Hcut1.
+ assumption*)
+ ]
+ | auto
+ (*apply minus_m_minus_mn.
+ apply le_S.
+ assumption*)
+ ]
+ | apply (eq_map_iter_p_transpose1 p f H H1)
+ [ assumption
+ | assumption
+ | auto
+ (*apply le_S.
+ assumption*)
+ | assumption
+ | assumption
+ | (*qui auto non chiude il goal, chiuso dall'invocazione della singola
+ * tattica assumption
+ *)
+ assumption
+ ]
+ ]
+ | apply (trans_eq ? ?
+ (map_iter_p (S n) p (\lambda i.f1 (transpose a1 (S n) (transpose (S n -m) a1 (transpose (S n -(S n -a1)) (S n) i)))) a f))
+ [ cut (a1 = (S n) -(S n -a1))
+ [ apply H2
+ [ apply lt_plus_to_lt_minus
+ [ auto
+ (*apply le_S.
+ assumption*)
+ | rewrite < sym_plus.
+ auto
+ (*apply lt_minus_to_lt_plus.
+ assumption*)
+ ]
+ | rewrite < Hcut1.
+ auto
+ (*apply (trans_lt ? (S n -m))
+ [ assumption
+ | assumption
+ ]*)
+ | rewrite < Hcut1.
+ assumption
+ | assumption
+ | auto
+ (*rewrite < Hcut1.
+ assumption*)
+ ]
+ | auto
+ (*apply minus_m_minus_mn.
+ apply le_S.
+ assumption*)
+ ]
+ | apply eq_map_iter_p.
+ cut (a1 = (S n) -(S n -a1))
+ [ intros.
+ apply eq_f.
+ rewrite < Hcut1.
+ rewrite < transpose_i_j_j_i.
+ rewrite > (transpose_i_j_j_i (S n -m)).
+ rewrite > (transpose_i_j_j_i a1 (S n)).
+ rewrite > (transpose_i_j_j_i (S n -m)).
+ apply sym_eq.
+ apply eq_transpose
+ [ unfold.
+ intro.
+ apply (not_le_Sn_n n).
+ rewrite < H12.
+ assumption
+ | unfold.
+ intro.
+ apply (not_le_Sn_n n).
+ rewrite > H12.
+ assumption
+ | unfold.
+ intro.
+ apply (not_le_Sn_n a1).
+ rewrite < H12 in \vdash (? (? %) ?).
+ assumption
+ ]
+ | auto
+ (*apply minus_m_minus_mn.
+ apply le_S.
+ assumption*)
+ ]
+ ]
+ ]
+ ]
+| auto
+ (*apply minus_m_minus_mn.
+ apply le_S.
+ assumption*)
+]
+qed.
+
+theorem eq_map_iter_p_transpose3: \forall p.\forall f.associative nat f \to
+symmetric2 nat nat f \to \forall g. \forall a,k,n:nat. O < k \to k \le (S n) \to
+(p (S n) = true) \to (p k) = true
+\to map_iter_p (S n) p g a f = map_iter_p (S n) p (\lambda m. g (transpose k (S n) m)) a f.
+intros.
+elim (le_to_or_lt_eq ? ? H3)
+[ apply (eq_map_iter_p_transpose2 p f H H1 g a k n H2);auto
+ (*[ apply le_S_S_to_le.
+ assumption
+ | assumption
+ | assumption
+ ]*)
+| rewrite > H6.
+ apply eq_map_iter_p.
+ intros.
+ auto
+ (*apply eq_f.
+ apply sym_eq.
+ apply transpose_i_i.*)
+]
+qed.
+
+lemma permut_p_O: \forall p.\forall h.\forall n.
+permut_p h p n \to p O = false \to \forall m. (S m) \le n \to p (S m) = true \to O < h(S m).
+intros.
+unfold permut_p in H.
+apply not_le_to_lt.unfold.
+intro.
+elim (H (S m) H2 H3).
+elim H5.
+absurd (p (h (S m)) = true)
+[ assumption
+| apply (le_n_O_elim ? H4).
+ unfold.
+ intro.
+ (*l'invocazione di auto in questo punto genera segmentation fault*)
+ apply not_eq_true_false.
+ (*l'invocazione di auto in questo punto genera segmentation fault*)
+ rewrite < H9.
+ (*l'invocazione di auto in questo punto genera segmentation fault*)
+ rewrite < H1.
+ (*l'invocazione di auto in questo punto genera segmentation fault*)
+ reflexivity
+]
+qed.
+
+theorem eq_map_iter_p_permut: \forall p.\forall f.associative nat f \to
+symmetric2 nat nat f \to \forall n.\forall g. \forall h.\forall a:nat.
+permut_p h p n \to p O = false \to
+map_iter_p n p g a f = map_iter_p n p (compose ? ? ? g h) a f .
+intros 5.
+elim n
+[ auto
+ (*simplify.
+ reflexivity*)
+| apply (bool_elim ? (p (S n1)))
+ [ intro.
+ apply (trans_eq ? ? (map_iter_p (S n1) p (\lambda m.g ((transpose (h (S n1)) (S n1)) m)) a f))
+ [ unfold permut_p in H3.
+ elim (H3 (S n1) (le_n ?) H5).
+ elim H6.
+ clear H6.
+ apply (eq_map_iter_p_transpose3 p f H H1 g a (h(S n1)) n1)
+ [ apply (permut_p_O ? ? ? H3 H4);auto
+ (*[ apply le_n
+ | assumption
+ ]*)
+ | assumption
+ | assumption
+ | assumption
+ ]
+ | apply (trans_eq ? ? (map_iter_p (S n1) p (\lambda m.
+ (g(transpose (h (S n1)) (S n1)
+ (transpose (h (S n1)) (S n1) (h m)))) ) a f))
+ [ rewrite > (map_iter_p_S_true ? ? ? ? ? H5).
+ rewrite > (map_iter_p_S_true ? ? ? ? ? H5).
+ apply eq_f2
+ [ rewrite > transpose_i_j_j.
+ rewrite > transpose_i_j_i.
+ rewrite > transpose_i_j_j.
+ reflexivity
+ | apply (H2 (\lambda m.(g(transpose (h (S n1)) (S n1) m))) ?)
+ [ unfold.
+ intros.
+ split
+ [ split
+ [ simplify.
+ unfold permut_p in H3.
+ elim (H3 i (le_S ? ? H6) H7).
+ elim H8.
+ clear H8.
+ elim (le_to_or_lt_eq ? ? H10)
+ [ unfold transpose.
+ rewrite > (not_eq_to_eqb_false ? ? (lt_to_not_eq ? ? H8)).
+ cut (h i \neq h (S n1))
+ [ rewrite > (not_eq_to_eqb_false ? ? Hcut).
+ simplify.
+ auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ | apply H9
+ [ apply H5
+ | apply le_n
+ | apply lt_to_not_eq.
+ auto
+ (*unfold.
+ apply le_S_S.
+ assumption*)
+ ]
+ ]
+ | rewrite > H8.
+ apply (eqb_elim (S n1) (h (S n1)))
+ [ intro.
+ absurd (h i = h (S n1))
+ [ auto
+ (*rewrite > H8.
+ assumption*)
+ | apply H9
+ [ assumption
+ | apply le_n
+ | apply lt_to_not_eq.
+ auto
+ (*unfold.
+ apply le_S_S.
+ assumption*)
+ ]
+ ]
+ | intro.
+ unfold transpose.
+ rewrite > (not_eq_to_eqb_false ? ? H12).
+ rewrite > (eq_to_eqb_true ? ? (refl_eq ? (S n1))).
+ simplify.
+ elim (H3 (S n1) (le_n ? ) H5).
+ elim H13.clear H13.
+ elim (le_to_or_lt_eq ? ? H15)
+ [ auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ | apply False_ind.
+ auto
+ (*apply H12.
+ apply sym_eq.
+ assumption*)
+ ]
+ ]
+ ]
+
+ | simplify.
+ unfold permut_p in H3.
+ unfold transpose.
+ apply (eqb_elim (h i) (S n1))
+ [ intro.
+ apply (eqb_elim (h i) (h (S n1)))
+ [ intro.
+ (*NB: qui auto non chiude il goal*)
+ simplify.
+ assumption
+ | intro.
+ simplify.
+ elim (H3 (S n1) (le_n ? ) H5).
+ auto
+ (*elim H10.
+ assumption*)
+ ]
+ | intro.
+ apply (eqb_elim (h i) (h (S n1)))
+ [ intro.
+ (*NB: qui auto non chiude il goal*)
+ simplify.
+ assumption
+ | intro.
+ simplify.
+ elim (H3 i (le_S ? ? H6) H7).
+ auto
+ (*elim H10.
+ assumption*)
+ ]
+ ]
+ ]
+ | simplify.
+ intros.
+ unfold Not.
+ intro.
+ unfold permut_p in H3.
+ elim (H3 i (le_S i ? H6) H7).
+ apply (H13 j H8 (le_S j ? H9) H10).
+ apply (injective_transpose ? ? ? ? H11)
+ ]
+ | assumption
+ ]
+ ]
+ | apply eq_map_iter_p.
+ intros.
+ auto
+ (*rewrite > transpose_transpose.
+ reflexivity*)
+ ]
+ ]
+ | intro.
+ rewrite > (map_iter_p_S_false ? ? ? ? ? H5).
+ rewrite > (map_iter_p_S_false ? ? ? ? ? H5).
+ apply H2
+ [ unfold permut_p.
+ unfold permut_p in H3.
+ intros.
+ elim (H3 i (le_S i ? H6) H7).
+ elim H8.
+ split
+ [ split
+ [ elim (le_to_or_lt_eq ? ? H10)
+ [ auto
+ (*apply le_S_S_to_le.assumption*)
+ | absurd (p (h i) = true)
+ [ assumption
+ | rewrite > H12.
+ rewrite > H5.
+ unfold.intro.
+ (*l'invocazione di auto qui genera segmentation fault*)
+ apply not_eq_true_false.
+ (*l'invocazione di auto qui genera segmentation fault*)
+ apply sym_eq.
+ (*l'invocazione di auto qui genera segmentation fault*)
+ assumption
+ ]
+ ]
+ | assumption
+ ]
+ | intros.
+ apply H9;auto
+ (*[ assumption
+ | apply (le_S ? ? H13)
+ | assumption
+ ]*)
+ ]
+ | assumption
+
+ ]
+
+ ]
+
+]
+qed.
+
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / Matita is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/minimization".
+
+include "nat/minus.ma".
+
+let rec max i f \def
+ match (f i) with
+ [ true \Rightarrow i
+ | false \Rightarrow
+ match i with
+ [ O \Rightarrow O
+ | (S j) \Rightarrow max j f ]].
+
+theorem max_O_f : \forall f: nat \to bool. max O f = O.
+intro. simplify.
+elim (f O); auto.
+(*[ simplify.
+ reflexivity
+| simplify.
+ reflexivity
+]*)
+qed.
+
+theorem max_S_max : \forall f: nat \to bool. \forall n:nat.
+(f (S n) = true \land max (S n) f = (S n)) \lor
+(f (S n) = false \land max (S n) f = max n f).
+intros.simplify.elim (f (S n));auto.
+(*[ simplify.
+ left.
+ split;reflexivity
+| simplify.
+ right.
+ split;reflexivity.
+]*)
+qed.
+
+theorem le_max_n : \forall f: nat \to bool. \forall n:nat.
+max n f \le n.
+intros.
+elim n
+[ rewrite > max_O_f.
+ apply le_n
+| simplify.
+ elim (f (S n1));simplify;auto.
+ (*[ simplify.
+ apply le_n
+ | simplify.
+ apply le_S.
+ assumption
+ ]*)
+]
+qed.
+
+theorem le_to_le_max : \forall f: nat \to bool. \forall n,m:nat.
+n\le m \to max n f \le max m f.
+intros.
+elim H
+[ apply le_n
+| apply (trans_le ? (max n1 f))
+ [ apply H2
+ | cut ((f (S n1) = true \land max (S n1) f = (S n1)) \lor
+ (f (S n1) = false \land max (S n1) f = max n1 f))
+ [ elim Hcut;elim H3;rewrite > H5;auto
+ (*[ elim H3.
+ rewrite > H5.
+ apply le_S.
+ apply le_max_n.
+ | elim H3.
+ rewrite > H5.
+ apply le_n.
+ ]*)
+ | apply max_S_max.
+ ]
+ ]
+]
+qed.
+
+theorem f_m_to_le_max: \forall f: nat \to bool. \forall n,m:nat.
+m\le n \to f m = true \to m \le max n f.
+intros 3.
+elim n
+[ auto.
+ (*apply (le_n_O_elim m H).
+ apply le_O_n.*)
+| apply (le_n_Sm_elim m n1 H1);intro
+ [ apply (trans_le ? (max n1 f)); auto
+ (*[ apply H
+ [apply le_S_S_to_le.
+ assumption
+ | assumption
+ ]
+ | apply le_to_le_max.
+ apply le_n_Sn.
+ ]*)
+ | simplify.
+ rewrite < H3.
+ rewrite > H2.
+ auto
+ (*simplify.
+ apply le_n.*)
+ ]
+]
+qed.
+
+
+definition max_spec \def \lambda f:nat \to bool.\lambda n: nat.
+\exists i. (le i n) \land (f i = true) \to
+(f n) = true \land (\forall i. i < n \to (f i = false)).
+
+theorem f_max_true : \forall f:nat \to bool. \forall n:nat.
+(\exists i:nat. le i n \land f i = true) \to f (max n f) = true.
+intros 2.
+elim n
+[ elim H.
+ elim H1.
+ generalize in match H3.
+ apply (le_n_O_elim a H2).
+ auto
+ (*intro.
+ simplify.
+ rewrite > H4.
+ simplify.
+ assumption.*)
+| simplify.
+ apply (bool_ind (\lambda b:bool.
+ (f (S n1) = b) \to (f (match b in bool with
+ [ true \Rightarrow (S n1)
+ | false \Rightarrow (max n1 f)])) = true))
+
+ [ auto
+ (*simplify.
+ intro.
+ assumption.*)
+ | simplify.
+ intro.
+ apply H.
+ elim H1.
+ elim H3.
+ generalize in match H5.
+ apply (le_n_Sm_elim a n1 H4)
+ [ intros.
+ apply (ex_intro nat ? a).
+ auto
+ (*split
+ [ apply le_S_S_to_le.
+ assumption.
+ | assumption.
+ ]*)
+ | intros.
+ (* una chiamata di auto in questo punto genera segmentation fault*)
+ apply False_ind.
+ (* una chiamata di auto in questo punto genera segmentation fault*)
+ apply not_eq_true_false.
+ (* una chiamata di auto in questo punto genera segmentation fault*)
+ rewrite < H2.
+ (* una chiamata di auto in questo punto genera segmentation fault*)
+ rewrite < H7.
+ (* una chiamata di auto in questo punto genera segmentation fault*)
+ rewrite > H6.
+ reflexivity.
+ ]
+ | reflexivity.
+ ]
+]
+qed.
+
+theorem lt_max_to_false : \forall f:nat \to bool.
+\forall n,m:nat. (max n f) < m \to m \leq n \to f m = false.
+intros 2.
+elim n
+[ absurd (le m O);auto
+ (*[ assumption
+ | cut (O < m)
+ [ apply (lt_O_n_elim m Hcut).
+ exact not_le_Sn_O.
+ | rewrite < (max_O_f f).
+ assumption.
+ ]
+ ]*)
+| generalize in match H1.
+ elim (max_S_max f n1)
+ [ elim H3.
+ absurd (m \le S n1)
+ [ assumption
+ | apply lt_to_not_le.
+ rewrite < H6.
+ assumption
+ ]
+ | elim H3.
+ apply (le_n_Sm_elim m n1 H2)
+ [ intro.
+ apply H;auto
+ (*[ rewrite < H6.
+ assumption
+ | apply le_S_S_to_le.
+ assumption
+ ]*)
+ | intro.
+ auto
+ (*rewrite > H7.
+ assumption*)
+ ]
+ ]
+]qed.
+
+let rec min_aux off n f \def
+ match f (n-off) with
+ [ true \Rightarrow (n-off)
+ | false \Rightarrow
+ match off with
+ [ O \Rightarrow n
+ | (S p) \Rightarrow min_aux p n f]].
+
+definition min : nat \to (nat \to bool) \to nat \def
+\lambda n.\lambda f. min_aux n n f.
+
+theorem min_aux_O_f: \forall f:nat \to bool. \forall i :nat.
+min_aux O i f = i.
+intros.simplify.
+(*una chiamata di auto a questo punto porta ad un'elaborazione molto lunga (forse va
+ in loop): dopo circa 3 minuti non era ancora terminata.
+ *)
+rewrite < minus_n_O.
+(*una chiamata di auto a questo punto porta ad un'elaborazione molto lunga (forse va
+ in loop): dopo circa 3 minuti non era ancora terminata.
+ *)
+elim (f i); auto.
+(*[ reflexivity.
+ simplify
+| reflexivity
+]*)
+qed.
+
+theorem min_O_f : \forall f:nat \to bool.
+min O f = O.
+intro.
+(* una chiamata di auto a questo punto NON conclude la dimostrazione*)
+apply (min_aux_O_f f O).
+qed.
+
+theorem min_aux_S : \forall f: nat \to bool. \forall i,n:nat.
+(f (n -(S i)) = true \land min_aux (S i) n f = (n - (S i))) \lor
+(f (n -(S i)) = false \land min_aux (S i) n f = min_aux i n f).
+intros.simplify.elim (f (n - (S i)));auto.
+(*[ simplify.
+ left.
+ split;reflexivity.
+| simplify.
+ right.
+ split;reflexivity.
+]*)
+qed.
+
+theorem f_min_aux_true: \forall f:nat \to bool. \forall off,m:nat.
+(\exists i. le (m-off) i \land le i m \land f i = true) \to
+f (min_aux off m f) = true.
+intros 2.
+elim off
+[ elim H.
+ elim H1.
+ elim H2.
+ cut (a = m)
+ [ auto.
+ (*rewrite > (min_aux_O_f f).
+ rewrite < Hcut.
+ assumption*)
+ | apply (antisym_le a m)
+ [ assumption
+ | rewrite > (minus_n_O m).
+ assumption
+ ]
+ ]
+| simplify.
+ apply (bool_ind (\lambda b:bool.
+ (f (m-(S n)) = b) \to (f (match b in bool with
+ [ true \Rightarrow m-(S n)
+ | false \Rightarrow (min_aux n m f)])) = true))
+ [ auto
+ (*simplify.
+ intro.
+ assumption.*)
+ | simplify.
+ intro.
+ apply H.
+ elim H1.
+ elim H3.
+ elim H4.
+ elim (le_to_or_lt_eq (m-(S n)) a H6)
+ [ apply (ex_intro nat ? a).
+ split;auto
+ (*[ auto.split
+ [ apply lt_minus_S_n_to_le_minus_n.
+ assumption
+ | assumption
+ ]
+ | assumption
+ ]*)
+ | absurd (f a = false)
+ [ (* una chiamata di auto in questo punto genera segmentation fault*)
+ rewrite < H8.
+ assumption.
+ | rewrite > H5.
+ apply not_eq_true_false
+ ]
+ ]
+ | reflexivity.
+ ]
+]
+qed.
+
+theorem lt_min_aux_to_false : \forall f:nat \to bool.
+\forall n,off,m:nat. (n-off) \leq m \to m < (min_aux off n f) \to f m = false.
+intros 3.
+elim off
+[absurd (le n m)
+ [ rewrite > minus_n_O.
+ assumption
+ | apply lt_to_not_le.
+ rewrite < (min_aux_O_f f n).
+ assumption
+ ]
+| generalize in match H1.
+ elim (min_aux_S f n1 n)
+ [ elim H3.
+ absurd (n - S n1 \le m)
+ [ assumption
+ | apply lt_to_not_le.
+ rewrite < H6.
+ assumption
+ ]
+ | elim H3.
+ elim (le_to_or_lt_eq (n -(S n1)) m)
+ [ apply H
+ [ auto
+ (*apply lt_minus_S_n_to_le_minus_n.
+ assumption*)
+ | rewrite < H6.
+ assumption
+ ]
+ | rewrite < H7.
+ assumption
+ | assumption
+ ]
+ ]
+]
+qed.
+
+theorem le_min_aux : \forall f:nat \to bool.
+\forall n,off:nat. (n-off) \leq (min_aux off n f).
+intros 3.
+elim off
+[ rewrite < minus_n_O.
+ auto
+ (*rewrite > (min_aux_O_f f n).
+ apply le_n.*)
+| elim (min_aux_S f n1 n)
+ [ elim H1.
+ auto
+ (*rewrite > H3.
+ apply le_n.*)
+ | elim H1.
+ rewrite > H3.
+ auto
+ (*apply (trans_le (n-(S n1)) (n-n1))
+ [ apply monotonic_le_minus_r.
+ apply le_n_Sn.
+ | assumption.
+ ]*)
+ ]
+]
+qed.
+
+theorem le_min_aux_r : \forall f:nat \to bool.
+\forall n,off:nat. (min_aux off n f) \le n.
+intros.
+elim off
+[ simplify.
+ rewrite < minus_n_O.
+ elim (f n);auto
+ (*[simplify.
+ apply le_n.
+ | simplify.
+ apply le_n.
+ ]*)
+| simplify.
+ elim (f (n -(S n1)));simplify;auto
+ (*[ apply le_plus_to_minus.
+ rewrite < sym_plus.
+ apply le_plus_n
+ | assumption
+ ]*)
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+
+set "baseuri" "cic:/matita/nat/minus".
+
+include "nat/le_arith.ma".
+include "nat/compare.ma".
+
+let rec minus n m \def
+ match n with
+ [ O \Rightarrow O
+ | (S p) \Rightarrow
+ match m with
+ [O \Rightarrow (S p)
+ | (S q) \Rightarrow minus p q ]].
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "natural minus" 'minus x y = (cic:/matita/nat/minus/minus.con x y).
+
+theorem minus_n_O: \forall n:nat.n=n-O.
+intros.
+elim n;
+auto. (* applico auto su entrambi i goal aperti*)
+(*simplify;reflexivity.*)
+qed.
+
+theorem minus_n_n: \forall n:nat.O=n-n.
+intros.
+elim n;simplify
+[ reflexivity
+| apply H
+]
+qed.
+
+theorem minus_Sn_n: \forall n:nat. S O = (S n)-n.
+intro.
+elim n
+[ auto
+ (*simplify.reflexivity.*)
+| elim H.
+ reflexivity
+]
+qed.
+
+
+theorem minus_Sn_m: \forall n,m:nat. m \leq n \to (S n)-m = S (n-m).
+intros 2.
+apply (nat_elim2
+(\lambda n,m.m \leq n \to (S n)-m = S (n-m)));intros
+[ apply (le_n_O_elim n1 H).
+ auto
+ (*simplify.
+ reflexivity.*)
+| auto
+ (*simplify.
+ reflexivity.*)
+| rewrite < H
+ [ reflexivity
+ | auto
+ (*apply le_S_S_to_le.
+ assumption.*)
+ ]
+]
+qed.
+
+theorem plus_minus:
+\forall n,m,p:nat. m \leq n \to (n-m)+p = (n+p)-m.
+intros 2.
+apply (nat_elim2
+(\lambda n,m.\forall p:nat.m \leq n \to (n-m)+p = (n+p)-m));intros
+[ apply (le_n_O_elim ? H).
+ auto
+ (*simplify.
+ rewrite < minus_n_O.
+ reflexivity.*)
+| auto
+ (*simplify.
+ reflexivity.*)
+| simplify.
+ auto
+ (*apply H.
+ apply le_S_S_to_le.
+ assumption.*)
+]
+qed.
+
+theorem minus_plus_m_m: \forall n,m:nat.n = (n+m)-m.
+intros 2.
+generalize in match n.
+elim m
+[ rewrite < minus_n_O.
+ apply plus_n_O.
+| elim n2
+ [ auto
+ (*simplify.
+ apply minus_n_n.*)
+ | rewrite < plus_n_Sm.
+ change with (S n3 = (S n3 + n1)-n1).
+ apply H
+ ]
+]
+qed.
+
+theorem plus_minus_m_m: \forall n,m:nat.
+m \leq n \to n = (n-m)+m.
+intros 2.
+apply (nat_elim2 (\lambda n,m.m \leq n \to n = (n-m)+m));intros
+[ apply (le_n_O_elim n1 H).
+ reflexivity
+| auto
+ (*simplify.
+ rewrite < plus_n_O.
+ reflexivity.*)
+| simplify.
+ rewrite < sym_plus.
+ simplify.
+ apply eq_f.
+ rewrite < sym_plus.
+ auto
+ (*apply H.
+ apply le_S_S_to_le.
+ assumption.*)
+]
+qed.
+
+theorem minus_to_plus :\forall n,m,p:nat.m \leq n \to n-m = p \to
+n = m+p.
+intros.apply (trans_eq ? ? ((n-m)+m));auto.
+(*[ apply plus_minus_m_m.
+ apply H.
+| elim H1.
+ apply sym_plus.
+]*)
+qed.
+
+theorem plus_to_minus :\forall n,m,p:nat.
+n = m+p \to n-m = p.
+intros.
+apply (inj_plus_r m).
+rewrite < H.
+rewrite < sym_plus.
+symmetry.
+auto.
+(*apply plus_minus_m_m.
+rewrite > H.
+rewrite > sym_plus.
+apply le_plus_n.*)
+qed.
+
+theorem minus_S_S : \forall n,m:nat.
+eq nat (minus (S n) (S m)) (minus n m).
+intros.
+reflexivity.
+qed.
+
+theorem minus_pred_pred : \forall n,m:nat. lt O n \to lt O m \to
+eq nat (minus (pred n) (pred m)) (minus n m).
+intros.
+apply (lt_O_n_elim n H).
+intro.
+apply (lt_O_n_elim m H1).
+intro.
+auto.
+(*simplify.reflexivity.*)
+qed.
+
+theorem eq_minus_n_m_O: \forall n,m:nat.
+n \leq m \to n-m = O.
+intros 2.
+apply (nat_elim2 (\lambda n,m.n \leq m \to n-m = O));intros
+[ auto
+ (*simplify.
+ reflexivity.*)
+| apply False_ind.
+ auto
+ (*apply not_le_Sn_O.
+ goal 15.*) (*prima goal 13*)
+(* effettuando un'esecuzione passo-passo, quando si arriva a dover
+ considerare questa tattica, la finestra di dimostrazione scompare
+ e viene generato il seguente errore:
+ Uncaught exception: File "matitaMathView.ml", line 677, characters
+ 6-12: Assertion failed.
+
+ tuttavia l'esecuzione continua, ed il teorema viene comunque
+ dimostrato.
+ *)
+ (*apply H.*)
+| simplify.
+ auto
+ (*apply H.
+ apply le_S_S_to_le.
+ apply H1.*)
+]
+qed.
+
+theorem le_SO_minus: \forall n,m:nat.S n \leq m \to S O \leq m-n.
+intros.
+elim H
+[ elim (minus_Sn_n n).apply le_n
+| rewrite > minus_Sn_m;auto
+ (*apply le_S.assumption.
+ apply lt_to_le.assumption.*)
+]
+qed.
+
+theorem minus_le_S_minus_S: \forall n,m:nat. m-n \leq S (m-(S n)).
+intros.apply (nat_elim2 (\lambda n,m.m-n \leq S (m-(S n))));intros
+[ elim n1;simplify
+ [ apply le_n_Sn.
+ | rewrite < minus_n_O.
+ apply le_n.
+ ]
+| auto
+ (*simplify.apply le_n_Sn.*)
+| simplify.apply H.
+]
+qed.
+
+theorem lt_minus_S_n_to_le_minus_n : \forall n,m,p:nat. m-(S n) < p \to m-n \leq p.
+intros 3.
+auto.
+(*simplify.
+intro.
+apply (trans_le (m-n) (S (m-(S n))) p).
+apply minus_le_S_minus_S.
+assumption.*)
+qed.
+
+theorem le_minus_m: \forall n,m:nat. n-m \leq n.
+intros.apply (nat_elim2 (\lambda m,n. n-m \leq n));intros
+[ auto
+ (*rewrite < minus_n_O.
+ apply le_n.*)
+| auto
+ (*simplify.
+ apply le_n.*)
+| simplify.
+ auto
+ (*apply le_S.
+ assumption.*)
+]
+qed.
+
+theorem lt_minus_m: \forall n,m:nat. O < n \to O < m \to n-m \lt n.
+intros.
+apply (lt_O_n_elim n H).
+intro.
+apply (lt_O_n_elim m H1).
+intro.
+simplify.
+auto.
+(*unfold lt.
+apply le_S_S.
+apply le_minus_m.*)
+qed.
+
+theorem minus_le_O_to_le: \forall n,m:nat. n-m \leq O \to n \leq m.
+intros 2.
+apply (nat_elim2 (\lambda n,m:nat.n-m \leq O \to n \leq m))
+[ intros.
+ apply le_O_n
+| simplify.
+ intros.
+ assumption
+| simplify.
+ intros.
+ auto
+ (*apply le_S_S.
+ apply H.
+ assumption.*)
+]
+qed.
+
+(* galois *)
+theorem monotonic_le_minus_r:
+\forall p,q,n:nat. q \leq p \to n-p \le n-q.
+(*simplify*).
+intros 2.
+apply (nat_elim2
+(\lambda p,q.\forall a.q \leq p \to a-p \leq a-q));intros
+[ apply (le_n_O_elim n H).
+ apply le_n.
+| rewrite < minus_n_O.
+ apply le_minus_m.
+| elim a
+ [ auto
+ (*simplify.
+ apply le_n.*)
+ | simplify.
+ auto
+ (*apply H.
+ apply le_S_S_to_le.
+ assumption.*)
+ ]
+]
+qed.
+
+theorem le_minus_to_plus: \forall n,m,p. (le (n-m) p) \to (le n (p+m)).
+intros 2.
+apply (nat_elim2 (\lambda n,m.\forall p.(le (n-m) p) \to (le n (p+m))));intros
+[ apply le_O_n.
+| rewrite < plus_n_O.
+ assumption.
+| rewrite < plus_n_Sm.
+ apply le_S_S.
+ apply H.
+ exact H1.
+]
+qed.
+
+theorem le_plus_to_minus: \forall n,m,p. (le n (p+m)) \to (le (n-m) p).
+intros 2.
+apply (nat_elim2 (\lambda n,m.\forall p.(le n (p+m)) \to (le (n-m) p)))
+[ intros.
+ auto
+ (*simplify.
+ apply le_O_n.*)
+| intros 2.
+ rewrite < plus_n_O.
+ auto
+ (*intro.
+ simplify.
+ assumption.*)
+| intros.
+ simplify.
+ apply H.
+ apply le_S_S_to_le.
+ rewrite > plus_n_Sm.
+ assumption
+]
+qed.
+
+(* the converse of le_plus_to_minus does not hold *)
+theorem le_plus_to_minus_r: \forall n,m,p. (le (n+m) p) \to (le n (p-m)).
+intros 3.
+apply (nat_elim2 (\lambda m,p.(le (n+m) p) \to (le n (p-m))));intro
+[ rewrite < plus_n_O.
+ rewrite < minus_n_O.
+ auto
+ (*intro.
+ assumption.*)
+| intro.
+ cut (n=O)
+ [ auto
+ (*rewrite > Hcut.
+ apply le_O_n.*)
+ | apply sym_eq.
+ apply le_n_O_to_eq.
+ auto
+ (*apply (trans_le ? (n+(S n1)))
+ [ rewrite < sym_plus.
+ apply le_plus_n
+ | assumption
+ ]*)
+ ]
+| intros.
+ simplify.
+ apply H.
+ apply le_S_S_to_le.
+ rewrite > plus_n_Sm.
+ assumption
+]
+qed.
+
+(* minus and lt - to be completed *)
+theorem lt_minus_to_plus: \forall n,m,p. (lt n (p-m)) \to (lt (n+m) p).
+intros 3.
+apply (nat_elim2 (\lambda m,p.(lt n (p-m)) \to (lt (n+m) p)))
+[ intro.
+ rewrite < plus_n_O.
+ rewrite < minus_n_O.
+ auto
+ (*intro.
+ assumption.*)
+| simplify.
+ intros.
+ apply False_ind.
+ apply (not_le_Sn_O n H)
+| (*simplify.*)
+ intros.
+ unfold lt.
+ apply le_S_S.
+ rewrite < plus_n_Sm.
+ auto
+ (*apply H.
+ apply H1.*)
+]
+qed.
+
+theorem distributive_times_minus: distributive nat times minus.
+unfold distributive.
+intros.
+apply ((leb_elim z y));intro
+[ cut (x*(y-z)+x*z = (x*y-x*z)+x*z)
+ [ auto
+ (*apply (inj_plus_l (x*z)).
+ assumption.*)
+ | apply (trans_eq nat ? (x*y))
+ [ rewrite < distr_times_plus.
+ auto
+ (*rewrite < (plus_minus_m_m ? ? H).
+ reflexivity.*)
+ | rewrite < plus_minus_m_m;auto
+ (*[ reflexivity.
+ | apply le_times_r.
+ assumption.
+ ]*)
+ ]
+ ]
+| rewrite > eq_minus_n_m_O
+ [ rewrite > (eq_minus_n_m_O (x*y))
+ [ auto
+ (*rewrite < sym_times.
+ simplify.
+ reflexivity.*)
+ | apply le_times_r.
+ apply lt_to_le.
+ auto
+ (*apply not_le_to_lt.
+ assumption.*)
+ ]
+ | auto
+ (*apply lt_to_le.
+ apply not_le_to_lt.
+ assumption.*)
+ ]
+]
+qed.
+
+theorem distr_times_minus: \forall n,m,p:nat. n*(m-p) = n*m-n*p
+\def distributive_times_minus.
+
+theorem eq_minus_plus_plus_minus: \forall n,m,p:nat. p \le m \to (n+m)-p = n+(m-p).
+intros.
+apply plus_to_minus.
+rewrite > sym_plus in \vdash (? ? ? %).
+rewrite > assoc_plus.
+auto.
+(*rewrite < plus_minus_m_m.
+reflexivity.
+assumption.
+*)
+qed.
+
+theorem eq_minus_minus_minus_plus: \forall n,m,p:nat. (n-m)-p = n-(m+p).
+intros.
+cut (m+p \le n \or m+p \nleq n)
+[ elim Hcut
+ [ symmetry.
+ apply plus_to_minus.
+ rewrite > assoc_plus.
+ rewrite > (sym_plus p).
+ rewrite < plus_minus_m_m
+ [ rewrite > sym_plus.
+ rewrite < plus_minus_m_m ; auto
+ (*[ reflexivity.
+ | apply (trans_le ? (m+p))
+ [ rewrite < sym_plus.
+ apply le_plus_n
+ | assumption
+ ]
+ ]*)
+ | apply le_plus_to_minus_r.
+ rewrite > sym_plus.
+ assumption.
+ ]
+ | rewrite > (eq_minus_n_m_O n (m+p))
+ [ rewrite > (eq_minus_n_m_O (n-m) p)
+ [ reflexivity
+ | apply le_plus_to_minus.
+ apply lt_to_le.
+ rewrite < sym_plus.
+ auto
+ (*apply not_le_to_lt.
+ assumption.*)
+ ]
+ | auto
+ (*apply lt_to_le.
+ apply not_le_to_lt.
+ assumption.*)
+ ]
+ ]
+| apply (decidable_le (m+p) n)
+]
+qed.
+
+theorem eq_plus_minus_minus_minus: \forall n,m,p:nat. p \le m \to m \le n \to
+p+(n-m) = n-(m-p).
+intros.
+apply sym_eq.
+apply plus_to_minus.
+rewrite < assoc_plus.
+rewrite < plus_minus_m_m;
+[ rewrite < sym_plus.
+ auto
+ (*rewrite < plus_minus_m_m
+ [ reflexivity
+ | assumption
+ ]*)
+| assumption
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/nat".
+
+include "higher_order_defs/functions.ma".
+
+inductive nat : Set \def
+ | O : nat
+ | S : nat \to nat.
+
+definition pred: nat \to nat \def
+ \lambda n:nat. match n with
+ [ O \Rightarrow O
+ | (S p) \Rightarrow p ].
+
+theorem pred_Sn : \forall n:nat.n=(pred (S n)).
+ auto.
+ (*intros. reflexivity.*)
+qed.
+
+theorem injective_S : injective nat nat S.
+ unfold injective.
+ intros.
+ rewrite > pred_Sn.
+ rewrite > (pred_Sn y).
+ auto.
+ (*apply eq_f.
+ assumption.*)
+qed.
+
+theorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m \def
+ injective_S.
+
+theorem not_eq_S : \forall n,m:nat.
+ \lnot n=m \to S n \neq S m.
+ intros.
+ unfold Not.
+ intros.
+ auto.
+ (*apply H.
+ apply injective_S.
+ assumption.*)
+qed.
+
+definition not_zero : nat \to Prop \def
+ \lambda n: nat.
+ match n with
+ [ O \Rightarrow False
+ | (S p) \Rightarrow True ].
+
+theorem not_eq_O_S : \forall n:nat. O \neq S n.
+ intros.
+ unfold Not.
+ intros.
+ cut (not_zero O)
+ [exact Hcut
+ |rewrite > H.exact I]
+qed.
+
+theorem not_eq_n_Sn : \forall n:nat. n \neq S n.
+ intros.
+ elim n
+ [apply not_eq_O_S
+ |apply not_eq_S.assumption]
+qed.
+
+theorem nat_case:
+ \forall n:nat.\forall P:nat \to Prop.
+ P O \to (\forall m:nat. P (S m)) \to P n.
+ intros.
+
+ elim n
+ [ assumption
+ | apply H1 ]
+qed.
+
+theorem nat_case1:
+ \forall n:nat.\forall P:nat \to Prop.
+ (n=O \to P O) \to (\forall m:nat. (n=(S m) \to P (S m))) \to P n.
+intros 2;
+elim n
+ [ auto
+ (*apply H;
+ reflexivity*)
+ | auto
+ (*apply H2;
+ reflexivity*) ]
+qed.
+
+theorem nat_elim2 :
+ \forall R:nat \to nat \to Prop.
+ (\forall n:nat. R O n)
+ \to (\forall n:nat. R (S n) O)
+ \to (\forall n,m:nat. R n m \to R (S n) (S m))
+ \to \forall n,m:nat. R n m.
+ intros 5;
+ elim n
+ [ apply H
+ | apply (nat_case m)
+ [ apply H1
+ |intro;
+ auto
+ (*apply H2;
+ apply H3*)
+ ]
+ ]
+qed.
+
+theorem decidable_eq_nat : \forall n,m:nat.decidable (n=m).
+ intros.unfold decidable.
+ apply (nat_elim2 (\lambda n,m.(Or (n=m) ((n=m) \to False))))
+ [ intro; elim n1
+ [auto
+ (*left;
+ reflexivity*)
+ |auto
+ (*right;
+ apply not_eq_O_S*) ]
+ | intro;
+ right;
+ intro;
+ apply (not_eq_O_S n1);
+ auto
+ (*apply sym_eq;
+ assumption*)
+ | intros; elim H
+ [ auto
+ (*left;
+ apply eq_f;
+ assumption*)
+ | right;
+ intro;
+ auto
+ (*apply H1;
+ apply inj_S;
+ assumption*)
+ ]
+ ]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / Matita is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/nth_prime".
+
+include "nat/primes.ma".
+include "nat/lt_arith.ma".
+
+(* upper bound by Bertrand's conjecture. *)
+(* Too difficult to prove.
+let rec nth_prime n \def
+match n with
+ [ O \Rightarrow (S(S O))
+ | (S p) \Rightarrow
+ let previous_prime \def S (nth_prime p) in
+ min_aux previous_prime ((S(S O))*previous_prime) primeb].
+
+theorem example8 : nth_prime (S(S O)) = (S(S(S(S(S O))))).
+normalize.reflexivity.
+qed.
+
+theorem example9 : nth_prime (S(S(S O))) = (S(S(S(S(S(S(S O))))))).
+normalize.reflexivity.
+qed.
+
+theorem example10 : nth_prime (S(S(S(S O)))) = (S(S(S(S(S(S(S(S(S(S(S O))))))))))).
+normalize.reflexivity.
+qed. *)
+
+theorem smallest_factor_fact: \forall n:nat.
+n < smallest_factor (S n!).
+intros.
+apply not_le_to_lt.
+unfold Not.
+intro.
+apply (not_divides_S_fact n (smallest_factor(S n!)))
+[ apply lt_SO_smallest_factor.
+ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_SO_fact*)
+| assumption
+| auto
+ (*apply divides_smallest_factor_n.
+ unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+]
+qed.
+
+theorem ex_prime: \forall n. (S O) \le n \to \exists m.
+n < m \land m \le S n! \land (prime m).
+intros.
+elim H
+[ apply (ex_intro nat ? (S(S O))).
+ split;auto
+ (*[ split
+ [ apply (le_n (S(S O)))
+ | apply (le_n (S(S O)))
+ ]
+ | apply (primeb_to_Prop (S(S O)))
+ ]*)
+| apply (ex_intro nat ? (smallest_factor (S (S n1)!))).
+ split
+ [ auto
+ (*split
+ [ apply smallest_factor_fact
+ | apply le_smallest_factor_n
+ ]*)
+ | (* Andrea: ancora hint non lo trova *)
+ apply prime_smallest_factor_n.
+ auto
+ (*unfold lt.
+ apply le_S.
+ apply le_SSO_fact.
+ unfold lt.
+ apply le_S_S.
+ assumption*)
+ ]
+]
+qed.
+
+let rec nth_prime n \def
+match n with
+ [ O \Rightarrow (S(S O))
+ | (S p) \Rightarrow
+ let previous_prime \def (nth_prime p) in
+ let upper_bound \def S previous_prime! in
+ min_aux (upper_bound - (S previous_prime)) upper_bound primeb].
+
+(* it works, but nth_prime 4 takes already a few minutes -
+it must compute factorial of 7 ...*)
+
+theorem example11 : nth_prime (S(S O)) = (S(S(S(S(S O))))).
+auto.
+(*normalize.reflexivity.*)
+qed.
+
+theorem example12: nth_prime (S(S(S O))) = (S(S(S(S(S(S(S O))))))).
+auto.
+(*normalize.reflexivity.*)
+qed.
+
+theorem example13 : nth_prime (S(S(S(S O)))) = (S(S(S(S(S(S(S(S(S(S(S O))))))))))).
+auto.
+(*normalize.reflexivity.*)
+qed.
+
+(*
+theorem example14 : nth_prime (S(S(S(S(S O))))) = (S(S(S(S(S(S(S(S(S(S(S O))))))))))).
+normalize.reflexivity.
+*)
+
+theorem prime_nth_prime : \forall n:nat.prime (nth_prime n).
+intro.
+apply (nat_case n)
+[ auto
+ (*simplify.
+ apply (primeb_to_Prop (S(S O)))*)
+| intro.
+ change with
+ (let previous_prime \def (nth_prime m) in
+ let upper_bound \def S previous_prime! in
+ prime (min_aux (upper_bound - (S previous_prime)) upper_bound primeb)).
+ apply primeb_true_to_prime.
+ apply f_min_aux_true.
+ apply (ex_intro nat ? (smallest_factor (S (nth_prime m)!))).
+ split
+ [ split
+ [ cut (S (nth_prime m)!-(S (nth_prime m)! - (S (nth_prime m))) = (S (nth_prime m)))
+ [ rewrite > Hcut.
+ exact (smallest_factor_fact (nth_prime m))
+ | (* maybe we could factorize this proof *)
+ apply plus_to_minus.
+ auto
+ (*apply plus_minus_m_m.
+ apply le_S_S.
+ apply le_n_fact_n*)
+ ]
+ | apply le_smallest_factor_n
+ ]
+ | apply prime_to_primeb_true.
+ apply prime_smallest_factor_n.
+ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_SO_fact*)
+ ]
+]
+qed.
+
+(* properties of nth_prime *)
+theorem increasing_nth_prime: increasing nth_prime.
+unfold increasing.
+intros.
+change with
+(let previous_prime \def (nth_prime n) in
+let upper_bound \def S previous_prime! in
+(S previous_prime) \le min_aux (upper_bound - (S previous_prime)) upper_bound primeb).
+intros.
+cut (upper_bound - (upper_bound -(S previous_prime)) = (S previous_prime))
+[ rewrite < Hcut in \vdash (? % ?).
+ apply le_min_aux
+| apply plus_to_minus.
+ auto
+ (*apply plus_minus_m_m.
+ apply le_S_S.
+ apply le_n_fact_n*)
+]
+qed.
+
+variant lt_nth_prime_n_nth_prime_Sn :\forall n:nat.
+(nth_prime n) < (nth_prime (S n)) \def increasing_nth_prime.
+
+theorem injective_nth_prime: injective nat nat nth_prime.
+auto.
+(*apply increasing_to_injective.
+apply increasing_nth_prime.*)
+qed.
+
+theorem lt_SO_nth_prime_n : \forall n:nat. (S O) \lt nth_prime n.
+intros.
+(*usando la tattica auto qui, dopo svariati minuti la computazione non era
+ * ancora terminata
+ *)
+elim n
+[ auto
+ (*unfold lt.
+ apply le_n*)
+| auto
+ (*apply (trans_lt ? (nth_prime n1))
+ [ assumption
+ | apply lt_nth_prime_n_nth_prime_Sn
+ ]*)
+]
+qed.
+
+theorem lt_O_nth_prime_n : \forall n:nat. O \lt nth_prime n.
+intros.
+auto.
+(*apply (trans_lt O (S O))
+[ unfold lt.
+ apply le_n
+| apply lt_SO_nth_prime_n
+]*)
+qed.
+
+theorem ex_m_le_n_nth_prime_m:
+\forall n: nat. nth_prime O \le n \to
+\exists m. nth_prime m \le n \land n < nth_prime (S m).
+auto.
+(*intros.
+apply increasing_to_le2
+[ exact lt_nth_prime_n_nth_prime_Sn
+| assumption
+]*)
+qed.
+
+theorem lt_nth_prime_to_not_prime: \forall n,m. nth_prime n < m \to m < nth_prime (S n)
+\to \lnot (prime m).
+intros.
+apply primeb_false_to_not_prime.
+letin previous_prime \def (nth_prime n).
+letin upper_bound \def (S previous_prime!).
+apply (lt_min_aux_to_false primeb upper_bound (upper_bound - (S previous_prime)) m)
+[ cut (S (nth_prime n)!-(S (nth_prime n)! - (S (nth_prime n))) = (S (nth_prime n)))
+ [ rewrite > Hcut.
+ assumption
+ | apply plus_to_minus.
+ auto
+ (*apply plus_minus_m_m.
+ apply le_S_S.
+ apply le_n_fact_n*)
+ ]
+| assumption
+]
+qed.
+
+(* nth_prime enumerates all primes *)
+theorem prime_to_nth_prime : \forall p:nat. prime p \to
+\exists i. nth_prime i = p.
+intros.
+cut (\exists m. nth_prime m \le p \land p < nth_prime (S m))
+[ elim Hcut.
+ elim H1.
+ cut (nth_prime a < p \lor nth_prime a = p)
+ [ elim Hcut1
+ [ absurd (prime p)
+ [ assumption
+ | auto
+ (*apply (lt_nth_prime_to_not_prime a);assumption*)
+ ]
+ | auto
+ (*apply (ex_intro nat ? a).
+ assumption*)
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ assumption*)
+ ]
+| apply ex_m_le_n_nth_prime_m.
+ simplify.
+ unfold prime in H.
+ elim H.
+ assumption
+]
+qed.
+
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / Matita is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/ord".
+
+include "datatypes/constructors.ma".
+include "nat/exp.ma".
+include "nat/gcd.ma".
+include "nat/relevant_equations.ma". (* required by auto paramod *)
+
+(* this definition of log is based on pairs, with a remainder *)
+
+let rec p_ord_aux p n m \def
+ match n \mod m with
+ [ O \Rightarrow
+ match p with
+ [ O \Rightarrow pair nat nat O n
+ | (S p) \Rightarrow
+ match (p_ord_aux p (n / m) m) with
+ [ (pair q r) \Rightarrow pair nat nat (S q) r] ]
+ | (S a) \Rightarrow pair nat nat O n].
+
+(* p_ord n m = <q,r> if m divides n q times, with remainder r *)
+definition p_ord \def \lambda n,m:nat.p_ord_aux n n m.
+
+theorem p_ord_aux_to_Prop: \forall p,n,m. O < m \to
+ match p_ord_aux p n m with
+ [ (pair q r) \Rightarrow n = m \sup q *r ].
+intro.
+elim p
+[ simplify.
+ apply (nat_case (n \mod m))
+ [ simplify.
+ apply plus_n_O
+ | intros.
+ simplify.
+ apply plus_n_O
+ ]
+| simplify.
+ apply (nat_case1 (n1 \mod m))
+ [ intro.
+ simplify.
+ generalize in match (H (n1 / m) m).
+ elim (p_ord_aux n (n1 / m) m).
+ simplify.
+ rewrite > assoc_times.
+ rewrite < H3
+ [ rewrite > (plus_n_O (m*(n1 / m))).
+ rewrite < H2.
+ rewrite > sym_times.
+ auto
+ (*rewrite < div_mod
+ [ reflexivity
+ | assumption
+ ]*)
+ | assumption
+ ]
+ | intros.
+ simplify.
+ apply plus_n_O
+ ]
+]
+qed.
+
+theorem p_ord_aux_to_exp: \forall p,n,m,q,r. O < m \to
+ (pair nat nat q r) = p_ord_aux p n m \to n = m \sup q * r.
+intros.
+change with
+match (pair nat nat q r) with
+ [ (pair q r) \Rightarrow n = m \sup q * r ].
+rewrite > H1.
+apply p_ord_aux_to_Prop.
+assumption.
+qed.
+
+(* questo va spostato in primes1.ma *)
+theorem p_ord_exp: \forall n,m,i. O < m \to n \mod m \neq O \to
+\forall p. i \le p \to p_ord_aux p (m \sup i * n) m = pair nat nat i n.
+intros 5.
+elim i
+[ simplify.
+ rewrite < plus_n_O.
+ apply (nat_case p)
+ [ simplify.
+ elim (n \mod m);auto
+ (*[ simplify.
+ reflexivity
+ | simplify.
+ reflexivity
+ ]*)
+ | intro.
+ simplify.
+ cut (O < n \mod m \lor O = n \mod m)
+ [ elim Hcut
+ [ apply (lt_O_n_elim (n \mod m) H3).
+ intros.auto
+ (*simplify.
+ reflexivity*)
+ | apply False_ind.auto
+ (*apply H1.
+ apply sym_eq.
+ assumption*)
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ apply le_O_n*)
+ ]
+ ]
+| generalize in match H3.
+ apply (nat_case p)
+ [ intro.
+ apply False_ind.
+ apply (not_le_Sn_O n1 H4)
+ | intros.
+ simplify.
+ fold simplify (m \sup (S n1)).
+ cut (((m \sup (S n1)*n) \mod m) = O)
+ [ rewrite > Hcut.
+ simplify.
+ fold simplify (m \sup (S n1)).
+ cut ((m \sup (S n1) *n) / m = m \sup n1 *n)
+ [ rewrite > Hcut1.
+ rewrite > (H2 m1);auto
+ (*[ simplify.
+ reflexivity
+ | apply le_S_S_to_le.
+ assumption
+ ]*)
+ | (* div_exp *)
+ simplify.
+ rewrite > assoc_times.
+ apply (lt_O_n_elim m H).
+ intro.
+ apply div_times
+ ]
+ | (* mod_exp = O *)
+ apply divides_to_mod_O
+ [ assumption
+ | simplify.auto
+ (*rewrite > assoc_times.
+ apply (witness ? ? (m \sup n1 *n)).
+ reflexivity*)
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem p_ord_aux_to_Prop1: \forall p,n,m. (S O) < m \to O < n \to n \le p \to
+ match p_ord_aux p n m with
+ [ (pair q r) \Rightarrow r \mod m \neq O].
+intro.
+elim p
+[ absurd (O < n);auto
+ (*[ assumption
+ | apply le_to_not_lt.
+ assumption
+ ]*)
+| simplify.
+ apply (nat_case1 (n1 \mod m))
+ [ intro.
+ generalize in match (H (n1 / m) m).
+ elim (p_ord_aux n (n1 / m) m).
+ apply H5
+ [ assumption
+ | auto
+ (*apply eq_mod_O_to_lt_O_div
+ [ apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]
+ | assumption
+ | assumption
+ ]*)
+ | apply le_S_S_to_le.auto
+ (*apply (trans_le ? n1)
+ [ change with (n1 / m < n1).
+ apply lt_div_n_m_n;assumption
+ | assumption
+ ]*)
+ ]
+ | intros.
+ simplify.auto
+ (*rewrite > H4.
+ unfold Not.
+ intro.
+ apply (not_eq_O_S m1).
+ rewrite > H5.
+ reflexivity.*)
+ ]
+]
+qed.
+
+theorem p_ord_aux_to_not_mod_O: \forall p,n,m,q,r. (S O) < m \to O < n \to n \le p \to
+ pair nat nat q r = p_ord_aux p n m \to r \mod m \neq O.
+intros.
+change with
+ match (pair nat nat q r) with
+ [ (pair q r) \Rightarrow r \mod m \neq O].
+rewrite > H3.
+apply p_ord_aux_to_Prop1;
+ assumption.
+qed.
+
+theorem p_ord_exp1: \forall p,n,q,r. O < p \to \lnot p \divides r \to
+n = p \sup q * r \to p_ord n p = pair nat nat q r.
+intros.
+unfold p_ord.
+rewrite > H2.
+apply p_ord_exp
+[ assumption
+| unfold.
+ intro.
+ auto
+ (*apply H1.
+ apply mod_O_to_divides
+ [ assumption
+ | assumption
+ ]*)
+| apply (trans_le ? (p \sup q))
+ [ cut ((S O) \lt p)
+ [ auto
+ (*elim q
+ [ simplify.
+ apply le_n_Sn
+ | simplify.
+ generalize in match H3.
+ apply (nat_case n1)
+ [ simplify.
+ rewrite < times_n_SO.
+ intro.
+ assumption
+ | intros.
+ apply (trans_le ? (p*(S m)))
+ [ apply (trans_le ? ((S (S O))*(S m)))
+ [ simplify.
+ rewrite > plus_n_Sm.
+ rewrite < plus_n_O.
+ apply le_plus_n
+ | apply le_times_l.
+ assumption
+ ]
+ | apply le_times_r.
+ assumption
+ ]
+ ]
+ ]*)
+ | alias id "not_eq_to_le_to_lt" = "cic:/matita/algebra/finite_groups/not_eq_to_le_to_lt.con".
+ apply not_eq_to_le_to_lt
+ [ unfold.
+ intro.
+ auto
+ (*apply H1.
+ rewrite < H3.
+ apply (witness ? r r ?).
+ simplify.
+ apply plus_n_O*)
+ | assumption
+ ]
+ ]
+ | rewrite > times_n_SO in \vdash (? % ?).
+ apply le_times_r.
+ change with (O \lt r).
+ apply not_eq_to_le_to_lt
+ [ unfold.
+ intro.auto
+ (*apply H1.rewrite < H3.
+ apply (witness ? ? O ?).rewrite < times_n_O.
+ reflexivity*)
+ | apply le_O_n
+ ]
+ ]
+]
+qed.
+
+theorem p_ord_to_exp1: \forall p,n,q,r. (S O) \lt p \to O \lt n \to p_ord n p = pair nat nat q r\to
+\lnot p \divides r \land n = p \sup q * r.
+intros.
+unfold p_ord in H2.
+split
+[ unfold.intro.
+ apply (p_ord_aux_to_not_mod_O n n p q r);auto
+ (*[ assumption
+ | assumption
+ | apply le_n
+ | symmetry.
+ assumption
+ | apply divides_to_mod_O
+ [ apply (trans_lt ? (S O))
+ [ unfold.
+ apply le_n
+ | assumption
+ ]
+ | assumption
+ ]
+ ]*)
+| apply (p_ord_aux_to_exp n);auto
+ (*[ apply (trans_lt ? (S O))
+ [ unfold.
+ apply le_n
+ | assumption
+ ]
+ | symmetry.
+ assumption
+ ]*)
+]
+qed.
+
+theorem p_ord_times: \forall p,a,b,qa,ra,qb,rb. prime p
+\to O \lt a \to O \lt b
+\to p_ord a p = pair nat nat qa ra
+\to p_ord b p = pair nat nat qb rb
+\to p_ord (a*b) p = pair nat nat (qa + qb) (ra*rb).
+intros.
+cut ((S O) \lt p)
+[ elim (p_ord_to_exp1 ? ? ? ? Hcut H1 H3).
+ elim (p_ord_to_exp1 ? ? ? ? Hcut H2 H4).
+ apply p_ord_exp1
+ [ auto
+ (*apply (trans_lt ? (S O))
+ [ unfold.
+ apply le_n
+ | assumption
+ ]*)
+ | unfold.
+ intro.
+ elim (divides_times_to_divides ? ? ? H H9);auto
+ (*[ apply (absurd ? ? H10 H5)
+ | apply (absurd ? ? H10 H7)
+ ]*)
+ | (* rewrite > H6.
+ rewrite > H8. *)
+ auto paramodulation
+ ]
+| unfold prime in H.
+ elim H.
+ assumption
+]
+qed.
+
+theorem fst_p_ord_times: \forall p,a,b. prime p
+\to O \lt a \to O \lt b
+\to fst ? ? (p_ord (a*b) p) = (fst ? ? (p_ord a p)) + (fst ? ? (p_ord b p)).
+intros.
+rewrite > (p_ord_times p a b (fst ? ? (p_ord a p)) (snd ? ? (p_ord a p))
+(fst ? ? (p_ord b p)) (snd ? ? (p_ord b p)) H H1 H2);auto.
+(*[ simplify.
+ reflexivity
+| apply eq_pair_fst_snd
+| apply eq_pair_fst_snd
+]*)
+qed.
+
+theorem p_ord_p : \forall p:nat. (S O) \lt p \to p_ord p p = pair ? ? (S O) (S O).
+intros.
+apply p_ord_exp1
+[ auto
+ (*apply (trans_lt ? (S O))
+ [ unfold.
+ apply le_n
+ | assumption
+ ]*)
+| unfold.
+ intro.
+ apply (absurd ? ? H).auto
+ (*apply le_to_not_lt.
+ apply divides_to_le
+ [ unfold.
+ apply le_n
+ | assumption
+ ]*)
+| auto
+ (*rewrite < times_n_SO.
+ apply exp_n_SO*)
+]
+qed.
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/orders".
+
+include "nat/nat.ma".
+include "higher_order_defs/ordering.ma".
+
+(* definitions *)
+inductive le (n:nat) : nat \to Prop \def
+ | le_n : le n n
+ | le_S : \forall m:nat. le n m \to le n (S m).
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "natural 'less or equal to'" 'leq x y = (cic:/matita/nat/orders/le.ind#xpointer(1/1) x y).
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "natural 'neither less nor equal to'" 'nleq x y =
+ (cic:/matita/logic/connectives/Not.con
+ (cic:/matita/nat/orders/le.ind#xpointer(1/1) x y)).
+
+definition lt: nat \to nat \to Prop \def
+\lambda n,m:nat.(S n) \leq m.
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "natural 'less than'" 'lt x y = (cic:/matita/nat/orders/lt.con x y).
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "natural 'not less than'" 'nless x y =
+ (cic:/matita/logic/connectives/Not.con (cic:/matita/nat/orders/lt.con x y)).
+
+definition ge: nat \to nat \to Prop \def
+\lambda n,m:nat.m \leq n.
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "natural 'greater or equal to'" 'geq x y = (cic:/matita/nat/orders/ge.con x y).
+
+definition gt: nat \to nat \to Prop \def
+\lambda n,m:nat.m<n.
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "natural 'greater than'" 'gt x y = (cic:/matita/nat/orders/gt.con x y).
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "natural 'not greater than'" 'ngtr x y =
+ (cic:/matita/logic/connectives/Not.con (cic:/matita/nat/orders/gt.con x y)).
+
+theorem transitive_le : transitive nat le.
+unfold transitive.
+intros.
+elim H1;auto.
+(*[ assumption
+| apply le_S.
+ assumption
+]*)
+qed.
+
+theorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
+\def transitive_le.
+
+theorem transitive_lt: transitive nat lt.
+unfold transitive.
+unfold lt.
+intros.
+elim H1;auto.
+ (*apply le_S;assumption.*)
+
+qed.
+
+theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p
+\def transitive_lt.
+
+theorem le_S_S: \forall n,m:nat. n \leq m \to S n \leq S m.
+intros.
+elim H;auto.
+(*[ apply le_n.
+| apply le_S.
+ assumption
+]*)
+qed.
+
+theorem le_O_n : \forall n:nat. O \leq n.
+intros.
+elim n;auto.
+(*[ apply le_n
+| apply le_S.
+ assumption
+]*)
+qed.
+
+theorem le_n_Sn : \forall n:nat. n \leq S n.
+intros.auto.
+(*apply le_S.
+apply le_n.*)
+qed.
+
+theorem le_pred_n : \forall n:nat. pred n \leq n.
+intros.
+elim n;auto.
+(*[ simplify.
+ apply le_n
+| simplify.
+ apply le_n_Sn
+]*)
+qed.
+
+theorem le_S_S_to_le : \forall n,m:nat. S n \leq S m \to n \leq m.
+intros.
+change with (pred (S n) \leq pred (S m)).
+elim H;auto.
+(*[ apply le_n
+| apply (trans_le ? (pred n1))
+ [ assumption
+ | apply le_pred_n
+ ]
+]*)
+qed.
+
+theorem leS_to_not_zero : \forall n,m:nat. S n \leq m \to not_zero m.
+intros.
+elim H
+[ (*qui auto non chiude il goal*)
+ exact I
+| (*qui auto non chiude il goal*)
+ exact I
+]
+qed.
+
+(* not le *)
+theorem not_le_Sn_O: \forall n:nat. S n \nleq O.
+intros.
+unfold Not.
+simplify.
+intros.
+(*qui auto NON chiude il goal*)
+apply (leS_to_not_zero ? ? H).
+qed.
+
+theorem not_le_Sn_n: \forall n:nat. S n \nleq n.
+intros.
+elim n
+[ apply not_le_Sn_O
+| unfold Not.
+ simplify.
+ intros.auto
+ (*cut (S n1 \leq n1).
+ [ apply H.
+ assumption
+ | apply le_S_S_to_le.
+ assumption
+ ]*)
+]
+qed.
+
+(* le to lt or eq *)
+theorem le_to_or_lt_eq : \forall n,m:nat.
+n \leq m \to n < m \lor n = m.
+intros.
+elim H
+[ auto
+ (*right.
+ reflexivity*)
+| left.
+ unfold lt.
+ auto
+ (*apply le_S_S.
+ assumption*)
+]
+qed.
+
+(* not eq *)
+theorem lt_to_not_eq : \forall n,m:nat. n<m \to n \neq m.
+unfold Not.
+intros.
+cut ((le (S n) m) \to False)
+[ auto
+ (*apply Hcut.
+ assumption*)
+| rewrite < H1.
+ apply not_le_Sn_n
+]
+qed.
+
+(* le vs. lt *)
+theorem lt_to_le : \forall n,m:nat. n<m \to n \leq m.
+simplify.
+intros.
+auto.
+(*unfold lt in H.
+elim H
+[ apply le_S.
+ apply le_n
+| apply le_S.
+ assumption
+]*)
+qed.
+
+theorem lt_S_to_le : \forall n,m:nat. n < S m \to n \leq m.
+auto.
+(*simplify.
+intros.
+apply le_S_S_to_le.
+assumption.*)
+qed.
+
+theorem not_le_to_lt: \forall n,m:nat. n \nleq m \to m<n.
+intros 2.
+apply (nat_elim2 (\lambda n,m.n \nleq m \to m<n))
+[ intros.
+ apply (absurd (O \leq n1));auto
+ (*[ apply le_O_n
+ | assumption
+ ]*)
+| unfold Not.
+ unfold lt.
+ intros.auto
+ (*apply le_S_S.
+ apply le_O_n*)
+| unfold Not.
+ unfold lt.
+ intros.
+ apply le_S_S.
+ apply H.
+ intros.
+ auto
+ (*apply H1.
+ apply le_S_S.
+ assumption*)
+]
+qed.
+
+theorem lt_to_not_le: \forall n,m:nat. n<m \to m \nleq n.
+unfold Not.
+unfold lt.
+intros 3.
+elim H;auto.
+(*[ apply (not_le_Sn_n n H1)
+| apply H2.
+ apply lt_to_le.
+ apply H3
+]*)
+qed.
+
+theorem not_lt_to_le: \forall n,m:nat. Not (lt n m) \to le m n.
+simplify.
+intros.
+apply lt_S_to_le.
+apply not_le_to_lt.
+(*qui auto non chiude il goal*)
+exact H.
+qed.
+
+theorem le_to_not_lt: \forall n,m:nat. le n m \to Not (lt m n).
+intros.
+unfold Not.
+unfold lt.
+apply lt_to_not_le.
+unfold lt.auto.
+(*apply le_S_S.
+assumption.*)
+qed.
+
+(* le elimination *)
+theorem le_n_O_to_eq : \forall n:nat. n \leq O \to O=n.
+intro.
+elim n
+[ reflexivity
+| apply False_ind.auto
+ (*apply not_le_Sn_O.
+ goal 17. apply H1.*)
+]
+qed.
+
+theorem le_n_O_elim: \forall n:nat.n \leq O \to \forall P: nat \to Prop.
+P O \to P n.
+intro.
+elim n
+[ assumption
+| apply False_ind.
+ apply (not_le_Sn_O ? H1)
+]
+qed.
+
+theorem le_n_Sm_elim : \forall n,m:nat.n \leq S m \to
+\forall P:Prop. (S n \leq S m \to P) \to (n=S m \to P) \to P.
+intros 4.
+elim H;auto.
+(*[ apply H2.
+ reflexivity
+| apply H3.
+ apply le_S_S.
+ assumption
+]*)
+qed.
+
+(* le to eq *)
+lemma le_to_le_to_eq: \forall n,m. n \le m \to m \le n \to n = m.
+apply nat_elim2
+[ intros.
+ auto
+ (*apply le_n_O_to_eq.
+ assumption*)
+| intros.auto
+ (*apply sym_eq.
+ apply le_n_O_to_eq.
+ assumption*)
+| intros.
+ apply eq_f.auto
+ (*apply H
+ [ apply le_S_S_to_le.assumption
+ | apply le_S_S_to_le.assumption
+ ]*)
+]
+qed.
+
+(* lt and le trans *)
+theorem lt_O_S : \forall n:nat. O < S n.
+intro.auto.
+(*unfold.
+apply le_S_S.
+apply le_O_n.*)
+qed.
+
+theorem lt_to_le_to_lt: \forall n,m,p:nat. lt n m \to le m p \to lt n p.
+intros.
+elim H1;auto.
+(*[ assumption
+| unfold lt.
+ apply le_S.
+ assumption
+]*)
+qed.
+
+theorem le_to_lt_to_lt: \forall n,m,p:nat. le n m \to lt m p \to lt n p.
+intros 4.
+elim H
+[ assumption
+| apply H2.auto
+ (*unfold lt.
+ apply lt_to_le.
+ assumption*)
+]
+qed.
+
+theorem lt_S_to_lt: \forall n,m. S n < m \to n < m.
+intros.auto.
+(*apply (trans_lt ? (S n))
+[ apply le_n
+| assumption
+]*)
+qed.
+
+theorem ltn_to_ltO: \forall n,m:nat. lt n m \to lt O m.
+intros.
+apply (le_to_lt_to_lt O n)
+[ apply le_O_n
+| assumption
+]
+qed.
+
+theorem lt_O_n_elim: \forall n:nat. lt O n \to
+\forall P:nat\to Prop. (\forall m:nat.P (S m)) \to P n.
+intro.
+elim n
+[ apply False_ind.
+ exact (not_le_Sn_O O H)
+| apply H2
+]
+qed.
+
+(* other abstract properties *)
+theorem antisymmetric_le : antisymmetric nat le.
+unfold antisymmetric.
+auto.
+(*intros 2.
+apply (nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m)))
+[ intros.
+ apply le_n_O_to_eq.
+ assumption
+| intros.
+ apply False_ind.
+ apply (not_le_Sn_O ? H)
+| intros.
+ apply eq_f.
+ apply H;
+ apply le_S_S_to_le;assumption
+]*)
+qed.
+
+(*NOTA:
+ * auto chiude il goal prima della tattica intros 2, tuttavia non chiude gli ultimi
+ * 2 rami aperti dalla tattica apply (nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m))).
+ * evidentemente auto sfrutta una dimostrazione diversa rispetto a quella proposta
+ * nella libreria
+ *)
+
+theorem antisym_le: \forall n,m:nat. n \leq m \to m \leq n \to n=m
+\def antisymmetric_le.
+
+theorem decidable_le: \forall n,m:nat. decidable (n \leq m).
+intros.
+apply (nat_elim2 (\lambda n,m.decidable (n \leq m)))
+[ intros.
+ unfold decidable.auto
+ (*left.
+ apply le_O_n*)
+| intros.
+ unfold decidable.
+ auto
+ (*right.
+ exact (not_le_Sn_O n1)*)
+| intros 2.
+ unfold decidable.
+ intro.
+ elim H
+ [ auto
+ (*left.
+ apply le_S_S.
+ assumption*)
+ | right.
+ unfold Not.
+ auto
+ (*intro.
+ apply H1.
+ apply le_S_S_to_le.
+ assumption*)
+ ]
+]
+qed.
+
+theorem decidable_lt: \forall n,m:nat. decidable (n < m).
+intros.
+(*qui auto non chiude il goal*)
+exact (decidable_le (S n) m).
+qed.
+
+(* well founded induction principles *)
+
+theorem nat_elim1 : \forall n:nat.\forall P:nat \to Prop.
+(\forall m.(\forall p. (p \lt m) \to P p) \to P m) \to P n.
+intros.
+cut (\forall q:nat. q \le n \to P q)
+[ auto
+ (*apply (Hcut n).
+ apply le_n*)
+| elim n
+ [ apply (le_n_O_elim q H1).
+ apply H.
+ intros.
+ apply False_ind.
+ apply (not_le_Sn_O p H2)
+ | apply H.
+ intros.
+ apply H1.
+ auto
+ (*cut (p < S n1)
+ [ apply lt_S_to_le.
+ assumption
+ | apply (lt_to_le_to_lt p q (S n1) H3 H2).
+ ]*)
+ ]
+]
+qed.
+
+(* some properties of functions *)
+
+definition increasing \def \lambda f:nat \to nat.
+\forall n:nat. f n < f (S n).
+
+theorem increasing_to_monotonic: \forall f:nat \to nat.
+increasing f \to monotonic nat lt f.
+unfold monotonic.
+unfold lt.
+unfold increasing.
+unfold lt.
+intros.
+elim H1;auto.
+(*[ apply H
+| apply (trans_le ? (f n1))
+ [ assumption
+ | apply (trans_le ? (S (f n1)))
+ [ apply le_n_Sn
+ | apply H
+ ]
+ ]
+]*)
+qed.
+
+theorem le_n_fn: \forall f:nat \to nat. (increasing f)
+\to \forall n:nat. n \le (f n).
+intros.
+elim n;auto.
+(*[ apply le_O_n
+| apply (trans_le ? (S (f n1)))
+ [ apply le_S_S.
+ apply H1
+ | simplify in H.
+ unfold increasing in H.
+ unfold lt in H.
+ apply H
+ ]
+]*)
+qed.
+
+theorem increasing_to_le: \forall f:nat \to nat. (increasing f)
+\to \forall m:nat. \exists i. m \le (f i).
+intros.auto.
+(*elim m
+[ apply (ex_intro ? ? O).
+ apply le_O_n
+| elim H1.
+ apply (ex_intro ? ? (S a)).
+ apply (trans_le ? (S (f a)))
+ [ apply le_S_S.
+ assumption
+ | simplify in H.
+ unfold increasing in H.
+ unfold lt in H.
+ apply H
+ ]
+]*)
+qed.
+
+theorem increasing_to_le2: \forall f:nat \to nat. (increasing f)
+\to \forall m:nat. (f O) \le m \to
+\exists i. (f i) \le m \land m <(f (S i)).
+intros.
+elim H1
+[ auto.
+ (*apply (ex_intro ? ? O).
+ split
+ [ apply le_n
+ | apply H
+ ]*)
+| elim H3.
+ elim H4.
+ cut ((S n1) < (f (S a)) \lor (S n1) = (f (S a)))
+ [ elim Hcut
+ [ apply (ex_intro ? ? a).
+ auto
+ (*split
+ [ apply le_S.
+ assumption
+ | assumption
+ ]*)
+ | apply (ex_intro ? ? (S a)).
+ split
+ [ rewrite < H7.
+ apply le_n
+ | auto
+ (*rewrite > H7.
+ apply H*)
+ ]
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ apply H6*)
+ ]
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/permutation".
+
+include "nat/compare.ma".
+include "nat/sigma_and_pi.ma".
+
+definition injn: (nat \to nat) \to nat \to Prop \def
+\lambda f:nat \to nat.\lambda n:nat.\forall i,j:nat.
+i \le n \to j \le n \to f i = f j \to i = j.
+
+theorem injn_Sn_n: \forall f:nat \to nat. \forall n:nat.
+injn f (S n) \to injn f n.
+unfold injn.
+intros.
+apply H;auto.
+(*[ apply le_S.
+ assumption
+| apply le_S.
+ assumption
+| assumption
+]*)
+qed.
+
+theorem injective_to_injn: \forall f:nat \to nat. \forall n:nat.
+injective nat nat f \to injn f n.
+unfold injective.
+unfold injn.
+intros.auto.
+(*apply H.
+assumption.*)
+qed.
+
+definition permut : (nat \to nat) \to nat \to Prop
+\def \lambda f:nat \to nat. \lambda m:nat.
+(\forall i:nat. i \le m \to f i \le m )\land injn f m.
+
+theorem permut_O_to_eq_O: \forall h:nat \to nat.
+permut h O \to (h O) = O.
+intros.
+unfold permut in H.
+elim H.
+apply sym_eq.auto.
+(*apply le_n_O_to_eq.
+apply H1.
+apply le_n.*)
+qed.
+
+theorem permut_S_to_permut: \forall f:nat \to nat. \forall m:nat.
+permut f (S m) \to f (S m) = (S m) \to permut f m.
+unfold permut.
+intros.
+elim H.
+split
+[ intros.
+ cut (f i < S m \lor f i = S m)
+ [ elim Hcut
+ [ auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ | apply False_ind.
+ apply (not_le_Sn_n m).
+ cut ((S m) = i)
+ [ rewrite > Hcut1.
+ assumption
+ | apply H3
+ [ apply le_n
+ | auto
+ (*apply le_S.
+ assumption*)
+ | auto
+ (*rewrite > H5.
+ assumption*)
+ ]
+ ]
+ ]
+ | apply le_to_or_lt_eq.
+ auto
+ (*apply H2.
+ apply le_S.
+ assumption*)
+ ]
+| apply (injn_Sn_n f m H3)
+]
+qed.
+
+(* transpositions *)
+
+definition transpose : nat \to nat \to nat \to nat \def
+\lambda i,j,n:nat.
+match eqb n i with
+ [ true \Rightarrow j
+ | false \Rightarrow
+ match eqb n j with
+ [ true \Rightarrow i
+ | false \Rightarrow n]].
+
+notation < "(❲i↹j❳)n"
+ right associative with precedence 71
+for @{ 'transposition $i $j $n}.
+
+notation < "(❲i \atop j❳)n"
+ right associative with precedence 71
+for @{ 'transposition $i $j $n}.
+
+interpretation "natural transposition" 'transposition i j n =
+ (cic:/matita/nat/permutation/transpose.con i j n).
+
+lemma transpose_i_j_i: \forall i,j:nat. transpose i j i = j.
+intros.
+unfold transpose.
+(*dopo circa 6 minuti, l'esecuzione di auto in questo punto non era ancora terminata*)
+rewrite > (eqb_n_n i).auto.
+(*simplify.
+reflexivity.*)
+qed.
+
+lemma transpose_i_j_j: \forall i,j:nat. transpose i j j = i.
+intros.
+unfold transpose.
+apply (eqb_elim j i)
+[ auto
+ (*simplify.
+ intro.
+ assumption*)
+| rewrite > (eqb_n_n j).
+ simplify.
+ intros.
+ reflexivity
+]
+qed.
+
+theorem transpose_i_i: \forall i,n:nat. (transpose i i n) = n.
+intros.
+unfold transpose.
+apply (eqb_elim n i)
+[ auto
+ (*intro.
+ simplify.
+ apply sym_eq.
+ assumption*)
+| intro.
+ auto
+ (*simplify.
+ reflexivity*)
+]
+qed.
+
+theorem transpose_i_j_j_i: \forall i,j,n:nat.
+transpose i j n = transpose j i n.
+intros.
+unfold transpose.
+apply (eqb_elim n i)
+[ apply (eqb_elim n j)
+ [ intros.
+ (*l'esecuzione di auto in questo punto, dopo circa 300 secondi, non era ancora terminata*)
+ simplify.auto
+ (*rewrite < H.
+ rewrite < H1.
+ reflexivity*)
+ | intros.
+ auto
+ (*simplify.
+ reflexivity*)
+ ]
+| apply (eqb_elim n j)
+ [ intros.auto
+ (*simplify.reflexivity *)
+ | intros.auto
+ (*simplify.reflexivity*)
+ ]
+]
+qed.
+
+theorem transpose_transpose: \forall i,j,n:nat.
+(transpose i j (transpose i j n)) = n.
+intros.
+unfold transpose.
+unfold transpose.
+apply (eqb_elim n i)
+[ simplify.
+ intro.
+ apply (eqb_elim j i)
+ [ simplify.
+ intros.
+ auto
+ (*rewrite > H.
+ rewrite > H1.
+ reflexivity*)
+ | rewrite > (eqb_n_n j).
+ simplify.
+ intros.
+ auto
+ (*apply sym_eq.
+ assumption*)
+ ]
+| apply (eqb_elim n j)
+ [ simplify.
+ rewrite > (eqb_n_n i).
+ intros.
+ auto
+ (*simplify.
+ apply sym_eq.
+ assumption*)
+ | simplify.
+ intros.
+ (*l'esecuzione di auto in questo punto, dopo piu' di 6 minuti non era ancora terminata*)
+ rewrite > (not_eq_to_eqb_false n i H1).
+ (*l'esecuzione di auto in questo punto, dopo piu' alcuni minuti non era ancora terminata*)
+ rewrite > (not_eq_to_eqb_false n j H).auto
+ (*simplify.
+ reflexivity*)
+ ]
+]
+qed.
+
+theorem injective_transpose : \forall i,j:nat.
+injective nat nat (transpose i j).
+unfold injective.
+intros.auto.
+(*rewrite < (transpose_transpose i j x).
+rewrite < (transpose_transpose i j y).
+apply eq_f.
+assumption.*)
+qed.
+
+variant inj_transpose: \forall i,j,n,m:nat.
+transpose i j n = transpose i j m \to n = m \def
+injective_transpose.
+
+theorem permut_transpose: \forall i,j,n:nat. i \le n \to j \le n \to
+permut (transpose i j) n.
+unfold permut.
+intros.
+split
+[ unfold transpose.
+ intros.
+ elim (eqb i1 i)
+ [ (*qui auto non chiude il goal*)
+ simplify.
+ assumption
+ | elim (eqb i1 j)
+ [ (*aui auto non chiude il goal*)
+ simplify.
+ assumption
+ | (*aui auto non chiude il goal*)
+ simplify.
+ assumption
+ ]
+ ]
+| auto
+ (*apply (injective_to_injn (transpose i j) n).
+ apply injective_transpose*)
+]
+qed.
+
+theorem permut_fg: \forall f,g:nat \to nat. \forall n:nat.
+permut f n \to permut g n \to permut (\lambda m.(f(g m))) n.
+unfold permut.
+intros.
+elim H.
+elim H1.
+split
+[ intros.
+ simplify.
+ auto
+ (*apply H2.
+ apply H4.
+ assumption*)
+| simplify.
+ intros.
+ apply H5
+ [ assumption
+ | assumption
+ | apply H3
+ [ auto
+ (*apply H4.
+ assumption*)
+ | auto
+ (*apply H4.
+ assumption*)
+ | assumption
+ ]
+ ]
+]
+qed.
+
+theorem permut_transpose_l:
+\forall f:nat \to nat. \forall m,i,j:nat.
+i \le m \to j \le m \to permut f m \to permut (\lambda n.transpose i j (f n)) m.
+intros.
+auto.
+(*apply (permut_fg (transpose i j) f m ? ?)
+[ apply permut_transpose;assumption
+| assumption
+]*)
+qed.
+
+theorem permut_transpose_r:
+\forall f:nat \to nat. \forall m,i,j:nat.
+i \le m \to j \le m \to permut f m \to permut (\lambda n.f (transpose i j n)) m.
+intros.auto.
+(*apply (permut_fg f (transpose i j) m ? ?)
+[ assumption
+| apply permut_transpose;assumption
+]*)
+qed.
+
+theorem eq_transpose : \forall i,j,k,n:nat. \lnot j=i \to
+ \lnot i=k \to \lnot j=k \to
+transpose i j n = transpose i k (transpose k j (transpose i k n)).
+(* uffa: triplo unfold? *)
+intros.unfold transpose.
+unfold transpose.
+unfold transpose.
+apply (eqb_elim n i)
+[ intro.
+ simplify.
+ rewrite > (eqb_n_n k).
+ simplify.
+ rewrite > (not_eq_to_eqb_false j i H).
+ rewrite > (not_eq_to_eqb_false j k H2).
+ reflexivity
+| intro.
+ apply (eqb_elim n j)
+ [ intro.
+ cut (\lnot n = k)
+ [ cut (\lnot n = i)
+ [ rewrite > (not_eq_to_eqb_false n k Hcut).
+ simplify.
+ rewrite > (not_eq_to_eqb_false n k Hcut).
+ rewrite > (eq_to_eqb_true n j H4).
+ simplify.
+ rewrite > (not_eq_to_eqb_false k i)
+ [ rewrite > (eqb_n_n k).
+ auto
+ (*simplify.
+ reflexivity*)
+ | unfold Not.
+ intro.auto
+ (*apply H1.
+ apply sym_eq.
+ assumption*)
+ ]
+ | assumption
+ ]
+ | unfold Not.
+ intro.auto
+ (*apply H2.
+ apply (trans_eq ? ? n)
+ [ apply sym_eq.
+ assumption
+ | assumption
+ ]*)
+ ]
+ | intro.
+ apply (eqb_elim n k)
+ [ intro.
+ simplify.
+ rewrite > (not_eq_to_eqb_false i k H1).
+ rewrite > (not_eq_to_eqb_false i j)
+ [ simplify.
+ rewrite > (eqb_n_n i).
+ auto
+ (*simplify.
+ assumption*)
+ | unfold Not.
+ intro.auto
+ (*apply H.
+ apply sym_eq.
+ assumption*)
+ ]
+ | intro.
+ simplify.
+ rewrite > (not_eq_to_eqb_false n k H5).
+ rewrite > (not_eq_to_eqb_false n j H4).
+ simplify.
+ rewrite > (not_eq_to_eqb_false n i H3).
+ rewrite > (not_eq_to_eqb_false n k H5).
+ auto
+ (*simplify.
+ reflexivity*)
+ ]
+ ]
+]
+qed.
+
+theorem permut_S_to_permut_transpose: \forall f:nat \to nat.
+\forall m:nat. permut f (S m) \to permut (\lambda n.transpose (f (S m)) (S m)
+(f n)) m.
+unfold permut.
+intros.
+elim H.
+split
+[ intros.
+ simplify.
+ unfold transpose.
+ apply (eqb_elim (f i) (f (S m)))
+ [ intro.
+ apply False_ind.
+ cut (i = (S m))
+ [ apply (not_le_Sn_n m).
+ rewrite < Hcut.
+ assumption
+ | apply H2;auto
+ (*[ apply le_S.
+ assumption
+ | apply le_n
+ | assumption
+ ]*)
+ ]
+ | intro.
+ simplify.
+ apply (eqb_elim (f i) (S m))
+ [ intro.
+ cut (f (S m) \lt (S m) \lor f (S m) = (S m))
+ [ elim Hcut
+ [ apply le_S_S_to_le.
+ (*NB qui auto non chiude il goal*)
+ assumption
+ | apply False_ind.
+ auto
+ (*apply H4.
+ rewrite > H6.
+ assumption*)
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ apply H1.
+ apply le_n*)
+ ]
+ | intro.simplify.
+ cut (f i \lt (S m) \lor f i = (S m))
+ [ elim Hcut
+ [ auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ | apply False_ind.
+ auto
+ (*apply H5.
+ assumption*)
+ ]
+ | apply le_to_or_lt_eq.
+ auto
+ (*apply H1.
+ apply le_S.
+ assumption*)
+ ]
+ ]
+ ]
+| unfold injn.
+ intros.
+ apply H2;auto
+ (*[ apply le_S.
+ assumption
+ | apply le_S.
+ assumption
+ | apply (inj_transpose (f (S m)) (S m)).
+ apply H5
+ ]*)
+]
+qed.
+
+(* bounded bijectivity *)
+
+definition bijn : (nat \to nat) \to nat \to Prop \def
+\lambda f:nat \to nat. \lambda n. \forall m:nat. m \le n \to
+ex nat (\lambda p. p \le n \land f p = m).
+
+theorem eq_to_bijn: \forall f,g:nat\to nat. \forall n:nat.
+(\forall i:nat. i \le n \to (f i) = (g i)) \to
+bijn f n \to bijn g n.
+intros 4.
+unfold bijn.
+intros.
+elim (H1 m)
+[ apply (ex_intro ? ? a).
+ rewrite < (H a)
+ [ assumption
+ | elim H3.
+ assumption
+ ]
+| assumption
+]
+qed.
+
+theorem bijn_Sn_n: \forall f:nat \to nat. \forall n:nat.
+bijn f (S n) \to f (S n) = (S n) \to bijn f n.
+unfold bijn.
+intros.
+elim (H m)
+[ elim H3.
+ apply (ex_intro ? ? a).
+ split
+ [ cut (a < S n \lor a = S n)
+ [ elim Hcut
+ [ auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ | apply False_ind.
+ apply (not_le_Sn_n n).
+ rewrite < H1.
+ rewrite < H6.
+ rewrite > H5.
+ assumption
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ assumption*)
+ ]
+ | assumption
+ ]
+| auto
+ (*apply le_S.
+ assumption*)
+]
+qed.
+
+theorem bijn_n_Sn: \forall f:nat \to nat. \forall n:nat.
+bijn f n \to f (S n) = (S n) \to bijn f (S n).
+unfold bijn.
+intros.
+cut (m < S n \lor m = S n)
+[ elim Hcut
+ [ elim (H m)
+ [ elim H4.
+ apply (ex_intro ? ? a).
+ auto
+ (*split
+ [ apply le_S.
+ assumption
+ | assumption
+ ]*)
+ | auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ ]
+ | auto
+ (*apply (ex_intro ? ? (S n)).
+ split
+ [ apply le_n
+ | rewrite > H3.
+ assumption
+ ]*)
+ ]
+| auto
+ (*apply le_to_or_lt_eq.
+ assumption*)
+]
+qed.
+
+theorem bijn_fg: \forall f,g:nat\to nat. \forall n:nat.
+bijn f n \to bijn g n \to bijn (\lambda p.f(g p)) n.
+unfold bijn.
+intros.
+simplify.
+elim (H m)
+[ elim H3.
+ elim (H1 a)
+ [ elim H6.
+ auto
+ (*apply (ex_intro ? ? a1).
+ split
+ [ assumption
+ | rewrite > H8.
+ assumption
+ ]*)
+ | assumption
+ ]
+| assumption
+]
+qed.
+
+theorem bijn_transpose : \forall n,i,j. i \le n \to j \le n \to
+bijn (transpose i j) n.
+intros.
+unfold bijn.
+unfold transpose.
+intros.
+cut (m = i \lor \lnot m = i)
+[ elim Hcut
+ [ apply (ex_intro ? ? j).
+ split
+ [ assumption
+ | apply (eqb_elim j i)
+ [ intro.
+ (*dopo circa 360 secondi l'esecuzione di auto in questo punto non era ancora terminata*)
+ simplify.
+ auto
+ (*rewrite > H3.
+ rewrite > H4.
+ reflexivity*)
+ | rewrite > (eqb_n_n j).
+ simplify.
+ intros.
+ auto
+ (*apply sym_eq.
+ assumption*)
+ ]
+ ]
+ | cut (m = j \lor \lnot m = j)
+ [ elim Hcut1
+ [ apply (ex_intro ? ? i).
+ split
+ [ assumption
+ | (*dopo circa 5 minuti, l'esecuzione di auto in questo punto non era ancora terminata*)
+ rewrite > (eqb_n_n i).
+ auto
+ (*simplify.
+ apply sym_eq.
+ assumption*)
+ ]
+ | apply (ex_intro ? ? m).
+ split
+ [ assumption
+ | rewrite > (not_eq_to_eqb_false m i)
+ [ (*dopo circa 5 minuti, l'esecuzione di auto in questo punto non era ancora terminata*)
+ rewrite > (not_eq_to_eqb_false m j)
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | assumption
+ ]
+ | assumption
+ ]
+ ]
+ ]
+ | apply (decidable_eq_nat m j)
+ ]
+ ]
+| apply (decidable_eq_nat m i)
+]
+qed.
+
+theorem bijn_transpose_r: \forall f:nat\to nat.\forall n,i,j. i \le n \to j \le n \to
+bijn f n \to bijn (\lambda p.f (transpose i j p)) n.
+intros.auto.
+(*apply (bijn_fg f ?)
+[ assumption
+| apply (bijn_transpose n i j)
+ [ assumption
+ | assumption
+ ]
+]*)
+qed.
+
+theorem bijn_transpose_l: \forall f:nat\to nat.\forall n,i,j. i \le n \to j \le n \to
+bijn f n \to bijn (\lambda p.transpose i j (f p)) n.
+intros.
+auto.
+(*apply (bijn_fg ? f)
+[ apply (bijn_transpose n i j)
+ [ assumption
+ | assumption
+ ]
+| assumption
+]*)
+qed.
+
+theorem permut_to_bijn: \forall n:nat.\forall f:nat\to nat.
+permut f n \to bijn f n.
+intro.
+elim n
+[ unfold bijn.
+ intros.
+ apply (ex_intro ? ? m).
+ split
+ [ assumption
+ | apply (le_n_O_elim m ? (\lambda p. f p = p))
+ [ assumption
+ | unfold permut in H.
+ elim H.
+ apply sym_eq.
+ auto
+ (*apply le_n_O_to_eq.
+ apply H2.
+ apply le_n*)
+ ]
+ ]
+| apply (eq_to_bijn (\lambda p.
+ (transpose (f (S n1)) (S n1)) (transpose (f (S n1)) (S n1) (f p))) f)
+ [ intros.
+ apply transpose_transpose
+ | apply (bijn_fg (transpose (f (S n1)) (S n1)))
+ [ apply bijn_transpose
+ [ unfold permut in H1.
+ elim H1.auto
+ (*apply H2.
+ apply le_n*)
+ | apply le_n
+ ]
+ | apply bijn_n_Sn
+ [ apply H.
+ auto
+ (*apply permut_S_to_permut_transpose.
+ assumption*)
+ | auto
+ (*unfold transpose.
+ rewrite > (eqb_n_n (f (S n1))).
+ simplify.
+ reflexivity*)
+ ]
+ ]
+ ]
+]
+qed.
+
+let rec invert_permut n f m \def
+ match eqb m (f n) with
+ [true \Rightarrow n
+ |false \Rightarrow
+ match n with
+ [O \Rightarrow O
+ |(S p) \Rightarrow invert_permut p f m]].
+
+theorem invert_permut_f: \forall f:nat \to nat. \forall n,m:nat.
+m \le n \to injn f n\to invert_permut n f (f m) = m.
+intros 4.
+elim H
+[ apply (nat_case1 m)
+ [ intro.
+ simplify.
+ (*l'applicazione di auto in questo punto, dopo alcuni minuti, non aveva ancora dato risultati*)
+ rewrite > (eqb_n_n (f O)).
+ auto
+ (*simplify.
+ reflexivity*)
+ | intros.simplify.
+ (*l'applicazione di auto in questo punto, dopo alcuni minuti, non aveva ancora dato risultati*)
+ rewrite > (eqb_n_n (f (S m1))).
+ auto
+ (*simplify.
+ reflexivity*)
+ ]
+| simplify.
+ rewrite > (not_eq_to_eqb_false (f m) (f (S n1)))
+ [ (*l'applicazione di auto in questo punto, dopo parecchi secondi, non aveva ancora prodotto un risultato*)
+ simplify.
+ auto
+ (*apply H2.
+ apply injn_Sn_n.
+ assumption*)
+ | unfold Not.
+ intro.
+ absurd (m = S n1)
+ [ apply H3;auto
+ (*[ apply le_S.
+ assumption
+ | apply le_n
+ | assumption
+ ]*)
+ | unfold Not.
+ intro.
+ apply (not_le_Sn_n n1).
+ rewrite < H5.
+ assumption
+ ]
+ ]
+]
+qed.
+
+theorem injective_invert_permut: \forall f:nat \to nat. \forall n:nat.
+permut f n \to injn (invert_permut n f) n.
+intros.
+unfold injn.
+intros.
+cut (bijn f n)
+[ unfold bijn in Hcut.
+ generalize in match (Hcut i H1).
+ intro.
+ generalize in match (Hcut j H2).
+ intro.
+ elim H4.
+ elim H6.
+ elim H5.
+ elim H9.
+ rewrite < H8.
+ rewrite < H11.
+ apply eq_f.
+ rewrite < (invert_permut_f f n a)
+ [ rewrite < (invert_permut_f f n a1)
+ [ rewrite > H8.
+ rewrite > H11.
+ assumption
+ | assumption
+ | unfold permut in H.elim H.
+ assumption
+ ]
+ | assumption
+ | unfold permut in H.
+ elim H.
+ assumption
+ ]
+| auto
+ (*apply permut_to_bijn.
+ assumption*)
+]
+qed.
+
+theorem permut_invert_permut: \forall f:nat \to nat. \forall n:nat.
+permut f n \to permut (invert_permut n f) n.
+intros.
+unfold permut.
+split
+[ intros.
+ simplify.
+ elim n
+ [ simplify.
+ elim (eqb i (f O));auto
+ (*[ simplify.
+ apply le_n
+ | simplify.
+ apply le_n
+ ]*)
+ | simplify.
+ elim (eqb i (f (S n1)))
+ [ auto
+ (*simplify.
+ apply le_n*)
+ | simplify.
+ auto
+ (*apply le_S.
+ assumption*)
+ ]
+ ]
+| auto
+ (*apply injective_invert_permut.
+ assumption.*)
+]
+qed.
+
+theorem f_invert_permut: \forall f:nat \to nat. \forall n,m:nat.
+m \le n \to permut f n\to f (invert_permut n f m) = m.
+intros.
+apply (injective_invert_permut f n H1)
+[ unfold permut in H1.
+ elim H1.
+ apply H2.
+ cut (permut (invert_permut n f) n)
+ [ unfold permut in Hcut.
+ elim Hcut.auto
+ (*apply H4.
+ assumption*)
+ | apply permut_invert_permut.
+ (*NB qui auto non chiude il goal*)
+ assumption
+ ]
+| assumption
+| apply invert_permut_f
+ [ cut (permut (invert_permut n f) n)
+ [ unfold permut in Hcut.
+ elim Hcut.
+ auto
+ (*apply H2.
+ assumption*)
+ | auto
+ (*apply permut_invert_permut.
+ assumption*)
+ ]
+ | unfold permut in H1.
+ elim H1.
+ assumption
+ ]
+]
+qed.
+
+theorem permut_n_to_eq_n: \forall h:nat \to nat.\forall n:nat.
+permut h n \to (\forall m:nat. m < n \to h m = m) \to h n = n.
+intros.
+unfold permut in H.
+elim H.
+cut (invert_permut n h n < n \lor invert_permut n h n = n)
+[ elim Hcut
+ [ rewrite < (f_invert_permut h n n) in \vdash (? ? ? %)
+ [ apply eq_f.
+ rewrite < (f_invert_permut h n n) in \vdash (? ? % ?)
+ [ auto
+ (*apply H1.
+ assumption*)
+ | apply le_n
+ | (*qui auto NON chiude il goal*)
+ assumption
+ ]
+ | apply le_n
+ | (*qui auto NON chiude il goal*)
+ assumption
+ ]
+ | rewrite < H4 in \vdash (? ? % ?).
+ apply (f_invert_permut h)
+ [ apply le_n
+ | (*qui auto NON chiude il goal*)
+ assumption
+ ]
+ ]
+| apply le_to_or_lt_eq.
+ cut (permut (invert_permut n h) n)
+ [ unfold permut in Hcut.
+ elim Hcut.
+ auto
+ (*apply H4.
+ apply le_n*)
+ | apply permut_invert_permut.
+ (*NB aui auto non chiude il goal*)
+ assumption
+ ]
+]
+qed.
+
+theorem permut_n_to_le: \forall h:nat \to nat.\forall k,n:nat.
+k \le n \to permut h n \to (\forall m:nat. m < k \to h m = m) \to
+\forall j. k \le j \to j \le n \to k \le h j.
+intros.
+unfold permut in H1.
+elim H1.
+cut (h j < k \lor \not(h j < k))
+[ elim Hcut
+ [ absurd (k \le j)
+ [ assumption
+ | apply lt_to_not_le.
+ cut (h j = j)
+ [ rewrite < Hcut1.
+ assumption
+ | apply H6;auto
+ (*[ apply H5.
+ assumption
+ | assumption
+ | apply H2.
+ assumption
+ ]*)
+ ]
+ ]
+ | auto
+ (*apply not_lt_to_le.
+ assumption*)
+ ]
+| apply (decidable_lt (h j) k)
+]
+qed.
+
+(* applications *)
+
+let rec map_iter_i k (g:nat \to nat) f (i:nat) \def
+ match k with
+ [ O \Rightarrow g i
+ | (S k) \Rightarrow f (g (S (k+i))) (map_iter_i k g f i)].
+
+theorem eq_map_iter_i: \forall g1,g2:nat \to nat.
+\forall f:nat \to nat \to nat. \forall n,i:nat.
+(\forall m:nat. i\le m \to m \le n+i \to g1 m = g2 m) \to
+map_iter_i n g1 f i = map_iter_i n g2 f i.
+intros 5.
+elim n
+[ simplify.
+ auto
+ (*apply H
+ [ apply le_n
+ | apply le_n
+ ]*)
+| simplify.
+ apply eq_f2
+ [ auto
+ (*apply H1
+ [ simplify.
+ apply le_S.
+ apply le_plus_n
+ | simplify.
+ apply le_n
+ ]*)
+ | apply H.
+ intros.
+ apply H1;auto
+ (*[ assumption
+ | simplify.
+ apply le_S.
+ assumption
+ ]*)
+ ]
+]
+qed.
+
+(* map_iter examples *)
+
+theorem eq_map_iter_i_sigma: \forall g:nat \to nat. \forall n,m:nat.
+map_iter_i n g plus m = sigma n g m.
+intros.
+elim n
+[ auto
+ (*simplify.
+ reflexivity*)
+| simplify.
+ auto
+ (*apply eq_f.
+ assumption*)
+]
+qed.
+
+theorem eq_map_iter_i_pi: \forall g:nat \to nat. \forall n,m:nat.
+map_iter_i n g times m = pi n g m.
+intros.
+elim n
+[ auto
+ (*simplify.
+ reflexivity*)
+| simplify.
+ auto
+ (*apply eq_f.
+ assumption*)
+]
+qed.
+
+theorem eq_map_iter_i_fact: \forall n:nat.
+map_iter_i n (\lambda m.m) times (S O) = (S n)!.
+intros.
+elim n
+[ auto
+ (*simplify.
+ reflexivity*)
+| change with
+ (((S n1)+(S O))*(map_iter_i n1 (\lambda m.m) times (S O)) = (S(S n1))*(S n1)!).
+ rewrite < plus_n_Sm.
+ rewrite < plus_n_O.
+ apply eq_f.
+ (*NB: qui auto non chiude il goal!!!*)
+ assumption
+]
+qed.
+
+
+theorem eq_map_iter_i_transpose_l : \forall f:nat\to nat \to nat.associative nat f \to
+symmetric2 nat nat f \to \forall g:nat \to nat. \forall n,k:nat.
+map_iter_i (S k) g f n = map_iter_i (S k) (\lambda m. g (transpose (k+n) (S k+n) m)) f n.
+intros.
+apply (nat_case1 k)
+[ intros.
+ simplify.
+ fold simplify (transpose n (S n) (S n)).
+ auto
+ (*rewrite > transpose_i_j_i.
+ rewrite > transpose_i_j_j.
+ apply H1*)
+| intros.
+ change with
+ (f (g (S (S (m+n)))) (f (g (S (m+n))) (map_iter_i m g f n)) =
+ f (g (transpose (S m + n) (S (S m) + n) (S (S m)+n)))
+ (f (g (transpose (S m + n) (S (S m) + n) (S m+n)))
+ (map_iter_i m (\lambda m1. g (transpose (S m+n) (S (S m)+n) m1)) f n))).
+ rewrite > transpose_i_j_i.
+ rewrite > transpose_i_j_j.
+ rewrite < H.
+ rewrite < H.
+ rewrite < (H1 (g (S m + n))).
+ apply eq_f.
+ apply eq_map_iter_i.
+ intros.
+ simplify.
+ unfold transpose.
+ rewrite > (not_eq_to_eqb_false m1 (S m+n))
+ [ rewrite > (not_eq_to_eqb_false m1 (S (S m)+n))
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | apply (lt_to_not_eq m1 (S ((S m)+n))).
+ auto
+ (*unfold lt.
+ apply le_S_S.
+ change with (m1 \leq S (m+n)).
+ apply le_S.
+ assumption*)
+ ]
+ | apply (lt_to_not_eq m1 (S m+n)).
+ simplify.auto
+ (*unfold lt.
+ apply le_S_S.
+ assumption*)
+ ]
+]
+qed.
+
+theorem eq_map_iter_i_transpose_i_Si : \forall f:nat\to nat \to nat.associative nat f \to
+symmetric2 nat nat f \to \forall g:nat \to nat. \forall n,k,i:nat. n \le i \to i \le k+n \to
+map_iter_i (S k) g f n = map_iter_i (S k) (\lambda m. g (transpose i (S i) m)) f n.
+intros 6.
+elim k
+[ cut (i=n)
+ [ rewrite > Hcut.
+ (*qui auto non chiude il goal*)
+ apply (eq_map_iter_i_transpose_l f H H1 g n O)
+ | apply antisymmetric_le
+ [ assumption
+ | assumption
+ ]
+ ]
+| cut (i < S n1 + n \lor i = S n1 + n)
+ [ elim Hcut
+ [ change with
+ (f (g (S (S n1)+n)) (map_iter_i (S n1) g f n) =
+ f (g (transpose i (S i) (S (S n1)+n))) (map_iter_i (S n1) (\lambda m. g (transpose i (S i) m)) f n)).
+ apply eq_f2
+ [ unfold transpose.
+ rewrite > (not_eq_to_eqb_false (S (S n1)+n) i)
+ [ rewrite > (not_eq_to_eqb_false (S (S n1)+n) (S i))
+ [ auto
+ (*simplify.
+ reflexivity*)
+ | simplify.
+ unfold Not.
+ intro.
+ apply (lt_to_not_eq i (S n1+n))
+ [ assumption
+ | auto
+ (*apply inj_S.
+ apply sym_eq.
+ assumption*)
+ ]
+ ]
+ | simplify.
+ unfold Not.
+ intro.
+ apply (lt_to_not_eq i (S (S n1+n)))
+ [ auto
+ (*simplify.
+ unfold lt.
+ apply le_S_S.
+ assumption*)
+ | auto
+ (*apply sym_eq.
+ assumption*)
+ ]
+ ]
+ | apply H2;auto
+ (*[ assumption
+ | apply le_S_S_to_le.
+ assumption
+ ]*)
+ ]
+ | rewrite > H5.
+ (*qui auto non chiude il goal*)
+ apply (eq_map_iter_i_transpose_l f H H1 g n (S n1)).
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ assumption*)
+ ]
+]
+qed.
+
+theorem eq_map_iter_i_transpose:
+\forall f:nat\to nat \to nat.
+associative nat f \to symmetric2 nat nat f \to \forall n,k,o:nat.
+\forall g:nat \to nat. \forall i:nat. n \le i \to S (o + i) \le S k+n \to
+map_iter_i (S k) g f n = map_iter_i (S k) (\lambda m. g (transpose i (S(o + i)) m)) f n.
+intros 6.
+apply (nat_elim1 o).
+intro.
+apply (nat_case m ?)
+[ intros.
+ apply (eq_map_iter_i_transpose_i_Si ? H H1);auto
+ (*[ exact H3
+ | apply le_S_S_to_le.
+ assumption
+ ]*)
+| intros.
+ apply (trans_eq ? ? (map_iter_i (S k) (\lambda m. g (transpose i (S(m1 + i)) m)) f n))
+ [ apply H2
+ [ auto
+ (*unfold lt.
+ apply le_n*)
+ | assumption
+ | apply (trans_le ? (S(S (m1+i))))
+ [ auto
+ (*apply le_S.
+ apply le_n*)
+ | (*qui auto non chiude il goal, chiuso invece da assumption*)
+ assumption
+ ]
+ ]
+ | apply (trans_eq ? ? (map_iter_i (S k) (\lambda m. g
+ (transpose i (S(m1 + i)) (transpose (S(m1 + i)) (S(S(m1 + i))) m))) f n))
+ [ (*qui auto dopo alcuni minuti non aveva ancora terminato la propria esecuzione*)
+ apply (H2 O ? ? (S(m1+i)))
+ [ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | auto
+ (*apply (trans_le ? i)
+ [ assumption
+ | change with (i \le (S m1)+i).
+ apply le_plus_n
+ ]*)
+ | (*qui auto non chiude il goal*)
+ exact H4
+ ]
+ | apply (trans_eq ? ? (map_iter_i (S k) (\lambda m. g
+ (transpose i (S(m1 + i))
+ (transpose (S(m1 + i)) (S(S(m1 + i)))
+ (transpose i (S(m1 + i)) m)))) f n))
+ [ (*qui auto dopo alcuni minuti non aveva ancora terminato la propria esecuzione*)
+ apply (H2 m1)
+ [ auto
+ (*unfold lt.
+ apply le_n*)
+ | assumption
+ | apply (trans_le ? (S(S (m1+i))))
+ [ auto
+ (*apply le_S.
+ apply le_n*)
+ | (*qui auto NON CHIUDE il goal*)
+ assumption
+ ]
+ ]
+ | apply eq_map_iter_i.
+ intros.
+ apply eq_f.
+ apply sym_eq.
+ apply eq_transpose
+ [ unfold Not.
+ intro.
+ apply (not_le_Sn_n i).
+ rewrite < H7 in \vdash (? ? %).
+ auto
+ (*apply le_S_S.
+ apply le_S.
+ apply le_plus_n*)
+ | unfold Not.
+ intro.
+ apply (not_le_Sn_n i).
+ rewrite > H7 in \vdash (? ? %).
+ auto
+ (*apply le_S_S.
+ apply le_plus_n*)
+ | unfold Not.
+ intro.
+ auto
+ (*apply (not_eq_n_Sn (S m1+i)).
+ apply sym_eq.
+ assumption*)
+ ]
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem eq_map_iter_i_transpose1: \forall f:nat\to nat \to nat.associative nat f \to
+symmetric2 nat nat f \to \forall n,k,i,j:nat.
+\forall g:nat \to nat. n \le i \to i < j \to j \le S k+n \to
+map_iter_i (S k) g f n = map_iter_i (S k) (\lambda m. g (transpose i j m)) f n.
+intros.
+simplify in H3.
+cut ((S i) < j \lor (S i) = j)
+[ elim Hcut
+ [ cut (j = S ((j - (S i)) + i))
+ [ rewrite > Hcut1.
+ apply (eq_map_iter_i_transpose f H H1 n k (j - (S i)) g i)
+ [ assumption
+ | rewrite < Hcut1.
+ assumption
+ ]
+ | rewrite > plus_n_Sm.
+ auto
+ (*apply plus_minus_m_m.
+ apply lt_to_le.
+ assumption*)
+ ]
+ | rewrite < H5.
+ apply (eq_map_iter_i_transpose_i_Si f H H1 g)
+ [ auto
+ (*simplify.
+ assumption*)
+ | apply le_S_S_to_le.
+ auto
+ (*apply (trans_le ? j)
+ [ assumption
+ | assumption
+ ]*)
+ ]
+ ]
+| auto
+ (*apply le_to_or_lt_eq.
+ assumption*)
+]
+qed.
+
+theorem eq_map_iter_i_transpose2: \forall f:nat\to nat \to nat.associative nat f \to
+symmetric2 nat nat f \to \forall n,k,i,j:nat.
+\forall g:nat \to nat. n \le i \to i \le (S k+n) \to n \le j \to j \le (S k+n) \to
+map_iter_i (S k) g f n = map_iter_i (S k) (\lambda m. g (transpose i j m)) f n.
+intros.
+apply (nat_compare_elim i j)
+[ intro.
+ (*qui auto non chiude il goal*)
+ apply (eq_map_iter_i_transpose1 f H H1 n k i j g H2 H6 H5)
+| intro.
+ rewrite > H6.
+ apply eq_map_iter_i.
+ intros.
+ auto
+ (*rewrite > (transpose_i_i j).
+ reflexivity*)
+| intro.
+ apply (trans_eq ? ? (map_iter_i (S k) (\lambda m:nat.g (transpose j i m)) f n))
+ [ apply (eq_map_iter_i_transpose1 f H H1 n k j i g H4 H6 H3)
+ | apply eq_map_iter_i.
+ intros.
+ auto
+ (*apply eq_f.
+ apply transpose_i_j_j_i*)
+ ]
+]
+qed.
+
+theorem permut_to_eq_map_iter_i:\forall f:nat\to nat \to nat.associative nat f \to
+symmetric2 nat nat f \to \forall k,n:nat.\forall g,h:nat \to nat.
+permut h (k+n) \to (\forall m:nat. m \lt n \to h m = m) \to
+map_iter_i k g f n = map_iter_i k (\lambda m.g(h m)) f n.
+intros 4.
+elim k
+[ simplify.
+ rewrite > (permut_n_to_eq_n h)
+ [ reflexivity
+ | (*qui auto non chiude il goal*)
+ assumption
+ | (*qui auto non chiude il goal*)
+ assumption
+ ]
+| apply (trans_eq ? ? (map_iter_i (S n) (\lambda m.g ((transpose (h (S n+n1)) (S n+n1)) m)) f n1))
+ [ unfold permut in H3.
+ elim H3.
+ apply (eq_map_iter_i_transpose2 f H H1 n1 n ? ? g)
+ [ apply (permut_n_to_le h n1 (S n+n1))
+ [ apply le_plus_n
+ | (*qui auto non chiude il goal*)
+ assumption
+ | (*qui auto non chiude il goal*)
+ assumption
+ | apply le_plus_n
+ | apply le_n
+ ]
+ | auto
+ (*apply H5.
+ apply le_n*)
+ | apply le_plus_n
+ | apply le_n
+ ]
+ | apply (trans_eq ? ? (map_iter_i (S n) (\lambda m.
+ (g(transpose (h (S n+n1)) (S n+n1)
+ (transpose (h (S n+n1)) (S n+n1) (h m)))) )f n1))
+ [ simplify.
+ fold simplify (transpose (h (S n+n1)) (S n+n1) (S n+n1)).
+ apply eq_f2
+ [ auto
+ (*apply eq_f.
+ rewrite > transpose_i_j_j.
+ rewrite > transpose_i_j_i.
+ rewrite > transpose_i_j_j.
+ reflexivity.*)
+ | apply (H2 n1 (\lambda m.(g(transpose (h (S n+n1)) (S n+n1) m))))
+ [ apply permut_S_to_permut_transpose.
+ (*qui auto non chiude il goal*)
+ assumption
+ | intros.
+ unfold transpose.
+ rewrite > (not_eq_to_eqb_false (h m) (h (S n+n1)))
+ [ rewrite > (not_eq_to_eqb_false (h m) (S n+n1))
+ [ simplify.
+ auto
+ (*apply H4.
+ assumption*)
+ | rewrite > H4
+ [ auto
+ (*apply lt_to_not_eq.
+ apply (trans_lt ? n1)
+ [ assumption
+ | simplify.
+ unfold lt.
+ apply le_S_S.
+ apply le_plus_n
+ ]*)
+ | assumption
+ ]
+ ]
+ | unfold permut in H3.
+ elim H3.
+ simplify.
+ unfold Not.
+ intro.
+ apply (lt_to_not_eq m (S n+n1))
+ [ auto
+ (*apply (trans_lt ? n1)
+ [ assumption
+ | simplify.
+ unfold lt.
+ apply le_S_S.
+ apply le_plus_n
+ ]*)
+ | unfold injn in H7.
+ apply (H7 m (S n+n1))
+ [ auto
+ (*apply (trans_le ? n1)
+ [ apply lt_to_le.
+ assumption
+ | apply le_plus_n
+ ]*)
+ | apply le_n
+ | assumption
+ ]
+ ]
+ ]
+ ]
+ ]
+ | apply eq_map_iter_i.
+ intros.
+ auto
+ (*rewrite > transpose_transpose.
+ reflexivity*)
+ ]
+ ]
+]
+qed.
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/plus".
+
+include "nat/nat.ma".
+
+let rec plus n m \def
+ match n with
+ [ O \Rightarrow m
+ | (S p) \Rightarrow S (plus p m) ].
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "natural plus" 'plus x y = (cic:/matita/nat/plus/plus.con x y).
+
+theorem plus_n_O: \forall n:nat. n = n+O.
+intros.elim n
+[ auto
+ (*simplify.
+ reflexivity*)
+| auto
+ (*simplify.
+ apply eq_f.
+ assumption.*)
+]
+qed.
+
+theorem plus_n_Sm : \forall n,m:nat. S (n+m) = n+(S m).
+intros.elim n
+[ auto
+ (*simplify.
+ reflexivity.*)
+| simplify.
+ auto
+ (*
+ apply eq_f.
+ assumption.*)]
+qed.
+
+theorem sym_plus: \forall n,m:nat. n+m = m+n.
+intros.elim n
+[ auto
+ (*simplify.
+ apply plus_n_O.*)
+| simplify.
+ auto
+ (*rewrite > H.
+ apply plus_n_Sm.*)]
+qed.
+
+theorem associative_plus : associative nat plus.
+unfold associative.intros.elim x
+[auto
+ (*simplify.
+ reflexivity.*)
+|simplify.
+ auto
+ (*apply eq_f.
+ assumption.*)
+]
+qed.
+
+theorem assoc_plus : \forall n,m,p:nat. (n+m)+p = n+(m+p)
+\def associative_plus.
+
+theorem injective_plus_r: \forall n:nat.injective nat nat (\lambda m.n+m).
+intro.simplify.intros 2.elim n
+[ exact H
+| auto
+ (*apply H.apply inj_S.apply H1.*)
+]
+qed.
+
+theorem inj_plus_r: \forall p,n,m:nat. p+n = p+m \to n=m
+\def injective_plus_r.
+
+theorem injective_plus_l: \forall m:nat.injective nat nat (\lambda n.n+m).
+intro.simplify.intros.auto.
+(*apply (injective_plus_r m).
+rewrite < sym_plus.
+rewrite < (sym_plus y).
+assumption.*)
+qed.
+
+theorem inj_plus_l: \forall p,n,m:nat. n+p = m+p \to n=m
+\def injective_plus_l.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / Matita is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/primes".
+
+include "nat/div_and_mod.ma".
+include "nat/minimization.ma".
+include "nat/sigma_and_pi.ma".
+include "nat/factorial.ma".
+
+inductive divides (n,m:nat) : Prop \def
+witness : \forall p:nat.m = times n p \to divides n m.
+
+interpretation "divides" 'divides n m = (cic:/matita/nat/primes/divides.ind#xpointer(1/1) n m).
+interpretation "not divides" 'ndivides n m =
+ (cic:/matita/logic/connectives/Not.con (cic:/matita/nat/primes/divides.ind#xpointer(1/1) n m)).
+
+theorem reflexive_divides : reflexive nat divides.
+unfold reflexive.
+intros.
+exact (witness x x (S O) (times_n_SO x)).
+qed.
+
+theorem divides_to_div_mod_spec :
+\forall n,m. O < n \to n \divides m \to div_mod_spec m n (m / n) O.
+intros.
+elim H1.
+rewrite > H2.
+constructor 1
+[ assumption
+| apply (lt_O_n_elim n H).
+ intros.
+ auto
+ (*rewrite < plus_n_O.
+ rewrite > div_times.
+ apply sym_times*)
+]
+qed.
+
+theorem div_mod_spec_to_divides :
+\forall n,m,p. div_mod_spec m n p O \to n \divides m.
+intros.
+elim H.
+auto.
+(*apply (witness n m p).
+rewrite < sym_times.
+rewrite > (plus_n_O (p*n)).
+assumption*)
+qed.
+
+theorem divides_to_mod_O:
+\forall n,m. O < n \to n \divides m \to (m \mod n) = O.
+intros.
+apply (div_mod_spec_to_eq2 m n (m / n) (m \mod n) (m / n) O)
+[ auto
+ (*apply div_mod_spec_div_mod.
+ assumption*)
+| auto
+ (*apply divides_to_div_mod_spec;assumption*)
+]
+qed.
+
+theorem mod_O_to_divides:
+\forall n,m. O< n \to (m \mod n) = O \to n \divides m.
+intros.
+apply (witness n m (m / n)).
+rewrite > (plus_n_O (n * (m / n))).
+rewrite < H1.
+rewrite < sym_times.
+auto.
+(* Andrea: perche' hint non lo trova ?*)
+(*apply div_mod.
+assumption.*)
+qed.
+
+theorem divides_n_O: \forall n:nat. n \divides O.
+intro.
+auto.
+(*apply (witness n O O).
+apply times_n_O.*)
+qed.
+
+theorem divides_n_n: \forall n:nat. n \divides n.
+auto.
+(*intro.
+apply (witness n n (S O)).
+apply times_n_SO.*)
+qed.
+
+theorem divides_SO_n: \forall n:nat. (S O) \divides n.
+intro.
+auto.
+(*apply (witness (S O) n n).
+simplify.
+apply plus_n_O.*)
+qed.
+
+theorem divides_plus: \forall n,p,q:nat.
+n \divides p \to n \divides q \to n \divides p+q.
+intros.
+elim H.
+elim H1.
+apply (witness n (p+q) (n2+n1)).
+auto.
+(*rewrite > H2.
+rewrite > H3.
+apply sym_eq.
+apply distr_times_plus.*)
+qed.
+
+theorem divides_minus: \forall n,p,q:nat.
+divides n p \to divides n q \to divides n (p-q).
+intros.
+elim H.
+elim H1.
+apply (witness n (p-q) (n2-n1)).
+auto.
+(*rewrite > H2.
+rewrite > H3.
+apply sym_eq.
+apply distr_times_minus.*)
+qed.
+
+theorem divides_times: \forall n,m,p,q:nat.
+n \divides p \to m \divides q \to n*m \divides p*q.
+intros.
+elim H.
+elim H1.
+apply (witness (n*m) (p*q) (n2*n1)).
+rewrite > H2.
+rewrite > H3.
+apply (trans_eq nat ? (n*(m*(n2*n1))))
+[ apply (trans_eq nat ? (n*(n2*(m*n1))))
+ [ apply assoc_times
+ | apply eq_f.
+ apply (trans_eq nat ? ((n2*m)*n1))
+ [ auto
+ (*apply sym_eq.
+ apply assoc_times*)
+ | rewrite > (sym_times n2 m).
+ apply assoc_times
+ ]
+ ]
+| auto
+ (*apply sym_eq.
+ apply assoc_times*)
+]
+qed.
+
+theorem transitive_divides: transitive ? divides.
+unfold.
+intros.
+elim H.
+elim H1.
+apply (witness x z (n2*n)).
+auto.
+(*rewrite > H3.
+rewrite > H2.
+apply assoc_times.*)
+qed.
+
+variant trans_divides: \forall n,m,p.
+ n \divides m \to m \divides p \to n \divides p \def transitive_divides.
+
+theorem eq_mod_to_divides:\forall n,m,p. O< p \to
+mod n p = mod m p \to divides p (n-m).
+intros.
+cut (n \le m \or \not n \le m)
+[ elim Hcut
+ [ cut (n-m=O)
+ [ auto
+ (*rewrite > Hcut1.
+ apply (witness p O O).
+ apply times_n_O*)
+ | auto
+ (*apply eq_minus_n_m_O.
+ assumption*)
+ ]
+ | apply (witness p (n-m) ((div n p)-(div m p))).
+ rewrite > distr_times_minus.
+ rewrite > sym_times.
+ rewrite > (sym_times p).
+ cut ((div n p)*p = n - (mod n p))
+ [ rewrite > Hcut1.
+ rewrite > eq_minus_minus_minus_plus.
+ rewrite > sym_plus.
+ rewrite > H1.
+ auto
+ (*rewrite < div_mod
+ [ reflexivity
+ | assumption
+ ]*)
+ | apply sym_eq.
+ apply plus_to_minus.
+ rewrite > sym_plus.
+ auto
+ (*apply div_mod.
+ assumption*)
+ ]
+ ]
+| apply (decidable_le n m)
+]
+qed.
+
+theorem antisymmetric_divides: antisymmetric nat divides.
+unfold antisymmetric.
+intros.
+elim H.
+elim H1.
+apply (nat_case1 n2)
+[ intro.
+ rewrite > H3.
+ rewrite > H2.
+ rewrite > H4.
+ rewrite < times_n_O.
+ reflexivity
+| intros.
+ apply (nat_case1 n)
+ [ intro.
+ rewrite > H2.
+ rewrite > H3.
+ rewrite > H5.
+ auto
+ (*rewrite < times_n_O.
+ reflexivity*)
+ | intros.
+ apply antisymmetric_le
+ [ rewrite > H2.
+ rewrite > times_n_SO in \vdash (? % ?).
+ apply le_times_r.
+ rewrite > H4.
+ auto
+ (*apply le_S_S.
+ apply le_O_n*)
+ | rewrite > H3.
+ rewrite > times_n_SO in \vdash (? % ?).
+ apply le_times_r.
+ rewrite > H5.
+ auto
+ (*apply le_S_S.
+ apply le_O_n*)
+ ]
+ ]
+]
+qed.
+
+(* divides le *)
+theorem divides_to_le : \forall n,m. O < m \to n \divides m \to n \le m.
+intros.
+elim H1.
+rewrite > H2.
+cut (O < n2)
+[ apply (lt_O_n_elim n2 Hcut).
+ intro.
+ auto
+ (*rewrite < sym_times.
+ simplify.
+ rewrite < sym_plus.
+ apply le_plus_n*)
+| elim (le_to_or_lt_eq O n2)
+ [ assumption
+ | absurd (O<m)
+ [ assumption
+ | rewrite > H2.
+ rewrite < H3.
+ rewrite < times_n_O.
+ apply (not_le_Sn_n O)
+ ]
+ | apply le_O_n
+ ]
+]
+qed.
+
+theorem divides_to_lt_O : \forall n,m. O < m \to n \divides m \to O < n.
+intros.
+elim H1.
+elim (le_to_or_lt_eq O n (le_O_n n))
+[ assumption
+| rewrite < H3.
+ absurd (O < m)
+ [ assumption
+ | rewrite > H2.
+ rewrite < H3.
+ auto
+ (*simplify.
+ exact (not_le_Sn_n O)*)
+ ]
+]
+qed.
+
+(* boolean divides *)
+definition divides_b : nat \to nat \to bool \def
+\lambda n,m :nat. (eqb (m \mod n) O).
+
+theorem divides_b_to_Prop :
+\forall n,m:nat. O < n \to
+match divides_b n m with
+[ true \Rightarrow n \divides m
+| false \Rightarrow n \ndivides m].
+intros.
+unfold divides_b.
+apply eqb_elim
+[ intro.
+ simplify.
+ auto
+ (*apply mod_O_to_divides;assumption*)
+| intro.
+ simplify.
+ unfold Not.
+ intro.
+ auto
+ (*apply H1.
+ apply divides_to_mod_O;assumption*)
+]
+qed.
+
+theorem divides_b_true_to_divides :
+\forall n,m:nat. O < n \to
+(divides_b n m = true ) \to n \divides m.
+intros.
+change with
+match true with
+[ true \Rightarrow n \divides m
+| false \Rightarrow n \ndivides m].
+rewrite < H1.
+apply divides_b_to_Prop.
+assumption.
+qed.
+
+theorem divides_b_false_to_not_divides :
+\forall n,m:nat. O < n \to
+(divides_b n m = false ) \to n \ndivides m.
+intros.
+change with
+match false with
+[ true \Rightarrow n \divides m
+| false \Rightarrow n \ndivides m].
+rewrite < H1.
+apply divides_b_to_Prop.
+assumption.
+qed.
+
+theorem decidable_divides: \forall n,m:nat.O < n \to
+decidable (n \divides m).
+intros.
+unfold decidable.
+cut
+(match divides_b n m with
+[ true \Rightarrow n \divides m
+| false \Rightarrow n \ndivides m] \to n \divides m \lor n \ndivides m)
+[ apply Hcut.
+ apply divides_b_to_Prop.
+ assumption
+| elim (divides_b n m)
+ [ left.
+ (*qui auto non chiude il goal, chiuso dalla sola apply H1*)
+ apply H1
+ | right.
+ (*qui auto non chiude il goal, chiuso dalla sola apply H1*)
+ apply H1
+ ]
+]
+qed.
+
+theorem divides_to_divides_b_true : \forall n,m:nat. O < n \to
+n \divides m \to divides_b n m = true.
+intros.
+cut (match (divides_b n m) with
+[ true \Rightarrow n \divides m
+| false \Rightarrow n \ndivides m] \to ((divides_b n m) = true))
+[ apply Hcut.
+ apply divides_b_to_Prop.
+ assumption
+| elim (divides_b n m)
+ [ reflexivity
+ | absurd (n \divides m)
+ [ assumption
+ | (*qui auto non chiude il goal, chiuso dalla sola applicazione di assumption*)
+ assumption
+ ]
+ ]
+]
+qed.
+
+theorem not_divides_to_divides_b_false: \forall n,m:nat. O < n \to
+\lnot(n \divides m) \to (divides_b n m) = false.
+intros.
+cut (match (divides_b n m) with
+[ true \Rightarrow n \divides m
+| false \Rightarrow n \ndivides m] \to ((divides_b n m) = false))
+[ apply Hcut.
+ apply divides_b_to_Prop.
+ assumption
+| elim (divides_b n m)
+ [ absurd (n \divides m)
+ [ (*qui auto non chiude il goal, chiuso dalla sola tattica assumption*)
+ assumption
+ | assumption
+ ]
+ | reflexivity
+ ]
+]
+qed.
+
+(* divides and pi *)
+theorem divides_f_pi_f : \forall f:nat \to nat.\forall n,m,i:nat.
+m \le i \to i \le n+m \to f i \divides pi n f m.
+intros 5.
+elim n
+[ simplify.
+ cut (i = m)
+ [ auto
+ (*rewrite < Hcut.
+ apply divides_n_n*)
+ | apply antisymmetric_le
+ [ assumption
+ | assumption
+ ]
+ ]
+| simplify.
+ cut (i < S n1+m \lor i = S n1 + m)
+ [ elim Hcut
+ [ apply (transitive_divides ? (pi n1 f m))
+ [ apply H1.
+ auto
+ (*apply le_S_S_to_le.
+ assumption*)
+ | auto
+ (*apply (witness ? ? (f (S n1+m))).
+ apply sym_times*)
+ ]
+ | auto
+ (*rewrite > H3.
+ apply (witness ? ? (pi n1 f m)).
+ reflexivity*)
+ ]
+ | auto
+ (*apply le_to_or_lt_eq.
+ assumption*)
+ ]
+]
+qed.
+
+(*
+theorem mod_S_pi: \forall f:nat \to nat.\forall n,i:nat.
+i < n \to (S O) < (f i) \to (S (pi n f)) \mod (f i) = (S O).
+intros.cut (pi n f) \mod (f i) = O.
+rewrite < Hcut.
+apply mod_S.apply trans_lt O (S O).apply le_n (S O).assumption.
+rewrite > Hcut.assumption.
+apply divides_to_mod_O.apply trans_lt O (S O).apply le_n (S O).assumption.
+apply divides_f_pi_f.assumption.
+qed.
+*)
+
+(* divides and fact *)
+theorem divides_fact : \forall n,i:nat.
+O < i \to i \le n \to i \divides n!.
+intros 3.
+elim n
+[ absurd (O<i)
+ [ assumption
+ | auto
+ (*apply (le_n_O_elim i H1).
+ apply (not_le_Sn_O O)*)
+ ]
+| change with (i \divides (S n1)*n1!).
+ apply (le_n_Sm_elim i n1 H2)
+ [ intro.
+ apply (transitive_divides ? n1!)
+ [ auto
+ (*apply H1.
+ apply le_S_S_to_le.
+ assumption*)
+ | auto
+ (*apply (witness ? ? (S n1)).
+ apply sym_times*)
+ ]
+ | intro.
+ auto
+ (*rewrite > H3.
+ apply (witness ? ? n1!).
+ reflexivity*)
+ ]
+]
+qed.
+
+theorem mod_S_fact: \forall n,i:nat.
+(S O) < i \to i \le n \to (S n!) \mod i = (S O).
+intros.
+cut (n! \mod i = O)
+[ rewrite < Hcut.
+ apply mod_S
+ [ auto
+ (*apply (trans_lt O (S O))
+ [ apply (le_n (S O))
+ | assumption
+ ]*)
+ | rewrite > Hcut.
+ assumption
+ ]
+| auto
+ (*apply divides_to_mod_O
+ [ apply (trans_lt O (S O))
+ [ apply (le_n (S O))
+ | assumption
+ ]
+ | apply divides_fact
+ [ apply (trans_lt O (S O))
+ [ apply (le_n (S O))
+ | assumption
+ ]
+ | assumption
+ ]
+ ]*)
+]
+qed.
+
+theorem not_divides_S_fact: \forall n,i:nat.
+(S O) < i \to i \le n \to i \ndivides S n!.
+intros.
+apply divides_b_false_to_not_divides
+[ auto
+ (*apply (trans_lt O (S O))
+ [ apply (le_n (S O))
+ | assumption
+ ]*)
+| unfold divides_b.
+ rewrite > mod_S_fact;auto
+ (*[ simplify.
+ reflexivity
+ | assumption
+ | assumption
+ ]*)
+]
+qed.
+
+(* prime *)
+definition prime : nat \to Prop \def
+\lambda n:nat. (S O) < n \land
+(\forall m:nat. m \divides n \to (S O) < m \to m = n).
+
+theorem not_prime_O: \lnot (prime O).
+unfold Not.
+unfold prime.
+intro.
+elim H.
+apply (not_le_Sn_O (S O) H1).
+qed.
+
+theorem not_prime_SO: \lnot (prime (S O)).
+unfold Not.
+unfold prime.
+intro.
+elim H.
+apply (not_le_Sn_n (S O) H1).
+qed.
+
+(* smallest factor *)
+definition smallest_factor : nat \to nat \def
+\lambda n:nat.
+match n with
+[ O \Rightarrow O
+| (S p) \Rightarrow
+ match p with
+ [ O \Rightarrow (S O)
+ | (S q) \Rightarrow min_aux q (S(S q)) (\lambda m.(eqb ((S(S q)) \mod m) O))]].
+
+(* it works !
+theorem example1 : smallest_prime_factor (S(S(S O))) = (S(S(S O))).
+normalize.reflexivity.
+qed.
+
+theorem example2: smallest_prime_factor (S(S(S(S O)))) = (S(S O)).
+normalize.reflexivity.
+qed.
+
+theorem example3 : smallest_prime_factor (S(S(S(S(S(S(S O))))))) = (S(S(S(S(S(S(S O))))))).
+simplify.reflexivity.
+qed. *)
+
+theorem lt_SO_smallest_factor:
+\forall n:nat. (S O) < n \to (S O) < (smallest_factor n).
+intro.
+apply (nat_case n)
+[ auto
+ (*intro.
+ apply False_ind.
+ apply (not_le_Sn_O (S O) H)*)
+| intro.
+ apply (nat_case m)
+ [ auto
+ (*intro. apply False_ind.
+ apply (not_le_Sn_n (S O) H)*)
+ | intros.
+ change with
+ (S O < min_aux m1 (S(S m1)) (\lambda m.(eqb ((S(S m1)) \mod m) O))).
+ apply (lt_to_le_to_lt ? (S (S O)))
+ [ apply (le_n (S(S O)))
+ | cut ((S(S O)) = (S(S m1)) - m1)
+ [ rewrite > Hcut.
+ apply le_min_aux
+ | apply sym_eq.
+ apply plus_to_minus.
+ auto
+ (*rewrite < sym_plus.
+ simplify.
+ reflexivity*)
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem lt_O_smallest_factor: \forall n:nat. O < n \to O < (smallest_factor n).
+intro.
+apply (nat_case n)
+[ auto
+ (*intro.
+ apply False_ind.
+ apply (not_le_Sn_n O H)*)
+| intro.
+ apply (nat_case m)
+ [ auto
+ (*intro.
+ simplify.
+ unfold lt.
+ apply le_n*)
+ | intros.
+ apply (trans_lt ? (S O))
+ [ auto
+ (*unfold lt.
+ apply le_n*)
+ | apply lt_SO_smallest_factor.
+ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ ]
+]
+qed.
+
+theorem divides_smallest_factor_n :
+\forall n:nat. O < n \to smallest_factor n \divides n.
+intro.
+apply (nat_case n)
+[ intro.
+ auto
+ (*apply False_ind.
+ apply (not_le_Sn_O O H)*)
+| intro.
+ apply (nat_case m)
+ [ intro.
+ auto
+ (*simplify.
+ apply (witness ? ? (S O)).
+ simplify.
+ reflexivity*)
+ | intros.
+ apply divides_b_true_to_divides
+ [ apply (lt_O_smallest_factor ? H)
+ | change with
+ (eqb ((S(S m1)) \mod (min_aux m1 (S(S m1))
+ (\lambda m.(eqb ((S(S m1)) \mod m) O)))) O = true).
+ apply f_min_aux_true.
+ apply (ex_intro nat ? (S(S m1))).
+ split
+ [ auto
+ (*split
+ [ apply le_minus_m
+ | apply le_n
+ ]*)
+ | auto
+ (*rewrite > mod_n_n
+ [ reflexivity
+ | apply (trans_lt ? (S O))
+ [ apply (le_n (S O))
+ | unfold lt.
+ apply le_S_S.
+ apply le_S_S.
+ apply le_O_n
+ ]
+ ]*)
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem le_smallest_factor_n :
+\forall n:nat. smallest_factor n \le n.
+intro.
+apply (nat_case n)
+[ auto
+ (*simplify.
+ apply le_n*)
+| intro.
+ auto
+ (*apply (nat_case m)
+ [ simplify.
+ apply le_n
+ | intro.
+ apply divides_to_le
+ [ unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ | apply divides_smallest_factor_n.
+ unfold lt.
+ apply le_S_S.
+ apply le_O_n
+ ]
+ ]*)
+]
+qed.
+
+theorem lt_smallest_factor_to_not_divides: \forall n,i:nat.
+(S O) < n \to (S O) < i \to i < (smallest_factor n) \to i \ndivides n.
+intros 2.
+apply (nat_case n)
+[ intro.
+ apply False_ind.
+ apply (not_le_Sn_O (S O) H)
+| intro.
+ apply (nat_case m)
+ [ intro.
+ apply False_ind.
+ apply (not_le_Sn_n (S O) H)
+ | intros.
+ apply divides_b_false_to_not_divides
+ [ auto
+ (*apply (trans_lt O (S O))
+ [ apply (le_n (S O))
+ | assumption
+ ]*)
+ | unfold divides_b.
+ apply (lt_min_aux_to_false
+ (\lambda i:nat.eqb ((S(S m1)) \mod i) O) (S(S m1)) m1 i)
+ [ cut ((S(S O)) = (S(S m1)-m1))
+ [ rewrite < Hcut.
+ exact H1
+ | apply sym_eq.
+ apply plus_to_minus.
+ auto
+ (*rewrite < sym_plus.
+ simplify.
+ reflexivity*)
+ ]
+ | exact H2
+ ]
+ ]
+ ]
+]
+qed.
+
+theorem prime_smallest_factor_n :
+\forall n:nat. (S O) < n \to prime (smallest_factor n).
+intro.
+change with ((S(S O)) \le n \to (S O) < (smallest_factor n) \land
+(\forall m:nat. m \divides smallest_factor n \to (S O) < m \to m = (smallest_factor n))).
+intro.
+split
+[ apply lt_SO_smallest_factor.
+ assumption
+| intros.
+ cut (le m (smallest_factor n))
+ [ elim (le_to_or_lt_eq m (smallest_factor n) Hcut)
+ [ absurd (m \divides n)
+ [ apply (transitive_divides m (smallest_factor n))
+ [ assumption
+ | apply divides_smallest_factor_n.
+ auto
+ (*apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | exact H
+ ]*)
+ ]
+ | apply lt_smallest_factor_to_not_divides;auto
+ (*[ exact H
+ | assumption
+ | assumption
+ ]*)
+ ]
+ | assumption
+ ]
+ | apply divides_to_le
+ [ apply (trans_lt O (S O))
+ [ apply (le_n (S O))
+ | apply lt_SO_smallest_factor.
+ exact H
+ ]
+ | assumption
+ ]
+ ]
+]
+qed.
+
+theorem prime_to_smallest_factor: \forall n. prime n \to
+smallest_factor n = n.
+intro.
+apply (nat_case n)
+[ intro.
+ auto
+ (*apply False_ind.
+ apply (not_prime_O H)*)
+| intro.
+ apply (nat_case m)
+ [ intro.
+ auto
+ (*apply False_ind.
+ apply (not_prime_SO H)*)
+ | intro.
+ change with
+ ((S O) < (S(S m1)) \land
+ (\forall m:nat. m \divides S(S m1) \to (S O) < m \to m = (S(S m1))) \to
+ smallest_factor (S(S m1)) = (S(S m1))).
+ intro.
+ elim H.
+ auto
+ (*apply H2
+ [ apply divides_smallest_factor_n.
+ apply (trans_lt ? (S O))
+ [ unfold lt.
+ apply le_n
+ | assumption
+ ]
+ | apply lt_SO_smallest_factor.
+ assumption
+ ]*)
+ ]
+]
+qed.
+
+(* a number n > O is prime iff its smallest factor is n *)
+definition primeb \def \lambda n:nat.
+match n with
+[ O \Rightarrow false
+| (S p) \Rightarrow
+ match p with
+ [ O \Rightarrow false
+ | (S q) \Rightarrow eqb (smallest_factor (S(S q))) (S(S q))]].
+
+(* it works!
+theorem example4 : primeb (S(S(S O))) = true.
+normalize.reflexivity.
+qed.
+
+theorem example5 : primeb (S(S(S(S(S(S O)))))) = false.
+normalize.reflexivity.
+qed.
+
+theorem example6 : primeb (S(S(S(S((S(S(S(S(S(S(S O)))))))))))) = true.
+normalize.reflexivity.
+qed.
+
+theorem example7 : primeb (S(S(S(S(S(S((S(S(S(S((S(S(S(S(S(S(S O))))))))))))))))))) = true.
+normalize.reflexivity.
+qed. *)
+
+theorem primeb_to_Prop: \forall n.
+match primeb n with
+[ true \Rightarrow prime n
+| false \Rightarrow \lnot (prime n)].
+intro.
+apply (nat_case n)
+[ simplify.
+ auto
+ (*unfold Not.
+ unfold prime.
+ intro.
+ elim H.
+ apply (not_le_Sn_O (S O) H1)*)
+| intro.
+ apply (nat_case m)
+ [ simplify.
+ auto
+ (*unfold Not.
+ unfold prime.
+ intro.
+ elim H.
+ apply (not_le_Sn_n (S O) H1)*)
+ | intro.
+ change with
+ match eqb (smallest_factor (S(S m1))) (S(S m1)) with
+ [ true \Rightarrow prime (S(S m1))
+ | false \Rightarrow \lnot (prime (S(S m1)))].
+ apply (eqb_elim (smallest_factor (S(S m1))) (S(S m1)))
+ [ intro.
+ simplify.
+ rewrite < H.
+ apply prime_smallest_factor_n.
+ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_S_S.
+ apply le_O_n*)
+ | intro.
+ simplify.
+ change with (prime (S(S m1)) \to False).
+ intro.
+ auto
+ (*apply H.
+ apply prime_to_smallest_factor.
+ assumption*)
+ ]
+ ]
+]
+qed.
+
+theorem primeb_true_to_prime : \forall n:nat.
+primeb n = true \to prime n.
+intros.
+change with
+match true with
+[ true \Rightarrow prime n
+| false \Rightarrow \lnot (prime n)].
+rewrite < H.
+(*qui auto non chiude il goal*)
+apply primeb_to_Prop.
+qed.
+
+theorem primeb_false_to_not_prime : \forall n:nat.
+primeb n = false \to \lnot (prime n).
+intros.
+change with
+match false with
+[ true \Rightarrow prime n
+| false \Rightarrow \lnot (prime n)].
+rewrite < H.
+(*qui auto non chiude il goal*)
+apply primeb_to_Prop.
+qed.
+
+theorem decidable_prime : \forall n:nat.decidable (prime n).
+intro.
+unfold decidable.
+cut
+(match primeb n with
+[ true \Rightarrow prime n
+| false \Rightarrow \lnot (prime n)] \to (prime n) \lor \lnot (prime n))
+[ apply Hcut.
+ (*qui auto non chiude il goal*)
+ apply primeb_to_Prop
+| elim (primeb n)
+ [ left.
+ (*qui auto non chiude il goal*)
+ apply H
+ | right.
+ (*qui auto non chiude il goal*)
+ apply H
+ ]
+]
+qed.
+
+theorem prime_to_primeb_true: \forall n:nat.
+prime n \to primeb n = true.
+intros.
+cut (match (primeb n) with
+[ true \Rightarrow prime n
+| false \Rightarrow \lnot (prime n)] \to ((primeb n) = true))
+[ apply Hcut.
+ (*qui auto non chiude il goal*)
+ apply primeb_to_Prop
+| elim (primeb n)
+ [ reflexivity.
+ | absurd (prime n)
+ [ assumption
+ | (*qui auto non chiude il goal*)
+ assumption
+ ]
+ ]
+]
+qed.
+
+theorem not_prime_to_primeb_false: \forall n:nat.
+\lnot(prime n) \to primeb n = false.
+intros.
+cut (match (primeb n) with
+[ true \Rightarrow prime n
+| false \Rightarrow \lnot (prime n)] \to ((primeb n) = false))
+[ apply Hcut.
+ (*qui auto non chiude il goal*)
+ apply primeb_to_Prop
+| elim (primeb n)
+ [ absurd (prime n)
+ [ (*qui auto non chiude il goal*)
+ assumption
+ | assumption
+ ]
+ | reflexivity
+ ]
+]
+qed.
+
--- /dev/null
+(**************************************************************************)
+(* __ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/relevant_equations".
+
+include "nat/times.ma".
+include "nat/minus.ma".
+include "nat/gcd.ma".
+(* if gcd is compiled before this, the applys will take too much *)
+
+theorem times_plus_l: \forall n,m,p:nat. (n+m)*p = n*p + m*p.
+intros.
+apply (trans_eq ? ? (p*(n+m)))
+[ apply sym_times
+| apply (trans_eq ? ? (p*n+p*m));auto
+ (*[ apply distr_times_plus
+ | apply eq_f2;
+ apply sym_times
+ ]*)
+]
+qed.
+
+theorem times_minus_l: \forall n,m,p:nat. (n-m)*p = n*p - m*p.
+intros.
+apply (trans_eq ? ? (p*(n-m)))
+[ apply sym_times
+| apply (trans_eq ? ? (p*n-p*m));auto
+ (*[ apply distr_times_minus
+ | apply eq_f2;
+ apply sym_times
+ ]*)
+]
+qed.
+
+theorem times_plus_plus: \forall n,m,p,q:nat. (n + m)*(p + q) =
+n*p + n*q + m*p + m*q.
+intros.
+auto.
+(*apply (trans_eq nat ? ((n*(p+q) + m*(p+q))))
+[ apply times_plus_l
+| rewrite > distr_times_plus.
+ rewrite > distr_times_plus.
+ rewrite < assoc_plus.
+ reflexivity
+]*)
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / Matita is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/sigma_and_pi".
+
+include "nat/factorial.ma".
+include "nat/exp.ma".
+include "nat/lt_arith.ma".
+
+let rec sigma n f m \def
+ match n with
+ [ O \Rightarrow (f m)
+ | (S p) \Rightarrow (f (S p+m))+(sigma p f m)].
+
+let rec pi n f m \def
+ match n with
+ [ O \Rightarrow f m
+ | (S p) \Rightarrow (f (S p+m))*(pi p f m)].
+
+theorem eq_sigma: \forall f,g:nat \to nat.
+\forall n,m:nat.
+(\forall i:nat. m \le i \to i \le m+n \to f i = g i) \to
+(sigma n f m) = (sigma n g m).
+intros 3.
+elim n
+[ simplify.
+ auto
+ (*apply H
+ [ apply le_n
+ | rewrite < plus_n_O.
+ apply le_n
+ ]*)
+| simplify.
+ apply eq_f2
+ [ apply H1
+ [ auto
+ (*change with (m \le (S n1)+m).
+ apply le_plus_n*)
+ | auto
+ (*rewrite > (sym_plus m).
+ apply le_n*)
+ ]
+ | apply H.
+ intros.
+ apply H1
+ [ assumption
+ | auto
+ (*rewrite < plus_n_Sm.
+ apply le_S.
+ assumption*)
+ ]
+ ]
+]
+qed.
+
+theorem eq_pi: \forall f,g:nat \to nat.
+\forall n,m:nat.
+(\forall i:nat. m \le i \to i \le m+n \to f i = g i) \to
+(pi n f m) = (pi n g m).
+intros 3.
+elim n
+[ simplify.
+ auto
+ (*apply H
+ [ apply le_n
+ | rewrite < plus_n_O.
+ apply le_n
+ ] *)
+| simplify.
+ apply eq_f2
+ [ apply H1
+ [ auto
+ (*change with (m \le (S n1)+m).
+ apply le_plus_n*)
+ | auto
+ (*rewrite > (sym_plus m).
+ apply le_n*)
+ ]
+ | apply H.
+ intros.
+ apply H1
+ [ assumption
+ | auto
+ (*rewrite < plus_n_Sm.
+ apply le_S.
+ assumption*)
+ ]
+ ]
+]
+qed.
+
+theorem eq_fact_pi: \forall n. (S n)! = pi n (\lambda m.m) (S O).
+intro.
+elim n
+[ auto
+ (*simplify.
+ reflexivity*)
+| change with ((S(S n1))*(S n1)! = ((S n1)+(S O))*(pi n1 (\lambda m.m) (S O))).
+ rewrite < plus_n_Sm.
+ rewrite < plus_n_O.
+ auto
+ (*apply eq_f.
+ assumption*)
+]
+qed.
+
+theorem exp_pi_l: \forall f:nat\to nat.\forall n,m,a:nat.
+(exp a (S n))*pi n f m= pi n (\lambda p.a*(f p)) m.
+intros.
+elim n
+[ auto
+ (*simplify.
+ rewrite < times_n_SO.
+ reflexivity*)
+| simplify.
+ rewrite < H.
+ rewrite > assoc_times.
+ rewrite > assoc_times in\vdash (? ? ? %).
+ apply eq_f.
+ rewrite < assoc_times.
+ rewrite < assoc_times.
+ auto
+ (*apply eq_f2
+ [ apply sym_times
+ | reflexivity
+ ]*)
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* __ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/times".
+
+include "nat/plus.ma".
+
+let rec times n m \def
+ match n with
+ [ O \Rightarrow O
+ | (S p) \Rightarrow m+(times p m) ].
+
+(*CSC: the URI must disappear: there is a bug now *)
+interpretation "natural times" 'times x y = (cic:/matita/nat/times/times.con x y).
+
+theorem times_n_O: \forall n:nat. O = n*O.
+intros.elim n
+[ auto
+ (*simplify.
+ reflexivity.*)
+| simplify. (* qui auto non funziona: Uncaught exception: Invalid_argument ("List.map2")*)
+ assumption.
+]
+qed.
+
+theorem times_n_Sm :
+\forall n,m:nat. n+(n*m) = n*(S m).
+intros.elim n
+[ auto.
+ (*simplify.reflexivity.*)
+| simplify.
+ auto
+ (*apply eq_f.
+ rewrite < H.
+ transitivity ((n1+m)+n1*m)
+ [ symmetry.
+ apply assoc_plus.
+ | transitivity ((m+n1)+n1*m)
+ [ apply eq_f2.
+ apply sym_plus.
+ reflexivity.
+ | apply assoc_plus.
+ ]
+ ]*)
+]
+qed.
+
+(* NOTA:
+ se non avessi semplificato con auto tutto il secondo ramo della tattica
+ elim n, avrei comunque potuto risolvere direttamente con auto entrambi
+ i rami generati dalla prima applicazione della tattica transitivity
+ (precisamente transitivity ((n1+m)+n1*m)
+ *)
+
+theorem times_n_SO : \forall n:nat. n = n * S O.
+intros.
+rewrite < times_n_Sm.
+auto paramodulation. (*termina la dim anche solo con auto*)
+(*rewrite < times_n_O.
+rewrite < plus_n_O.
+reflexivity.*)
+qed.
+
+theorem times_SSO_n : \forall n:nat. n + n = S (S O) * n.
+intros.
+simplify.
+auto paramodulation. (* termina la dim anche solo con auto*)
+(*rewrite < plus_n_O.
+reflexivity.*)
+qed.
+
+theorem symmetric_times : symmetric nat times.
+unfold symmetric.
+intros.elim x
+[ auto
+ (*simplify.apply times_n_O.*)
+| simplify.
+ auto
+ (*rewrite > H.apply times_n_Sm.*)
+]
+qed.
+
+variant sym_times : \forall n,m:nat. n*m = m*n \def
+symmetric_times.
+
+theorem distributive_times_plus : distributive nat times plus.
+unfold distributive.
+intros.elim x;simplify
+[ reflexivity.
+| auto
+ (*rewrite > H.
+ rewrite > assoc_plus.
+ rewrite > assoc_plus.
+ apply eq_f.
+ rewrite < assoc_plus.
+ rewrite < (sym_plus ? z).
+ rewrite > assoc_plus.
+ reflexivity.*)
+]
+qed.
+
+variant distr_times_plus: \forall n,m,p:nat. n*(m+p) = n*m + n*p
+\def distributive_times_plus.
+
+theorem associative_times: associative nat times.
+unfold associative.intros.
+elim x;simplify
+[ apply refl_eq
+| auto
+ (*rewrite < sym_times.
+ rewrite > distr_times_plus.
+ rewrite < sym_times.
+ rewrite < (sym_times (times n y) z).
+ rewrite < H.
+ apply refl_eq.*)
+]
+qed.
+
+variant assoc_times: \forall n,m,p:nat. (n*m)*p = n*(m*p) \def
+associative_times.
--- /dev/null
+(**************************************************************************)
+(* __ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| A.Asperti, C.Sacerdoti Coen, *)
+(* ||A|| E.Tassi, S.Zacchiroli *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU Lesser General Public License Version 2.1 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/totient".
+
+include "nat/count.ma".
+include "nat/chinese_reminder.ma".
+
+definition totient : nat \to nat \def
+\lambda n. count n (\lambda m. eqb (gcd m n) (S O)).
+
+theorem totient3: totient (S(S(S O))) = (S(S O)).
+reflexivity.
+qed.
+
+theorem totient6: totient (S(S(S(S(S(S O)))))) = (S(S O)).
+reflexivity.
+qed.
+
+theorem totient_times: \forall n,m:nat. (gcd m n) = (S O) \to
+totient (n*m) = (totient n)*(totient m).
+intro.
+apply (nat_case n)
+[ intros.
+ auto
+ (*simplify.
+ reflexivity*)
+| intros 2.
+ apply (nat_case m1)
+ [ rewrite < sym_times.
+ rewrite < (sym_times (totient O)).
+ simplify.
+ intro.
+ reflexivity
+ | intros.
+ unfold totient.
+ apply (count_times m m2 ? ? ?
+ (\lambda b,a. cr_pair (S m) (S m2) a b)
+ (\lambda x. x \mod (S m)) (\lambda x. x \mod (S m2)))
+ [ intros.
+ unfold cr_pair.
+ apply (le_to_lt_to_lt ? (pred ((S m)*(S m2))))
+ [ unfold min.
+ apply le_min_aux_r
+ | auto
+ (*unfold lt.
+ apply (nat_case ((S m)*(S m2)))
+ [ apply le_n
+ | intro.
+ apply le_n
+ ]*)
+ ]
+ | intros.
+ generalize in match (mod_cr_pair (S m) (S m2) a b H1 H2 H).
+ intro.
+ elim H3.
+ apply H4
+ | intros.
+ generalize in match (mod_cr_pair (S m) (S m2) a b H1 H2 H).
+ intro.
+ elim H3.
+ apply H5
+ | intros.
+ generalize in match (mod_cr_pair (S m) (S m2) a b H1 H2 H).
+ intro.
+ elim H3.
+ apply eqb_elim
+ [ intro.
+ rewrite > eq_to_eqb_true
+ [ rewrite > eq_to_eqb_true
+ [ reflexivity
+ | rewrite < H4.
+ rewrite > sym_gcd.
+ rewrite > gcd_mod
+ [ apply (gcd_times_SO_to_gcd_SO ? ? (S m2))
+ [ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | assumption
+ ]
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ ]
+ | rewrite < H5.
+ rewrite > sym_gcd.
+ rewrite > gcd_mod
+ [ apply (gcd_times_SO_to_gcd_SO ? ? (S m))
+ [ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | auto
+ (*rewrite > sym_times.
+ assumption*)
+ ]
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ ]
+ | intro.
+ apply eqb_elim
+ [ intro.
+ apply eqb_elim
+ [ intro.
+ apply False_ind.
+ apply H6.
+ apply eq_gcd_times_SO
+ [ auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ | rewrite < gcd_mod
+ [ auto
+ (*rewrite > H4.
+ rewrite > sym_gcd.
+ assumption*)
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ | rewrite < gcd_mod
+ [ auto
+ (*rewrite > H5.
+ rewrite > sym_gcd.
+ assumption*)
+ | auto
+ (*unfold lt.
+ apply le_S_S.
+ apply le_O_n*)
+ ]
+ ]
+ | intro.
+ reflexivity
+ ]
+ | intro.
+ reflexivity
+ ]
+ ]
+ ]
+ ]
+]
+qed.
\ No newline at end of file