src/Sequents/LK/Nat.thy
author wenzelm
Sat, 01 Feb 2014 18:00:28 +0100
changeset 55229 08f2ebb65078
parent 55228 901a6696cdd8
child 55230 cb5ef74b32f9
permissions -rw-r--r--
simplified sessions;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
     1
(*  Title:      Sequents/LK/Nat.thy
7091
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
7095
cfc11af6174a fixed the comments...
paulson
parents: 7091
diff changeset
     3
    Copyright   1999  University of Cambridge
7091
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
     4
*)
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
     5
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
     6
header {* Theory of the natural numbers: Peano's axioms, primitive recursion *}
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
     7
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
     8
theory Nat
55229
08f2ebb65078 simplified sessions;
wenzelm
parents: 55228
diff changeset
     9
imports "../LK"
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
    10
begin
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
    11
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
    12
typedecl nat
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
    13
arities nat :: "term"
7091
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    14
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    15
axiomatization
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    16
  Zero :: nat      ("0") and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    17
  Suc :: "nat=>nat" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    18
  rec :: "[nat, 'a, [nat,'a]=>'a] => 'a"
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    19
where
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
    20
  induct:  "[| $H |- $E, P(0), $F;
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    21
              !!x. $H, P(x) |- $E, P(Suc(x)), $F |] ==> $H |- $E, P(n), $F" and
7095
cfc11af6174a fixed the comments...
paulson
parents: 7091
diff changeset
    22
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    23
  Suc_inject:  "|- Suc(m)=Suc(n) --> m=n" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    24
  Suc_neq_0:   "|- Suc(m) ~= 0" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    25
  rec_0:       "|- rec(0,a,f) = a" and
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
    26
  rec_Suc:     "|- rec(Suc(m), a, f) = f(m, rec(m,a,f))"
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    27
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    28
definition add :: "[nat, nat] => nat"  (infixl "+" 60)
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 41310
diff changeset
    29
  where "m + n == rec(m, n, %x y. Suc(y))"
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
    30
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    31
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    32
declare Suc_neq_0 [simp]
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    33
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    34
lemma Suc_inject_rule: "$H, $G, m = n |- $E \<Longrightarrow> $H, Suc(m) = Suc(n), $G |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    35
  by (rule L_of_imp [OF Suc_inject])
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    36
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    37
lemma Suc_n_not_n: "|- Suc(k) ~= k"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    38
  apply (rule_tac n = k in induct)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51309
diff changeset
    39
  apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms Suc_neq_0}) 1 *})
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 51717
diff changeset
    40
  apply (fast add!: Suc_inject_rule)
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    41
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    42
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    43
lemma add_0: "|- 0+n = n"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    44
  apply (unfold add_def)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    45
  apply (rule rec_0)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    46
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    47
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    48
lemma add_Suc: "|- Suc(m)+n = Suc(m+n)"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    49
  apply (unfold add_def)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    50
  apply (rule rec_Suc)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    51
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    52
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    53
declare add_0 [simp] add_Suc [simp]
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    54
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    55
lemma add_assoc: "|- (k+m)+n = k+(m+n)"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    56
  apply (rule_tac n = "k" in induct)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51309
diff changeset
    57
  apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_0}) 1 *})
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51309
diff changeset
    58
  apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_Suc}) 1 *})
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    59
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    60
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    61
lemma add_0_right: "|- m+0 = m"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    62
  apply (rule_tac n = "m" in induct)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51309
diff changeset
    63
  apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_0}) 1 *})
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51309
diff changeset
    64
  apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_Suc}) 1 *})
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    65
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    66
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    67
lemma add_Suc_right: "|- m+Suc(n) = Suc(m+n)"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    68
  apply (rule_tac n = "m" in induct)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51309
diff changeset
    69
  apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_0}) 1 *})
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51309
diff changeset
    70
  apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_Suc}) 1 *})
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    71
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    72
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    73
lemma "(!!n. |- f(Suc(n)) = Suc(f(n))) ==> |- f(i+j) = i+f(j)"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    74
  apply (rule_tac n = "i" in induct)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51309
diff changeset
    75
  apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_0}) 1 *})
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51309
diff changeset
    76
  apply (tactic {* asm_simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_Suc}) 1 *})
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    77
  done
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 7123
diff changeset
    78
7091
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    79
end