src/Sequents/LK/Nat.ML
author wenzelm
Tue, 30 May 2000 16:08:38 +0200
changeset 9000 c20d58286a51
parent 7095 cfc11af6174a
child 17481 75166ebb619b
permissions -rw-r--r--
cleaned up;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7095
cfc11af6174a fixed the comments...
paulson
parents: 7091
diff changeset
     1
(*  Title:      Sequents/LK/Nat
7091
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
     2
    ID:         $Id$
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
7095
cfc11af6174a fixed the comments...
paulson
parents: 7091
diff changeset
     4
    Copyright   1999  University of Cambridge
7091
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
     5
7095
cfc11af6174a fixed the comments...
paulson
parents: 7091
diff changeset
     6
Theory of the natural numbers: Peano's axioms, primitive recursion
7091
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
     7
*)
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
     8
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
     9
Addsimps [Suc_neq_0];
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    10
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    11
Add_safes [Suc_inject RS L_of_imp];
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    12
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    13
Goal "|- Suc(k) ~= k";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    14
by (res_inst_tac [("n","k")] induct 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    15
by (Simp_tac 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    16
by (Fast_tac 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    17
qed "Suc_n_not_n";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    18
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    19
Goalw [add_def] "|- 0+n = n";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    20
by (rtac rec_0 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    21
qed "add_0";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    22
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    23
Goalw [add_def] "|- Suc(m)+n = Suc(m+n)";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    24
by (rtac rec_Suc 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    25
qed "add_Suc";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    26
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    27
Addsimps [add_0, add_Suc];
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    28
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    29
Goal "|- (k+m)+n = k+(m+n)";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    30
by (res_inst_tac [("n","k")] induct 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    31
by (Simp_tac 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    32
by (Asm_simp_tac 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    33
qed "add_assoc";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    34
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    35
Goal "|- m+0 = m";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    36
by (res_inst_tac [("n","m")] induct 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    37
by (Simp_tac 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    38
by (Asm_simp_tac 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    39
qed "add_0_right";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    40
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    41
Goal "|- m+Suc(n) = Suc(m+n)";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    42
by (res_inst_tac [("n","m")] induct 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    43
by (ALLGOALS (Asm_simp_tac));
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    44
qed "add_Suc_right";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    45
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    46
(*Example used in Reference Manual, Doc/Ref/simplifier.tex*)
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    47
val [prem] = Goal "(!!n. |- f(Suc(n)) = Suc(f(n))) ==> |- f(i+j) = i+f(j)";
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    48
by (res_inst_tac [("n","i")] induct 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    49
by (Simp_tac 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    50
by (simp_tac (simpset() addsimps [prem]) 1);
b76a26835a5c Sequents/LK/Nat: new example of simplification in LK
paulson
parents:
diff changeset
    51
result();