Sequents/LK/Nat: new example of simplification in LK
authorpaulson
Tue, 27 Jul 1999 18:50:14 +0200
changeset 7091 b76a26835a5c
parent 7090 dd30a72880c7
child 7092 d7958f38e9e0
Sequents/LK/Nat: new example of simplification in LK
src/Sequents/LK/Nat.ML
src/Sequents/LK/Nat.thy
src/Sequents/LK/ROOT.ML
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Sequents/LK/Nat.ML	Tue Jul 27 18:50:14 1999 +0200
@@ -0,0 +1,51 @@
+(*  Title:      FOL/ex/Nat.ML
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1992  University of Cambridge
+
+Proofs about the natural numbers.
+*)
+
+Addsimps [Suc_neq_0];
+
+Add_safes [Suc_inject RS L_of_imp];
+
+Goal "|- Suc(k) ~= k";
+by (res_inst_tac [("n","k")] induct 1);
+by (Simp_tac 1);
+by (Fast_tac 1);
+qed "Suc_n_not_n";
+
+Goalw [add_def] "|- 0+n = n";
+by (rtac rec_0 1);
+qed "add_0";
+
+Goalw [add_def] "|- Suc(m)+n = Suc(m+n)";
+by (rtac rec_Suc 1);
+qed "add_Suc";
+
+Addsimps [add_0, add_Suc];
+
+Goal "|- (k+m)+n = k+(m+n)";
+by (res_inst_tac [("n","k")] induct 1);
+by (Simp_tac 1);
+by (Asm_simp_tac 1);
+qed "add_assoc";
+
+Goal "|- m+0 = m";
+by (res_inst_tac [("n","m")] induct 1);
+by (Simp_tac 1);
+by (Asm_simp_tac 1);
+qed "add_0_right";
+
+Goal "|- m+Suc(n) = Suc(m+n)";
+by (res_inst_tac [("n","m")] induct 1);
+by (ALLGOALS (Asm_simp_tac));
+qed "add_Suc_right";
+
+(*Example used in Reference Manual, Doc/Ref/simplifier.tex*)
+val [prem] = Goal "(!!n. |- f(Suc(n)) = Suc(f(n))) ==> |- f(i+j) = i+f(j)";
+by (res_inst_tac [("n","i")] induct 1);
+by (Simp_tac 1);
+by (simp_tac (simpset() addsimps [prem]) 1);
+result();
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Sequents/LK/Nat.thy	Tue Jul 27 18:50:14 1999 +0200
@@ -0,0 +1,27 @@
+(*  Title:      FOL/ex/Nat.thy
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1992  University of Cambridge
+
+Examples for the manuals.
+
+Theory of the natural numbers: Peano's axioms, primitive recursion
+*)
+
+Nat = LK +
+types   nat
+arities nat :: term
+consts  "0" :: nat      ("0")
+        Suc :: nat=>nat  
+        rec :: [nat, 'a, [nat,'a]=>'a] => 'a  
+        "+" :: [nat, nat] => nat                (infixl 60)
+
+rules   
+  induct      "[| $H |- $E, P(0), $F;
+		  !!x. $H |- $E, P(x) --> P(Suc(x)), $F |] ==> $H |- $E, P(n), $F"
+  Suc_inject  "|- Suc(m)=Suc(n) --> m=n"
+  Suc_neq_0   "|- Suc(m) ~= 0"
+  rec_0       "|- rec(0,a,f) = a"
+  rec_Suc     "|- rec(Suc(m), a, f) = f(m, rec(m,a,f))"
+  add_def     "m+n == rec(m, n, %x y. Suc(y))"
+end
--- a/src/Sequents/LK/ROOT.ML	Tue Jul 27 17:30:25 1999 +0200
+++ b/src/Sequents/LK/ROOT.ML	Tue Jul 27 18:50:14 1999 +0200
@@ -12,3 +12,4 @@
 time_use "prop.ML";
 time_use "quant.ML";
 time_use "hardquant.ML";
+use_thy "Nat";