]> matita.cs.unibo.it Git - helm.git/blob - helm/matita/tests/match.ma
036cc393b45214e60a96442f8bd1a4477b01017d
[helm.git] / helm / matita / tests / match.ma
1 (**************************************************************************)
2 (*       ___                                                                *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||       A.Asperti, C.Sacerdoti Coen,                          *)
8 (*      ||A||       E.Tassi, S.Zacchiroli                                 *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15
16 inductive True: Prop \def
17 I : True.
18
19 inductive False: Prop \def .
20
21 definition Not: Prop \to Prop \def
22 \lambda A. (A \to False).
23
24 theorem absurd : \forall A,C:Prop. A \to Not A \to C.
25 intros. elim (H1 H).
26 qed.
27
28 inductive And (A,B:Prop) : Prop \def
29     conj : A \to B \to (And A B).
30
31 theorem proj1: \forall A,B:Prop. (And A B) \to A.
32 intros. elim H. assumption.
33 qed.
34
35 theorem proj2: \forall A,B:Prop. (And A B) \to A.
36 intros. elim H. assumption.
37 qed.
38
39 inductive Or (A,B:Prop) : Prop \def
40      or_introl : A \to (Or A B)
41    | or_intror : B \to (Or A B).
42
43 inductive ex (A:Type) (P:A \to Prop) : Prop \def
44     ex_intro: \forall x:A. P x \to ex A P.
45
46 inductive ex2 (A:Type) (P,Q:A \to Prop) : Prop \def
47     ex_intro2: \forall x:A. P x \to Q x \to ex2 A P Q.
48
49 inductive eq (A:Type) (x:A) : A \to Prop \def
50     refl_equal : eq A x x.
51
52 theorem sym_eq : \forall A:Type.\forall x,y:A. eq A x y  \to eq A y x.
53 intros. elim H. apply refl_equal.
54 qed.
55
56 theorem trans_eq : \forall A:Type.
57 \forall x,y,z:A. eq A x y  \to eq A y z \to eq A x z.
58 intros.elim H1.assumption.
59 qed.
60
61 theorem f_equal: \forall  A,B:Type.\forall f:A\to B.
62 \forall x,y:A. eq A x y \to eq B (f x) (f y).
63 intros.elim H.apply refl_equal.
64 qed.
65
66 theorem f_equal2: \forall  A,B,C:Type.\forall f:A\to B \to C.
67 \forall x1,x2:A. \forall y1,y2:B.
68 eq A x1 x2\to eq B y1 y2\to eq C (f x1 y1) (f x2 y2).
69 intros.elim H1.elim H.apply refl_equal.
70 qed.
71
72 inductive nat : Set \def
73   | O : nat
74   | S : nat \to nat.
75
76 definition pred: nat \to nat \def
77 \lambda n:nat. match n with
78 [ O \Rightarrow  O
79 | (S u) \Rightarrow u ].
80
81 theorem pred_Sn : \forall n:nat.
82 (eq nat n (pred (S n))).
83 intros.apply refl_equal.
84 qed.
85
86 theorem injective_S : \forall n,m:nat. 
87 (eq nat (S n) (S m)) \to (eq nat n m).
88 intros.(elim (sym_eq ? ? ? (pred_Sn n))).(elim (sym_eq ? ? ? (pred_Sn m))).
89 apply f_equal. assumption.
90 qed.
91
92 theorem not_eq_S  : \forall n,m:nat. 
93 Not (eq nat n m) \to Not (eq nat (S n) (S m)).
94 intros. simplify.intros.
95 apply H.apply injective_S.assumption.
96 qed.
97
98 definition not_zero : nat \to Prop \def
99 \lambda n: nat.
100   match n with
101   [ O \Rightarrow False
102   | (S p) \Rightarrow True ].
103
104 theorem O_S : \forall n:nat. Not (eq nat O (S n)).
105 intros.simplify.intros.
106 cut (not_zero O).exact Hcut.elim (sym_eq ? ? ? H).
107 exact I.
108 qed.
109
110 theorem n_Sn : \forall n:nat. Not (eq nat n (S n)).
111 intros.elim n.apply O_S.apply not_eq_S.assumption.
112 qed.
113
114
115 let rec plus n m \def 
116  match n with 
117  [ O \Rightarrow m
118  | (S p) \Rightarrow S (plus p m) ].
119
120 theorem plus_n_O: \forall n:nat. eq nat n (plus n O).
121 intros.elim n.simplify.apply refl_equal.simplify.apply f_equal.assumption.
122 qed.
123
124 theorem plus_n_Sm : \forall n,m:nat. eq nat (S (plus n  m)) (plus n (S m)).
125 intros.elim n.simplify.apply refl_equal.simplify.apply f_equal.assumption.
126 qed.
127
128 theorem sym_plus: \forall n,m:nat. eq nat (plus n m) (plus m n).
129 intros.elim n.simplify.apply plus_n_O.
130 simplify.elim (sym_eq ? ? ? H).apply plus_n_Sm.
131 qed.
132
133 theorem assoc_plus: 
134 \forall n,m,p:nat. eq nat (plus (plus n m) p) (plus n (plus m p)).
135 intros.elim n.simplify.apply refl_equal.simplify.apply f_equal.assumption.
136 qed.
137
138 let rec times n m \def 
139  match n with 
140  [ O \Rightarrow O
141  | (S p) \Rightarrow (plus m (times p m)) ].
142
143 theorem times_n_O: \forall n:nat. eq nat O (times n O).
144 intros.elim n.simplify.apply refl_equal.simplify.assumption.
145 qed.
146
147 theorem times_n_Sm : 
148 \forall n,m:nat. eq nat (plus n (times n  m)) (times n (S m)).
149 intros.elim n.simplify.apply refl_equal.
150 simplify.apply f_equal.elim H.
151 apply trans_eq ? ? (plus (plus e m) (times e m)).apply sym_eq.
152 apply assoc_plus.apply trans_eq ? ? (plus (plus m e) (times e m)).
153 apply f_equal2.
154 apply sym_plus.apply refl_equal.apply assoc_plus.
155 qed.
156
157 theorem sym_times : 
158 \forall n,m:nat. eq nat (times n m) (times m n).
159 intros.elim n.simplify.apply times_n_O.
160 simplify.elim (sym_eq ? ? ? H).apply times_n_Sm.
161 qed.
162
163 let rec minus n m \def 
164  match n with 
165  [ O \Rightarrow O
166  | (S p) \Rightarrow 
167         match m with
168         [O \Rightarrow (S p)
169         | (S q) \Rightarrow minus p q ]].
170
171 theorem nat_case :
172 \forall n:nat.\forall P:nat \to Prop. 
173 P O \to  (\forall m:nat. P (S m)) \to P n.
174 intros.elim n.assumption.apply H1.
175 qed.
176
177 theorem nat_double_ind :
178 \forall R:nat \to nat \to Prop.
179 (\forall n:nat. R O n) \to 
180 (\forall n:nat. R (S n) O) \to 
181 (\forall n,m:nat. R n m \to R (S n) (S m)) \to \forall n,m:nat. R n m.
182 intros.cut \forall m:nat.R n m.apply Hcut.elim n.apply H.
183 apply nat_case m1.apply H1.intros.apply H2. apply H3.
184 qed.
185
186 inductive bool : Set \def 
187   | true : bool
188   | false : bool.
189
190 definition notn : bool \to bool \def
191 \lambda b:bool. 
192  match b with 
193  [ true \Rightarrow false
194  | false \Rightarrow true ].
195
196 definition andb : bool \to bool \to bool\def
197 \lambda b1,b2:bool. 
198  match b1 with 
199  [ true \Rightarrow 
200         match b2 with [true \Rightarrow true | false \Rightarrow false]
201  | false \Rightarrow false ].
202
203 definition orb : bool \to bool \to bool\def
204 \lambda b1,b2:bool. 
205  match b1 with 
206  [ true \Rightarrow 
207         match b2 with [true \Rightarrow true | false \Rightarrow false]
208  | false \Rightarrow false ].
209
210 definition if_then_else : bool \to Prop \to Prop \to Prop \def 
211 \lambda b:bool.\lambda P,Q:Prop.
212 match b with
213 [ true \Rightarrow P
214 | false  \Rightarrow Q].
215
216 inductive le (n:nat) : nat \to Prop \def
217   | le_n : le n n
218   | le_S : \forall m:nat. le n m \to le n (S m).
219
220 theorem trans_le: \forall n,m,p:nat. le n m \to le m p \to le n p.
221 intros.
222 elim H1.assumption.
223 apply le_S.assumption.
224 qed.
225
226 theorem le_n_S: \forall n,m:nat. le n m \to le (S n) (S m).
227 intros.elim H.
228 apply le_n.apply le_S.assumption.
229 qed.
230
231 theorem le_O_n : \forall n:nat. le O n.
232 intros.elim n.apply le_n.apply le_S. assumption.
233 qed.
234
235 theorem le_n_Sn : \forall n:nat. le n (S n).
236 intros. apply le_S.apply le_n.
237 qed.
238
239 theorem le_pred_n : \forall n:nat. le (pred n) n.
240 intros.elim n.simplify.apply le_n.simplify.
241 apply le_n_Sn.
242 qed.
243
244 theorem not_zero_le : \forall n,m:nat. (le (S n) m ) \to not_zero m.
245 intros.elim H.exact I.exact I.
246 qed.
247
248 theorem le_Sn_O: \forall n:nat. Not (le (S n) O).
249 intros.simplify.intros.apply not_zero_le ? O H.
250 qed.
251
252 theorem le_n_O_eq : \forall n:nat. (le n O) \to (eq nat O n).
253 intros.cut (le n O) \to (eq nat O n).apply Hcut. assumption.
254 elim n.apply refl_equal.apply False_ind.apply  (le_Sn_O ? H2).
255 qed.
256
257 theorem le_S_n : \forall n,m:nat. le (S n) (S m) \to le n m.
258 intros.cut le (pred (S n)) (pred (S m)).exact Hcut.
259 elim H.apply le_n.apply trans_le ? (pred x).assumption.
260 apply le_pred_n.
261 qed.
262
263 theorem le_Sn_n : \forall n:nat. Not (le (S n) n).
264 intros.elim n.apply le_Sn_O.simplify.intros.
265 cut le (S e) e.apply H.assumption.apply le_S_n.assumption.
266 qed.
267
268 theorem le_antisym : \forall n,m:nat. (le n m) \to (le m n) \to (eq nat n m).
269 intros.cut (le n m) \to (le m n) \to (eq nat n m).exact Hcut H H1.
270 apply nat_double_ind (\lambda n,m.((le n m) \to (le m n) \to eq nat n m)).
271 intros.whd.intros.
272 apply le_n_O_eq.assumption.
273 intros.whd.intros.apply sym_eq.apply le_n_O_eq.assumption.
274 intros.whd.intros.apply f_equal.apply H2.
275 apply le_S_n.assumption.
276 apply le_S_n.assumption.
277 qed.
278
279 let rec leb n m \def 
280     match n with 
281     [ O \Rightarrow true
282     | (S p) \Rightarrow
283         match m with 
284         [ O \Rightarrow false
285         | (S q) \Rightarrow leb p q]].
286
287 theorem le_dec: \forall n,m:nat. if_then_else (leb n m) (le n m) (Not (le n m)).
288 intros.
289 apply (nat_double_ind 
290 (\lambda n,m:nat.if_then_else (leb n m) (le n m) (Not (le n m))) ? ? ? n m).
291 simplify.intros.apply le_O_n.
292 simplify.exact le_Sn_O.
293 intros 2.simplify.elim (leb n1 m1).
294 simplify.apply le_n_S.apply H.
295 simplify.intros.apply H.apply le_S_n.assumption.
296 qed.
297
298 inductive Z : Set \def
299   OZ : Z
300 | pos : nat \to Z
301 | neg : nat \to Z.
302
303 definition Z_of_nat \def
304 \lambda n. match n with
305 [ O \Rightarrow  OZ 
306 | (S n)\Rightarrow  pos n].
307
308 coercion Z_of_nat.
309
310 definition neg_Z_of_nat \def
311 \lambda n. match n with
312 [ O \Rightarrow  OZ 
313 | (S n)\Rightarrow  neg n].
314
315 definition absZ \def
316 \lambda z.
317  match z with 
318 [ OZ \Rightarrow O
319 | (pos n) \Rightarrow n
320 | (neg n) \Rightarrow n].
321
322 definition OZ_testb \def
323 \lambda z.
324 match z with 
325 [ OZ \Rightarrow true
326 | (pos n) \Rightarrow false
327 | (neg n) \Rightarrow false].
328
329 theorem OZ_discr :
330 \forall z. if_then_else (OZ_testb z) (eq Z z OZ) (Not (eq Z z OZ)). 
331 intros.elim z.simplify.apply refl_equal.
332 simplify.intros.
333 cut match neg e with 
334 [ OZ \Rightarrow True 
335 | (pos n) \Rightarrow False
336 | (neg n) \Rightarrow False].
337 apply Hcut. elim (sym_eq ? ? ? H).simplify.exact I.
338 simplify.intros.
339 cut match pos e with 
340 [ OZ \Rightarrow True 
341 | (pos n) \Rightarrow False
342 | (neg n) \Rightarrow False].
343 apply Hcut. elim (sym_eq ? ? ? H).simplify.exact I.
344 qed.
345
346 definition Zsucc \def
347 \lambda z. match z with
348 [ OZ \Rightarrow pos O
349 | (pos n) \Rightarrow pos (S n)
350 | (neg n) \Rightarrow 
351           match n with
352           [ O \Rightarrow OZ
353           | (S p) \Rightarrow neg p]].
354
355 definition Zpred \def
356 \lambda z. match z with
357 [ OZ \Rightarrow neg O
358 | (pos n) \Rightarrow 
359           match n with
360           [ O \Rightarrow OZ
361           | (S p) \Rightarrow pos p]
362 | (neg n) \Rightarrow neg (S n)].
363
364 theorem Zpred_succ: \forall z:Z. eq Z (Zpred (Zsucc z)) z.
365 intros.elim z.apply refl_equal.
366 elim e.apply refl_equal.
367 apply refl_equal.           
368 apply refl_equal.
369 qed.
370
371 theorem Zsucc_pred: \forall z:Z. eq Z (Zsucc (Zpred z)) z.
372 intros.elim z.apply refl_equal.
373 apply refl_equal.
374 elim e.apply refl_equal.
375 apply refl_equal.
376 qed.
377
378 inductive compare :Set \def
379 | LT : compare
380 | EQ : compare
381 | GT : compare.
382
383 let rec nat_compare n m: compare \def
384 match n with
385 [ O \Rightarrow 
386     match m with 
387       [ O \Rightarrow EQ
388       | (S q) \Rightarrow LT ]
389 | (S p) \Rightarrow 
390     match m with 
391       [ O \Rightarrow GT
392       | (S q) \Rightarrow nat_compare p q]].
393
394 definition compare_invert: compare \to compare \def
395   \lambda c.
396     match c with
397       [ LT \Rightarrow GT
398       | EQ \Rightarrow EQ
399       | GT \Rightarrow LT ].
400
401 theorem nat_compare_invert: \forall n,m:nat. 
402 eq compare (nat_compare n m) (compare_invert (nat_compare m n)).
403 intros. 
404 apply nat_double_ind (\lambda n,m.eq compare (nat_compare n m) (compare_invert (nat_compare m n))).
405 intros.elim n1.simplify.apply refl_equal.
406 simplify.apply refl_equal.
407 intro.elim n1.simplify.apply refl_equal.
408 simplify.apply refl_equal.
409 intros.simplify.elim H.apply refl_equal.
410 qed.
411
412 let rec Zplus x y : Z \def
413   match x with
414     [ OZ \Rightarrow y
415     | (pos m) \Rightarrow
416         match y with
417          [ OZ \Rightarrow x
418          | (pos n) \Rightarrow (pos (S (plus m n)))
419          | (neg n) \Rightarrow 
420               match nat_compare m n with
421                 [ LT \Rightarrow (neg (pred (minus n m)))
422                 | EQ \Rightarrow OZ
423                 | GT \Rightarrow (pos (pred (minus m n)))]]
424     | (neg m) \Rightarrow
425         match y with
426          [ OZ \Rightarrow x
427          | (pos n) \Rightarrow 
428               match nat_compare m n with
429                 [ LT \Rightarrow (pos (pred (minus n m)))
430                 | EQ \Rightarrow OZ
431                 | GT \Rightarrow (neg (pred (minus m n)))]     
432          | (neg n) \Rightarrow (neg (S (plus m n)))]].
433          
434 theorem Zplus_z_O:  \forall z:Z. eq Z (Zplus z OZ) z.
435 intro.elim z.
436 simplify.apply refl_equal.
437 simplify.apply refl_equal.
438 simplify.apply refl_equal.
439 qed.
440
441 theorem sym_Zplus : \forall x,y:Z. eq Z (Zplus x y) (Zplus y x).
442 intros.elim x.simplify.elim (sym_eq ? ? ? (Zplus_z_O y)).apply refl_equal.
443 elim y.simplify.reflexivity.
444 simplify.elim (sym_plus e e1).apply refl_equal.
445 simplify.elim (sym_eq ? ? ?(nat_compare_invert e e1)).
446 simplify.elim nat_compare e1 e.simplify.apply refl_equal.
447 simplify. apply refl_equal.
448 simplify. apply refl_equal.
449 elim y.simplify.reflexivity.
450 simplify.elim (sym_eq ? ? ?(nat_compare_invert e e1)).
451 simplify.elim nat_compare e1 e.simplify.apply refl_equal.
452 simplify. apply refl_equal.
453 simplify. apply refl_equal.
454 simplify.elim (sym_plus e1 e).apply refl_equal.
455 qed.
456
457 theorem Zpred_neg : \forall z:Z. eq Z (Zpred z) (Zplus (neg O) z).
458 intros.elim z.
459 simplify.apply refl_equal.
460 simplify.apply refl_equal.
461 elim e.simplify.apply refl_equal.
462 simplify.apply refl_equal.
463 qed.
464
465 theorem Zsucc_pos : \forall z:Z. eq Z (Zsucc z) (Zplus (pos O) z).
466 intros.elim z.
467 simplify.apply refl_equal.
468 elim e.simplify.apply refl_equal.
469 simplify.apply refl_equal.
470 simplify.apply refl_equal.
471 qed.
472
473 theorem Zplus_succ_pred_pp :
474 \forall n,m. eq Z (Zplus (pos n) (pos m)) (Zplus (Zsucc (pos n)) (Zpred (pos m))).
475 intros.
476 elim n.elim m.
477 simplify.apply refl_equal.
478 simplify.apply refl_equal.
479 elim m.
480 simplify.elim (plus_n_O ?).apply refl_equal.
481 simplify.elim (plus_n_Sm ? ?).apply refl_equal.
482 qed.
483
484 theorem Zplus_succ_pred_pn :
485 \forall n,m. eq Z (Zplus (pos n) (neg m)) (Zplus (Zsucc (pos n)) (Zpred (neg m))).
486 intros.apply refl_equal.
487 qed.
488
489 theorem Zplus_succ_pred_np :
490 \forall n,m. eq Z (Zplus (neg n) (pos m)) (Zplus (Zsucc (neg n)) (Zpred (pos m))).
491 intros.
492 elim n.elim m.
493 simplify.apply refl_equal.
494 simplify.apply refl_equal.
495 elim m.
496 simplify.apply refl_equal.
497 simplify.apply refl_equal.
498 qed.
499
500 theorem Zplus_succ_pred_nn:
501 \forall n,m. eq Z (Zplus (neg n) (neg m)) (Zplus (Zsucc (neg n)) (Zpred (neg m))).
502 intros.
503 elim n.elim m.
504 simplify.apply refl_equal.
505 simplify.apply refl_equal.
506 elim m.
507 simplify.elim (plus_n_Sm ? ?).apply refl_equal.
508 simplify.elim (sym_eq ? ? ? (plus_n_Sm ? ?)).apply refl_equal.
509 qed.
510
511 theorem Zplus_succ_pred:
512 \forall x,y. eq Z (Zplus x y) (Zplus (Zsucc x) (Zpred y)).
513 intros.
514 elim x. elim y.
515 simplify.apply refl_equal.
516 simplify.apply refl_equal.
517 elim (Zsucc_pos ?).elim (sym_eq ? ? ? (Zsucc_pred ?)).apply refl_equal.
518 elim y.elim sym_Zplus ? ?.elim sym_Zplus (Zpred OZ) ?.
519 elim (Zpred_neg ?).elim (sym_eq ? ? ? (Zpred_succ ?)).
520 simplify.apply refl_equal.
521 apply Zplus_succ_pred_nn.
522 apply Zplus_succ_pred_np.
523 elim y.simplify.apply refl_equal.
524 apply Zplus_succ_pred_pn.
525 apply Zplus_succ_pred_pp.
526 qed.
527
528 theorem Zsucc_plus_pp : 
529 \forall n,m. eq Z (Zplus (Zsucc (pos n)) (pos m)) (Zsucc (Zplus (pos n) (pos m))).
530 intros.apply refl_equal.
531 qed.
532
533 theorem Zsucc_plus_pn : 
534 \forall n,m. eq Z (Zplus (Zsucc (pos n)) (neg m)) (Zsucc (Zplus (pos n) (neg m))).
535 intros.
536 apply nat_double_ind
537 (\lambda n,m. eq Z (Zplus (Zsucc (pos n)) (neg m)) (Zsucc (Zplus (pos n) (neg m)))).intro.
538 intros.elim n1.
539 simplify. apply refl_equal.
540 elim e.simplify. apply refl_equal.
541 simplify. apply refl_equal.
542 intros. elim n1.
543 simplify. apply refl_equal.
544 simplify.apply refl_equal.
545 intros.
546 elim (Zplus_succ_pred_pn ? m1).
547 elim H.apply refl_equal.
548 qed.
549
550 theorem Zsucc_plus_nn : 
551 \forall n,m. eq Z (Zplus (Zsucc (neg n)) (neg m)) (Zsucc (Zplus (neg n) (neg m))).
552 intros.
553 apply nat_double_ind
554 (\lambda n,m. eq Z (Zplus (Zsucc (neg n)) (neg m)) (Zsucc (Zplus (neg n) (neg m)))).intro.
555 intros.elim n1.
556 simplify. apply refl_equal.
557 elim e.simplify. apply refl_equal.
558 simplify. apply refl_equal.
559 intros. elim n1.
560 simplify. apply refl_equal.
561 simplify.apply refl_equal.
562 intros.
563 elim (Zplus_succ_pred_nn ? m1).
564 apply refl_equal.
565 qed.
566
567 theorem Zsucc_plus_np : 
568 \forall n,m. eq Z (Zplus (Zsucc (neg n)) (pos m)) (Zsucc (Zplus (neg n) (pos m))).
569 intros.
570 apply nat_double_ind
571 (\lambda n,m. eq Z (Zplus (Zsucc (neg n)) (pos m)) (Zsucc (Zplus (neg n) (pos m)))).
572 intros.elim n1.
573 simplify. apply refl_equal.
574 elim e.simplify. apply refl_equal.
575 simplify. apply refl_equal.
576 intros. elim n1.
577 simplify. apply refl_equal.
578 simplify.apply refl_equal.
579 intros.
580 elim H.
581 elim (Zplus_succ_pred_np ? (S m1)).
582 apply refl_equal.
583 qed.
584
585
586 theorem Zsucc_plus : \forall x,y:Z. eq Z (Zplus (Zsucc x) y) (Zsucc (Zplus x y)).
587 intros.elim x.elim y.
588 simplify. apply refl_equal.
589 elim (Zsucc_pos ?).apply refl_equal.
590 simplify.apply refl_equal.
591 elim y.elim sym_Zplus ? ?.elim sym_Zplus OZ ?.simplify.apply refl_equal.
592 apply Zsucc_plus_nn.
593 apply Zsucc_plus_np.
594 elim y.
595 elim (sym_Zplus OZ ?).apply refl_equal.
596 apply Zsucc_plus_pn.
597 apply Zsucc_plus_pp.
598 qed.
599
600 theorem Zpred_plus : \forall x,y:Z. eq Z (Zplus (Zpred x) y) (Zpred (Zplus x y)).
601 intros.
602 cut eq Z (Zpred (Zplus x y)) (Zpred (Zplus (Zsucc (Zpred x)) y)).
603 elim (sym_eq ? ? ? Hcut).
604 elim (sym_eq ? ? ? (Zsucc_plus ? ?)).
605 elim (sym_eq ? ? ? (Zpred_succ ?)).
606 apply refl_equal.
607 elim (sym_eq ? ? ? (Zsucc_pred ?)).
608 apply refl_equal.
609 qed.
610
611 theorem assoc_Zplus : 
612 \forall x,y,z:Z. eq Z (Zplus x (Zplus y z)) (Zplus (Zplus x y) z).
613 intros.elim x.simplify.apply refl_equal.
614 elim e.elim (Zpred_neg (Zplus y z)).
615 elim (Zpred_neg y).
616 elim (Zpred_plus ? ?).
617 apply refl_equal.
618 elim (sym_eq ? ? ? (Zpred_plus (neg e1) ?)).
619 elim (sym_eq ? ? ? (Zpred_plus (neg e1) ?)).
620 elim (sym_eq ? ? ? (Zpred_plus (Zplus (neg e1) y) ?)).
621 apply f_equal.assumption.
622 elim e.elim (Zsucc_pos ?).
623 elim (Zsucc_pos ?).
624 apply (sym_eq ? ? ? (Zsucc_plus ? ?)) .
625 elim (sym_eq ? ? ? (Zsucc_plus (pos e1) ?)).
626 elim (sym_eq ? ? ? (Zsucc_plus (pos e1) ?)).
627 elim (sym_eq ? ? ? (Zsucc_plus (Zplus (pos e1) y) ?)).
628 apply f_equal.assumption.
629 qed.
630
631
632