1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
17 include "nat/times.ma".
18 include "nat/orders.ma".
20 theorem easy: ∀n,m. n * m = O → n = O ∨ m = O.
24 by (refl_eq ? O) we proved (O = O) (trivial).
25 by or_introl we proved (O = O ∨ m = O) (trivial2).
26 using (λ_.trivial2) we proved (O*m=O → O=O ∨ m=O) (base_case).
29 (∀n1. (n1 * m = O → n1 = O ∨ m = O) → (S n1) * m = O → (S n1) = O ∨ m = O)
32 suppose (n1 * m = O → n1 = O ∨ m = O) (inductive_hyp).
34 by or_intror we proved (S n1 = O ∨ O = O) (pre_base_case2).
35 using (λ_.pre_base_case2) we proved (S n1*O = O → S n1 = O ∨ O = O) (base_case2).
38 (∀m1. (S n1 * m1 = O → S n1 = O ∨ m1 = O) →
39 (S n1 * S m1 = O → S n1 = O ∨ S m1 = O)) (inductive_hyp2).
41 suppose (S n1 * m1 = O → S n1 = O ∨ m1 = O) (useless).
42 suppose (S n1 * S m1 = O) (absurd_hyp).
43 simplify in absurd_hyp.
44 by sym_eq we proved (O = S (m1+n1*S m1)) (absurd_hyp').
45 by not_eq_O_S we proved False (the_absurd).
46 by (False_ind ? the_absurd)
49 using (nat_ind (λm.S n1 * m = O → S n1 = O ∨ m = O) base_case2 inductive_hyp2 m)
52 using (nat_ind (λn.n*m=O → n=O ∨ m=O) base_case inductive_case n)
56 theorem easy2: ∀n,m. n * m = O → n = O ∨ m = O.
69 lapply (sym_eq ? ? ? H2);
70 elim (not_eq_O_S ? Hletin)
75 theorem easy15: ∀n,m. n * m = O → n = O ∨ m = O.
79 we proved (O = O) (trivial).
80 we proved (O = O ∨ m = O) (trivial2).
81 we proved (O*m=O → O=O ∨ m=O) (base_case).
84 (∀n1. (n1 * m = O → n1 = O ∨ m = O) → (S n1) * m = O → (S n1) = O ∨ m = O)
87 suppose (n1 * m = O → n1 = O ∨ m = O) (inductive_hyp).
89 we proved (S n1 = O ∨ O = O) (pre_base_case2).
90 we proved (S n1*O = O → S n1 = O ∨ O = O) (base_case2).
93 (∀m1. (S n1 * m1 = O → S n1 = O ∨ m1 = O) →
94 (S n1 * S m1 = O → S n1 = O ∨ S m1 = O)) (inductive_hyp2).
96 suppose (S n1 * m1 = O → S n1 = O ∨ m1 = O) (useless).
97 suppose (S n1 * S m1 = O) (absurd_hyp).
98 simplify in absurd_hyp.
99 we proved (O = S (m1+n1*S m1)) (absurd_hyp').
100 we proved False (the_absurd).
101 we proceed by cases on the_absurd to prove (S n1=O ∨ S m1=O).
103 using (nat_ind (λm.S n1 * m = O → S n1 = O ∨ m = O) base_case2 inductive_hyp2 m)
106 using (nat_ind (λn.n*m=O → n=O ∨ m=O) base_case inductive_case n)
110 theorem easy3: ∀A:Prop. (A ∧ ∃n:nat.n ≠ n) → True.
112 suppose (P ∧ ∃m:nat.m ≠ m) (H).
113 by H we have P (H1) and (∃x:nat.x≠x) (H2).
114 by H2 let q:nat such that (q ≠ q) (Ineq).
118 theorem easy4: ∀n,m,p. n = m → S m = S p → n = S p → S n = n.
123 suppose (S m = S p) (K).
124 suppose (n = S p) (L).
125 conclude (S n) = (S m) by H.
131 theorem easy45: ∀n,m,p. n = m → S m = S p → n = S p → S n = n.
136 suppose (S m = S p) (K).
137 suppose (n = S p) (L).
138 conclude (S n) = (S m).
144 theorem easy5: ∀n:nat. n*O=O.
146 (* Bug here: False should be n*0=0 *)
147 we proceed by induction on n to prove False.
149 the thesis becomes (O*O=O).
152 by induction hypothesis we know (m*O=O) (I).
153 the thesis becomes (S m * O = O).
154 (* Bug here: missing that is equivalent to *)
159 inductive tree : Type ≝
161 | Node: tree → tree → tree.
166 | (Node t1 t2) ⇒ S ((size t1) + (size t2))
169 theorem easy6: ∀t. O ≮ O → O < size t → t ≠ Empty.
171 suppose (O ≮ O) (trivial).
172 (*Bug here: False should be something else *)
173 we proceed by induction on t to prove False.
175 the thesis becomes (O < size Empty → Empty ≠ Empty).
176 suppose (O < size Empty) (absurd)
177 that is equivalent to (O < O).
178 (* Here the "natural" language is not natural at all *)
179 we proceed by induction on (trivial absurd) to prove False.
180 (*Bug here: this is what we want
181 case Node (t1:tree) (t2:tree).
182 by induction hypothesis we know (O < size t1 → t1 ≠ Empty) (Ht1).
183 by induction hypothesis we know (O < size t2 → t2 ≠ Empty) (Ht2). *)
184 (*This is the best we can do right now*)
187 by induction hypothesis we know (O < size t1 → t1 ≠ Empty) (Ht1).
189 by induction hypothesis we know (O < size t2 → t2 ≠ Empty) (Ht2).
190 suppose (O < size (Node t1 t2)) (useless).
191 we need to prove (Node t1 t2 ≠ Empty) (final)
192 or equivalently (Node t1 t2 = Empty → False).
193 suppose (Node t1 t2 = Empty) (absurd).
194 (* Discriminate should really generate a theorem to be useful with
195 declarative tactics *)