| author | wenzelm | 
| Wed, 12 Feb 2025 14:26:43 +0100 | |
| changeset 82146 | 143ff9527bac | 
| parent 80914 | d97fdabd9e2b | 
| 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 | ||
| 60770 | 6 | section \<open>Theory of the natural numbers: Peano's axioms, primitive recursion\<close> | 
| 17481 | 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: 
55230diff
changeset | 13 | instance nat :: "term" .. | 
| 7091 | 14 | |
| 51309 | 15 | axiomatization | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
61386diff
changeset | 16 | Zero :: nat (\<open>0\<close>) and | 
| 61385 | 17 | Suc :: "nat \<Rightarrow> nat" and | 
| 18 | rec :: "[nat, 'a, [nat,'a] \<Rightarrow> 'a] \<Rightarrow> 'a" | |
| 51309 | 19 | where | 
| 61386 | 20 | induct: "\<lbrakk>$H \<turnstile> $E, P(0), $F; | 
| 21 | \<And>x. $H, P(x) \<turnstile> $E, P(Suc(x)), $F\<rbrakk> \<Longrightarrow> $H \<turnstile> $E, P(n), $F" and | |
| 7095 | 22 | |
| 61386 | 23 | Suc_inject: "\<turnstile> Suc(m) = Suc(n) \<longrightarrow> m = n" and | 
| 24 | Suc_neq_0: "\<turnstile> Suc(m) \<noteq> 0" and | |
| 25 | rec_0: "\<turnstile> rec(0,a,f) = a" and | |
| 26 | rec_Suc: "\<turnstile> rec(Suc(m), a, f) = f(m, rec(m,a,f))" | |
| 51309 | 27 | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
61386diff
changeset | 28 | definition add :: "[nat, nat] \<Rightarrow> nat" (infixl \<open>+\<close> 60) | 
| 61385 | 29 | where "m + n == rec(m, n, \<lambda>x y. Suc(y))" | 
| 17481 | 30 | |
| 21426 | 31 | |
| 32 | declare Suc_neq_0 [simp] | |
| 33 | ||
| 61386 | 34 | lemma Suc_inject_rule: "$H, $G, m = n \<turnstile> $E \<Longrightarrow> $H, Suc(m) = Suc(n), $G \<turnstile> $E" | 
| 21426 | 35 | by (rule L_of_imp [OF Suc_inject]) | 
| 36 | ||
| 61386 | 37 | lemma Suc_n_not_n: "\<turnstile> Suc(k) \<noteq> k" | 
| 21426 | 38 | apply (rule_tac n = k in induct) | 
| 55230 | 39 | apply simp | 
| 55228 | 40 | apply (fast add!: Suc_inject_rule) | 
| 21426 | 41 | done | 
| 42 | ||
| 61386 | 43 | lemma add_0: "\<turnstile> 0 + n = n" | 
| 21426 | 44 | apply (unfold add_def) | 
| 45 | apply (rule rec_0) | |
| 46 | done | |
| 47 | ||
| 61386 | 48 | lemma add_Suc: "\<turnstile> Suc(m) + n = Suc(m + n)" | 
| 21426 | 49 | apply (unfold add_def) | 
| 50 | apply (rule rec_Suc) | |
| 51 | done | |
| 52 | ||
| 53 | declare add_0 [simp] add_Suc [simp] | |
| 54 | ||
| 61386 | 55 | lemma add_assoc: "\<turnstile> (k + m) + n = k + (m + n)" | 
| 21426 | 56 | apply (rule_tac n = "k" in induct) | 
| 55230 | 57 | apply simp_all | 
| 21426 | 58 | done | 
| 59 | ||
| 61386 | 60 | lemma add_0_right: "\<turnstile> m + 0 = m" | 
| 21426 | 61 | apply (rule_tac n = "m" in induct) | 
| 55230 | 62 | apply simp_all | 
| 21426 | 63 | done | 
| 64 | ||
| 61386 | 65 | lemma add_Suc_right: "\<turnstile> m + Suc(n) = Suc(m + n)" | 
| 21426 | 66 | apply (rule_tac n = "m" in induct) | 
| 55230 | 67 | apply simp_all | 
| 21426 | 68 | done | 
| 69 | ||
| 61386 | 70 | lemma "(\<And>n. \<turnstile> f(Suc(n)) = Suc(f(n))) \<Longrightarrow> \<turnstile> f(i + j) = i + f(j)" | 
| 21426 | 71 | apply (rule_tac n = "i" in induct) | 
| 55230 | 72 | apply simp_all | 
| 21426 | 73 | done | 
| 17481 | 74 | |
| 7091 | 75 | end |