should have been. This has been possible because of some bugs fixed.
intros.change with (frac (cons (x1 + neg m) f) = frac (cons (neg m + x1) f)).
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)
rewrite > Zplus_Zopp.reflexivity.
change with (Z_to_ratio (neg n + - (neg n)) = one).
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
(cic:/matita/logic/connectives/Not.con (cic:/matita/Z/orders/Zlt.con x y)).
theorem irreflexive_Zlt: irreflexive Z Zlt.
-change with (\forall x:Z. x < x \to False).
+unfold irreflexive.unfold Not.
intro.elim x.exact H.
cut (neg n < neg n \to False).
apply Hcut.apply H.simplify.unfold lt.apply not_le_Sn_n.
simplify.reflexivity.
simplify.reflexivity.
qed.
+
theorem symmetric_Ztimes : symmetric Z Ztimes.
change with (\forall x,y:Z. x*y = y*x).
intros.elim x.rewrite > Ztimes_z_OZ.reflexivity.
\def symmetric_Ztimes.
theorem associative_Ztimes: associative Z Ztimes.
-change with (\forall x,y,z:Z. (x*y)*z = x*(y*z)).
+unfold associative.
intros.elim x.
simplify.reflexivity.
elim y.
elim H3;
assumption
| intros;
- change in H5 with (ι(G \sub O · G \sub x) = ι(G \sub O · G \sub y));
+ simplify in H5;
cut (G \sub (ι(G \sub O · G \sub x)) = G \sub (ι(G \sub O · G \sub y)));
[ rewrite > (repr_index_of ? ? (G \sub O · G \sub x)) in Hcut;
rewrite > (repr_index_of ? ? (G \sub O · G \sub y)) in Hcut;
apply andb_elim.
apply eqb_elim.intro.
rewrite > H3.
-change with
-(eqb ((cr_pair m n a b) \mod n) b = true \to
-a = a \land (cr_pair m n a b) \mod n = b).
+simplify.
intro.split.reflexivity.
apply eqb_true_to_eq.assumption.
intro.
-change with (false = true \to
-(cr_pair m n a b) \mod m = a \land (cr_pair m n a b) \mod n = b).
+simplify.
intro.apply False_ind.
apply not_eq_true_false.apply sym_eq.assumption.
apply (f_min_aux_true
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.change with (congruent (f m) (f m \mod p) p).
+elim n. simplify.
apply congruent_n_mod_n.assumption.
-change with (congruent ((f (S n1+m))*(pi n1 f m))
-(((f (S n1+m))\mod p)*(pi n1 (\lambda m.(f m) \mod p) m)) p).
+simplify.
apply congruent_times.
assumption.
apply congruent_n_mod_n.assumption.
intro.elim n.simplify.
rewrite < plus_n_O.
apply eq_sigma.intros.reflexivity.
-change with
-(sigma (m+(S n1)*(S m)) f O =
-sigma m (\lambda a.(f ((S(n1+O))*(S m)+a)) + (sigma n1 (\lambda b.f (b*(S m)+a)) O)) O).
+simplify.
rewrite > sigma_f_g.
rewrite < plus_n_O.
rewrite < H.
(sigma m (\lambda n.(bool_to_nat (f2 n))*(bool_to_nat (f1 i))) O)).
reflexivity.
apply sym_eq. apply sigma_times.
-change in match (pred (S n)) with n.
-change in match (pred (S m)) with m.
+simplify.
apply sym_eq. apply sigma_times.
unfold permut.
split.
injective_times_r.
theorem lt_O_to_injective_times_r: \forall n:nat. O < n \to injective nat nat (\lambda m:nat.n*m).
-change with (\forall n. O < n \to \forall p,q:nat.n*p = n*q \to p=q).
+simplify.
intros 4.
apply (lt_O_n_elim n H).intros.
apply (inj_times_r m).assumption.
\def lt_O_to_injective_times_r.
theorem injective_times_l: \forall n:nat.injective nat nat (\lambda m:nat.m*(S n)).
-change with (\forall n,p,q:nat.p*(S n) = q*(S n) \to p=q).
+simplify.
intros.
-apply (inj_times_r n p q).
+apply (inj_times_r n x y).
rewrite < sym_times.
-rewrite < (sym_times q).
+rewrite < (sym_times y).
assumption.
qed.
injective_times_l.
theorem lt_O_to_injective_times_l: \forall n:nat. O < n \to injective nat nat (\lambda m:nat.m*n).
-change with (\forall n. O < n \to \forall p,q:nat.p*n = q*n \to p=q).
+simplify.
intros 4.
apply (lt_O_n_elim n H).intros.
apply (inj_times_l m).assumption.
| 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;
[ apply (trans_lt ? (S O));
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.change with
-((max n (\lambda p:nat.eqb (n \mod (nth_prime p)) O)) \le
-(max m (\lambda p:nat.eqb (m \mod (nth_prime p)) O))).
+intros.unfold max_prime_factor.
apply f_m_to_le_max.
apply (trans_le ? n).
apply le_max_n.apply divides_to_le.assumption.assumption.
absurd (nth_prime (max_prime_factor n) \divides r).
rewrite < H4.
apply divides_max_prime_factor_n.
-assumption.
-change with (nth_prime (max_prime_factor n) \divides r \to False).
+assumption.unfold Not.
intro.
cut (r \mod (nth_prime (max_prime_factor n)) \neq O).
apply Hcut1.apply divides_to_mod_O.
intro.
apply (nat_case n).reflexivity.
intro.apply (nat_case m).reflexivity.
-intro.change with
+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
apply sym_eq.apply eq_pair_fst_snd.
intros.
rewrite < H.
-change with
-(defactorize_aux (factorize_aux p r (nf_last (pred q))) O = (S(S m1))).
+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).
rewrite > Hcut2.
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).
intro.absurd ((nth_prime i) = (nth_prime j)).
apply (divides_exp_to_eq ? ? (S n)).
apply prime_nth_prime.apply prime_nth_prime.
-assumption.
-change with ((nth_prime i) = (nth_prime j) \to False).
+assumption.unfold Not.
intro.cut (i = j).
apply (not_le_Sn_n i).rewrite > Hcut in \vdash (? ? %).assumption.
apply (injective_nth_prime ? ? H2).
-change with
-(nth_prime i \divides (nth_prime j) \sup n *(defactorize_aux n1 (S j)) \to False).
+unfold Not.simplify.
intro.
cut (nth_prime i \divides (nth_prime j) \sup n
\lor nth_prime i \divides defactorize_aux n1 (S j)).
absurd ((nth_prime i) = (nth_prime j)).
apply (divides_exp_to_eq ? ? n).
apply prime_nth_prime.apply prime_nth_prime.
-assumption.
-change with ((nth_prime i) = (nth_prime j) \to False).
+assumption.unfold Not.
intro.
cut (i = j).
apply (not_le_Sn_n i).rewrite > Hcut1 in \vdash (? ? %).assumption.
simplify.apply le_S_S.
apply le_plus_n.
apply injective_nth_prime.
-(* uffa, perche' semplifica ? *)
-change with (nth_prime (S(max_p g)+i)= nth_prime i).
apply (divides_exp_to_eq ? ? (S n)).
apply prime_nth_prime.apply prime_nth_prime.
rewrite > H.
theorem injective_defactorize_aux: \forall i:nat.
injective nat_fact nat (\lambda f.defactorize_aux f i).
-change with (\forall i:nat.\forall f,g:nat_fact.
-defactorize_aux f i = defactorize_aux g i \to f = g).
+simplify.
intros.
-apply (eq_defactorize_aux_to_eq f g i H).
+apply (eq_defactorize_aux_to_eq x y i H).
qed.
theorem injective_defactorize:
injective nat_fact_all nat defactorize.
-change with (\forall f,g:nat_fact_all.
-defactorize f = defactorize g \to f = g).
+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 *)
\forall f,g: nat_fact_all. factorize (defactorize f) = f.
intros.
apply injective_defactorize.
-(* uffa: perche' semplifica ??? *)
-change with (defactorize(factorize (defactorize f)) = (defactorize f)).
apply defactorize_factorize.
qed.
assumption.
apply False_ind.
apply (prime_to_not_divides_fact p H (pred p)).
-change with (S (pred p) \le p).
+unfold lt.
rewrite < S_pred.apply le_n.
assumption.assumption.
apply divides_times_to_divides.
rewrite > Hcut3.rewrite < times_n_O.
apply mod_O_n.apply sym_eq.apply le_n_O_to_eq.
apply le_S_S_to_le.assumption.
-assumption.
-change with ((S O) \le pred p).
+assumption.unfold lt.
apply le_S_S_to_le.rewrite < S_pred.
unfold prime in H.elim H.assumption.assumption.
unfold prime in H.elim H.apply (trans_lt ? (S O)).
intro.elim p.
absurd (O < n).assumption.apply le_to_not_lt.assumption.
cut ((n1 \divides m) \lor (n1 \ndivides m)).
-change with
-((match divides_b n1 m with
-[ true \Rightarrow n1
-| false \Rightarrow gcd_aux n n1 (m \mod n1)]) \divides m \land
-(match divides_b n1 m with
-[ true \Rightarrow n1
-| false \Rightarrow gcd_aux n n1 (m \mod n1)]) \divides n1).
+simplify.
elim Hcut.rewrite > divides_to_divides_b_true.
simplify.
split.assumption.apply (witness n1 n1 (S O)).apply times_n_SO.
assumption.assumption.
rewrite > not_divides_to_divides_b_false.
-change with
-(gcd_aux n n1 (m \mod n1) \divides m \land
-gcd_aux n n1 (m \mod n1) \divides n1).
+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.
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
d \divides m \to d \divides n \to d \divides gcd_aux p m n.
intro.elim p.
absurd (O < n).assumption.apply le_to_not_lt.assumption.
-change with
-(d \divides
-(match divides_b n1 m with
-[ true \Rightarrow n1
-| false \Rightarrow gcd_aux n n1 (m \mod n1)])).
+simplify.
cut (n1 \divides m \lor n1 \ndivides m).
elim Hcut.
rewrite > divides_to_divides_b_true.
simplify.assumption.
assumption.assumption.
rewrite > not_divides_to_divides_b_false.
-change with (d \divides gcd_aux n n1 (m \mod n1)).
+simplify.
apply H.
cut (O \lt m \mod n1 \lor O = m \mod n1).
elim Hcut1.assumption.
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
absurd (O < n).assumption.apply le_to_not_lt.assumption.
cut (O < m).
cut (n1 \divides m \lor n1 \ndivides m).
-change with
-(\exists a,b.
-a*n1 - b*m = match divides_b n1 m with
-[ true \Rightarrow n1
-| false \Rightarrow gcd_aux n n1 (m \mod n1)]
-\lor
-b*m - a*n1 = match divides_b n1 m with
-[ true \Rightarrow n1
-| false \Rightarrow gcd_aux n n1 (m \mod n1)]).
+simplify.
elim Hcut1.
rewrite > divides_to_divides_b_true.
simplify.
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.
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.change with (gcd n m = (S O)).
+intros.unfold prime in H.
elim H.
apply antisym_le.
-apply not_lt_to_le.
-change with ((S (S O)) \le gcd n m \to False).intro.
+apply not_lt_to_le.unfold Not.unfold lt.
+intro.
apply H1.rewrite < (H3 (gcd n m)).
apply divides_gcd_m.
apply divides_gcd_n.assumption.
theorem monotonic_lt_plus_l:
\forall n:nat.monotonic nat lt (\lambda m.m+n).
-change with (\forall n,p,q:nat. p < q \to p + n < q + n).
+simplify.
intros.
rewrite < sym_plus. rewrite < (sym_plus n).
apply lt_plus_r.assumption.
(* times *)
theorem monotonic_lt_times_r:
\forall n:nat.monotonic nat lt (\lambda m.(S n)*m).
-change with (\forall n,p,q:nat. p < q \to (S n) * p < (S n) * q).
+simplify.
intros.elim n.
simplify.rewrite < plus_n_O.rewrite < plus_n_O.assumption.
-change with (p + (S n1) * p < q + (S n1) * q).
apply lt_plus.assumption.assumption.
qed.
theorem monotonic_lt_times_l:
\forall m:nat.monotonic nat lt (\lambda n.n * (S m)).
-change with
-(\forall n,p,q:nat. p < q \to p*(S n) < q*(S n)).
+simplify.
intros.
-rewrite < sym_times.rewrite < (sym_times (S n)).
+rewrite < sym_times.rewrite < (sym_times (S m)).
apply lt_times_r.assumption.
qed.
theorem smallest_factor_fact: \forall n:nat.
n < smallest_factor (S n!).
intros.
-apply not_le_to_lt.
-change with (smallest_factor (S n!) \le n \to False).intro.
+apply not_le_to_lt.unfold Not.
+intro.
apply (not_divides_S_fact n (smallest_factor(S n!))).
apply lt_SO_smallest_factor.
unfold lt.apply le_S_S.apply le_SO_fact.
apply smallest_factor_fact.
apply le_smallest_factor_n.
(* Andrea: ancora hint non lo trova *)
-apply prime_smallest_factor_n.
-change with ((S(S O)) \le S (S n1)!).
+apply prime_smallest_factor_n.unfold lt.
apply le_S.apply le_SSO_fact.
unfold lt.apply le_S_S.assumption.
qed.
theorem prime_nth_prime : \forall n:nat.prime (nth_prime n).
intro.
-apply (nat_case n).
-change with (prime (S(S O))).
+apply (nat_case n).simplify.
apply (primeb_to_Prop (S(S O))).
intro.
change with
apply le_n_fact_n.
apply le_smallest_factor_n.
apply prime_to_primeb_true.
-apply prime_smallest_factor_n.
-change with ((S(S O)) \le S (nth_prime m)!).
+apply prime_smallest_factor_n.unfold lt.
apply le_S_S.apply le_SO_fact.
qed.
(* properties of nth_prime *)
theorem increasing_nth_prime: increasing nth_prime.
-change with (\forall n:nat. (nth_prime n) < (nth_prime (S n))).
+unfold increasing.
intros.
change with
(let previous_prime \def (nth_prime n) in
match p_ord_aux p n m with
[ (pair q r) \Rightarrow n = m \sup q *r ].
intro.
-elim p.
-change with
-match (
-match n \mod m with
- [ O \Rightarrow pair nat nat O n
- | (S a) \Rightarrow pair nat nat O n] )
-with
- [ (pair q r) \Rightarrow n = m \sup q * r ].
+elim p.simplify.
apply (nat_case (n \mod m)).
simplify.apply plus_n_O.
intros.
-simplify.apply plus_n_O.
-change with
-match (
-match n1 \mod m with
- [ O \Rightarrow
- match (p_ord_aux n (n1 / m) m) with
- [ (pair q r) \Rightarrow pair nat nat (S q) r]
- | (S a) \Rightarrow pair nat nat O n1] )
-with
- [ (pair q r) \Rightarrow n1 = m \sup q * r].
+simplify.apply plus_n_O.
+simplify.
apply (nat_case1 (n1 \mod m)).intro.
-change with
-match (
- match (p_ord_aux n (n1 / m) m) with
- [ (pair q r) \Rightarrow pair nat nat (S q) r])
-with
- [ (pair q r) \Rightarrow n1 = m \sup q * r].
+simplify.
generalize in match (H (n1 / m) m).
elim (p_ord_aux n (n1 / m) m).
simplify.
simplify.
rewrite < plus_n_O.
apply (nat_case p).
-change with
- (match n \mod m with
- [ O \Rightarrow pair nat nat O n
- | (S a) \Rightarrow pair nat nat O n]
- = pair nat nat O n).
+simplify.
elim (n \mod m).simplify.reflexivity.simplify.reflexivity.
intro.
-change with
- (match n \mod m with
- [ O \Rightarrow
- match (p_ord_aux m1 (n / m) m) with
- [ (pair q r) \Rightarrow pair nat nat (S q) r]
- | (S a) \Rightarrow pair nat nat O n]
- = pair nat nat O n).
+simplify.
cut (O < n \mod m \lor O = n \mod m).
elim Hcut.apply (lt_O_n_elim (n \mod m) H3).
intros. simplify.reflexivity.
generalize in match H3.
apply (nat_case p).intro.apply False_ind.apply (not_le_Sn_O n1 H4).
intros.
-change with
- (match ((m \sup (S n1) *n) \mod m) with
- [ O \Rightarrow
- match (p_ord_aux m1 ((m \sup (S n1) *n) / m) m) with
- [ (pair q r) \Rightarrow pair nat nat (S q) r]
- | (S a) \Rightarrow pair nat nat O (m \sup (S n1) *n)]
- = pair nat nat (S n1) n).
+simplify. fold simplify (m \sup (S n1)).
cut (((m \sup (S n1)*n) \mod m) = O).
rewrite > Hcut.
-change with
-(match (p_ord_aux m1 ((m \sup (S n1)*n) / m) m) with
- [ (pair q r) \Rightarrow pair nat nat (S q) r]
- = pair nat nat (S n1) n).
+simplify.fold simplify (m \sup (S n1)).
cut ((m \sup (S n1) *n) / m = m \sup n1 *n).
rewrite > Hcut1.
rewrite > (H2 m1). simplify.reflexivity.
apply le_S_S_to_le.assumption.
(* div_exp *)
-change with ((m* m \sup n1 *n) / m = m \sup n1 * n).
+simplify.
rewrite > assoc_times.
apply (lt_O_n_elim m H).
intro.apply div_times.
[ (pair q r) \Rightarrow r \mod m \neq O].
intro.elim p.absurd (O < n).assumption.
apply le_to_not_lt.assumption.
-change with
-match
- (match n1 \mod m with
- [ O \Rightarrow
- match (p_ord_aux n(n1 / m) m) with
- [ (pair q r) \Rightarrow pair nat nat (S q) r]
- | (S a) \Rightarrow pair nat nat O n1])
-with
- [ (pair q r) \Rightarrow r \mod m \neq O].
+simplify.
apply (nat_case1 (n1 \mod m)).intro.
generalize in match (H (n1 / m) m).
elim (p_ord_aux n (n1 / m) m).
apply le_S_S_to_le.
apply (trans_le ? n1).change with (n1 / m < n1).
apply lt_div_n_m_n.assumption.assumption.assumption.
-intros.
-change with (n1 \mod m \neq O).
+intros.simplify.
rewrite > H4.
unfold Not.intro.
apply (not_eq_O_S m1).
qed.
theorem le_to_not_lt: \forall n,m:nat. le n m \to Not (lt m n).
-intros.
-change with (Not (le (S m) n)).
+intros.unfold Not.unfold lt.
apply lt_to_not_le.unfold lt.
apply le_S_S.assumption.
qed.
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.
-change with
-(f (g (S n)) (g n) =
-f (g (transpose n (S n) (S n))) (g (transpose n (S n) n))).
+intros.simplify.fold simplify (transpose n (S n) (S n)).
rewrite > transpose_i_j_i.
rewrite > transpose_i_j_j.
apply H1.
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)).
-change with
-(f (g (transpose (h (S n+n1)) (S n+n1) (S n+n1)))
-(map_iter_i n (\lambda m.
-g (transpose (h (S n+n1)) (S n+n1) m)) f n1)
-=
-f
-(g(transpose (h (S n+n1)) (S n+n1)
-(transpose (h (S n+n1)) (S n+n1) (h (S n+n1)))))
-(map_iter_i 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.apply eq_f.
rewrite > transpose_i_j_j.
rewrite > transpose_i_j_i.
match divides_b n m with
[ true \Rightarrow n \divides m
| false \Rightarrow n \ndivides m].
-intros.
-change with
-match eqb (m \mod n) O with
-[ true \Rightarrow n \divides m
-| false \Rightarrow n \ndivides m].
+intros.unfold divides_b.
apply eqb_elim.
intro.simplify.apply mod_O_to_divides.assumption.assumption.
intro.simplify.unfold Not.intro.apply H1.apply divides_to_mod_O.assumption.assumption.
theorem decidable_divides: \forall n,m:nat.O < n \to
decidable (n \divides m).
-intros.change with ((n \divides m) \lor n \ndivides m).
+intros.unfold decidable.
cut
(match divides_b n m with
[ true \Rightarrow n \divides m
intros.
apply divides_b_false_to_not_divides.
apply (trans_lt O (S O)).apply (le_n (S O)).assumption.
-change with ((eqb ((S n!) \mod i) O) = false).
+unfold divides_b.
rewrite > mod_S_fact.simplify.reflexivity.
assumption.assumption.
qed.
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.
-apply (trans_lt O (S O)).apply (le_n (S O)).assumption.
-change with ((eqb ((S(S m1)) \mod i) O) = false).
+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)).
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
+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.
[ 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.change with (prime (S(S m1))).
+intro.simplify.
rewrite < H.
apply prime_smallest_factor_n.
unfold lt.apply le_S_S.apply le_S_S.apply le_O_n.
-intro.change with (\lnot (prime (S(S m1)))).
+intro.simplify.
change with (prime (S(S m1)) \to False).
intro.apply H.
apply prime_to_smallest_factor.
qed.
theorem decidable_prime : \forall n:nat.decidable (prime n).
-intro.change with ((prime n) \lor \lnot (prime n)).
+intro.unfold decidable.
cut
(match primeb n with
[ true \Rightarrow prime n
intros.unfold cr_pair.
apply (le_to_lt_to_lt ? (pred ((S m)*(S m2)))).
unfold min.
-apply le_min_aux_r.
-change with ((S (pred ((S m)*(S m2)))) \le ((S m)*(S m2))).
+apply le_min_aux_r.unfold lt.
apply (nat_case ((S m)*(S m2))).apply le_n.
intro.apply le_n.
intros.