7095
|
1 |
(* Title: Sequents/LK/Nat
|
7091
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
7095
|
4 |
Copyright 1999 University of Cambridge
|
7091
|
5 |
|
7095
|
6 |
Theory of the natural numbers: Peano's axioms, primitive recursion
|
7091
|
7 |
*)
|
|
8 |
|
|
9 |
Addsimps [Suc_neq_0];
|
|
10 |
|
|
11 |
Add_safes [Suc_inject RS L_of_imp];
|
|
12 |
|
|
13 |
Goal "|- Suc(k) ~= k";
|
|
14 |
by (res_inst_tac [("n","k")] induct 1);
|
|
15 |
by (Simp_tac 1);
|
|
16 |
by (Fast_tac 1);
|
|
17 |
qed "Suc_n_not_n";
|
|
18 |
|
|
19 |
Goalw [add_def] "|- 0+n = n";
|
|
20 |
by (rtac rec_0 1);
|
|
21 |
qed "add_0";
|
|
22 |
|
|
23 |
Goalw [add_def] "|- Suc(m)+n = Suc(m+n)";
|
|
24 |
by (rtac rec_Suc 1);
|
|
25 |
qed "add_Suc";
|
|
26 |
|
|
27 |
Addsimps [add_0, add_Suc];
|
|
28 |
|
|
29 |
Goal "|- (k+m)+n = k+(m+n)";
|
|
30 |
by (res_inst_tac [("n","k")] induct 1);
|
|
31 |
by (Simp_tac 1);
|
|
32 |
by (Asm_simp_tac 1);
|
|
33 |
qed "add_assoc";
|
|
34 |
|
|
35 |
Goal "|- m+0 = m";
|
|
36 |
by (res_inst_tac [("n","m")] induct 1);
|
|
37 |
by (Simp_tac 1);
|
|
38 |
by (Asm_simp_tac 1);
|
|
39 |
qed "add_0_right";
|
|
40 |
|
|
41 |
Goal "|- m+Suc(n) = Suc(m+n)";
|
|
42 |
by (res_inst_tac [("n","m")] induct 1);
|
|
43 |
by (ALLGOALS (Asm_simp_tac));
|
|
44 |
qed "add_Suc_right";
|
|
45 |
|
|
46 |
(*Example used in Reference Manual, Doc/Ref/simplifier.tex*)
|
|
47 |
val [prem] = Goal "(!!n. |- f(Suc(n)) = Suc(f(n))) ==> |- f(i+j) = i+f(j)";
|
|
48 |
by (res_inst_tac [("n","i")] induct 1);
|
|
49 |
by (Simp_tac 1);
|
|
50 |
by (simp_tac (simpset() addsimps [prem]) 1);
|
|
51 |
result();
|