| author | blanchet | 
| Mon, 17 Feb 2014 13:31:42 +0100 | |
| changeset 55532 | b751e6d7f4e9 | 
| parent 55380 | 4de48353034e | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 17481 | 1  | 
(* Title: Sequents/LK/Nat.thy  | 
| 7091 | 2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 7095 | 3  | 
Copyright 1999 University of Cambridge  | 
| 7091 | 4  | 
*)  | 
5  | 
||
| 17481 | 6  | 
header {* Theory of the natural numbers: Peano's axioms, primitive recursion *}
 | 
7  | 
||
8  | 
theory Nat  | 
|
| 55229 | 9  | 
imports "../LK"  | 
| 17481 | 10  | 
begin  | 
11  | 
||
12  | 
typedecl nat  | 
|
| 
55380
 
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
 
wenzelm 
parents: 
55230 
diff
changeset
 | 
13  | 
instance nat :: "term" ..  | 
| 7091 | 14  | 
|
| 51309 | 15  | 
axiomatization  | 
16  | 
  Zero :: nat      ("0") and
 | 
|
17  | 
Suc :: "nat=>nat" and  | 
|
18  | 
rec :: "[nat, 'a, [nat,'a]=>'a] => 'a"  | 
|
19  | 
where  | 
|
| 17481 | 20  | 
induct: "[| $H |- $E, P(0), $F;  | 
| 51309 | 21  | 
!!x. $H, P(x) |- $E, P(Suc(x)), $F |] ==> $H |- $E, P(n), $F" and  | 
| 7095 | 22  | 
|
| 51309 | 23  | 
Suc_inject: "|- Suc(m)=Suc(n) --> m=n" and  | 
24  | 
Suc_neq_0: "|- Suc(m) ~= 0" and  | 
|
25  | 
rec_0: "|- rec(0,a,f) = a" and  | 
|
| 17481 | 26  | 
rec_Suc: "|- rec(Suc(m), a, f) = f(m, rec(m,a,f))"  | 
| 51309 | 27  | 
|
28  | 
definition add :: "[nat, nat] => nat" (infixl "+" 60)  | 
|
29  | 
where "m + n == rec(m, n, %x y. Suc(y))"  | 
|
| 17481 | 30  | 
|
| 21426 | 31  | 
|
32  | 
declare Suc_neq_0 [simp]  | 
|
33  | 
||
34  | 
lemma Suc_inject_rule: "$H, $G, m = n |- $E \<Longrightarrow> $H, Suc(m) = Suc(n), $G |- $E"  | 
|
35  | 
by (rule L_of_imp [OF Suc_inject])  | 
|
36  | 
||
37  | 
lemma Suc_n_not_n: "|- Suc(k) ~= k"  | 
|
38  | 
apply (rule_tac n = k in induct)  | 
|
| 55230 | 39  | 
apply simp  | 
| 55228 | 40  | 
apply (fast add!: Suc_inject_rule)  | 
| 21426 | 41  | 
done  | 
42  | 
||
43  | 
lemma add_0: "|- 0+n = n"  | 
|
44  | 
apply (unfold add_def)  | 
|
45  | 
apply (rule rec_0)  | 
|
46  | 
done  | 
|
47  | 
||
48  | 
lemma add_Suc: "|- Suc(m)+n = Suc(m+n)"  | 
|
49  | 
apply (unfold add_def)  | 
|
50  | 
apply (rule rec_Suc)  | 
|
51  | 
done  | 
|
52  | 
||
53  | 
declare add_0 [simp] add_Suc [simp]  | 
|
54  | 
||
55  | 
lemma add_assoc: "|- (k+m)+n = k+(m+n)"  | 
|
56  | 
apply (rule_tac n = "k" in induct)  | 
|
| 55230 | 57  | 
apply simp_all  | 
| 21426 | 58  | 
done  | 
59  | 
||
60  | 
lemma add_0_right: "|- m+0 = m"  | 
|
61  | 
apply (rule_tac n = "m" in induct)  | 
|
| 55230 | 62  | 
apply simp_all  | 
| 21426 | 63  | 
done  | 
64  | 
||
65  | 
lemma add_Suc_right: "|- m+Suc(n) = Suc(m+n)"  | 
|
66  | 
apply (rule_tac n = "m" in induct)  | 
|
| 55230 | 67  | 
apply simp_all  | 
| 21426 | 68  | 
done  | 
69  | 
||
70  | 
lemma "(!!n. |- f(Suc(n)) = Suc(f(n))) ==> |- f(i+j) = i+f(j)"  | 
|
71  | 
apply (rule_tac n = "i" in induct)  | 
|
| 55230 | 72  | 
apply simp_all  | 
| 21426 | 73  | 
done  | 
| 17481 | 74  | 
|
| 7091 | 75  | 
end  |