doc-src/AxClass/Nat/NatClass.ML
author wenzelm
Sat, 26 Jan 2008 20:01:37 +0100
changeset 25985 8d69087f6a4b
parent 24584 01e83ffa6c54
permissions -rw-r--r--
avoid redundant escaping of Isabelle symbols;

(*  Title:      Doc/AxClass/Nat/NatClass.ML
    ID:         $Id$
    Author:     Markus Wenzel, TU Muenchen

This is Nat.ML with some trivial modifications in order to make it
work with NatClass.thy.
*)

val induct = thm "induct";
val Suc_inject = thm "Suc_inject";
val Suc_neq_0 = thm "Suc_neq_0";
val rec_0 = thm "rec_0";
val rec_Suc = thm "rec_Suc";
val add_def = thm "add_def";


Goal "Suc(k) ~= (k::'a::nat)";
by (res_inst_tac [("n","k")] induct 1);
by (rtac notI 1);
by (etac Suc_neq_0 1);
by (rtac notI 1);
by (etac notE 1);
by (etac Suc_inject 1);
qed "Suc_n_not_n";


Goal "(k+m)+n = k+(m+n)";
prths ([induct] RL [topthm()]);  (*prints all 14 next states!*)
by (rtac induct 1);
back();
back();
back();
back();
back();
back();

Goalw [add_def] "\<zero>+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+\<zero> = 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";

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 (asm_simp_tac (simpset() addsimps [prem]) 1);
qed "";