src/FOL/ex/NatClass.ML
author clasohm
Fri, 13 Oct 1995 11:48:40 +0100
changeset 1280 909079af97b7
parent 1246 706cfddca75c
child 1459 d12da312eff4
permissions -rw-r--r--
corrected spelling of title
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1246
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
     1
(*  Title: 	FOL/ex/NatClass.ML
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
     5
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
     6
This is Nat.ML trivially modified to make it work with NatClass.thy.
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
     7
*)
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
     8
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
     9
open NatClass;
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    10
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    11
goal NatClass.thy "Suc(k) ~= (k::'a::nat)";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    12
by (res_inst_tac [("n","k")] induct 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    13
by (resolve_tac [notI] 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    14
by (eresolve_tac [Suc_neq_0] 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    15
by (resolve_tac [notI] 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    16
by (eresolve_tac [notE] 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    17
by (eresolve_tac [Suc_inject] 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    18
qed "Suc_n_not_n";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    19
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    20
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    21
goal NatClass.thy "(k+m)+n = k+(m+n)";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    22
prths ([induct] RL [topthm()]);  (*prints all 14 next states!*)
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    23
by (resolve_tac [induct] 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    24
back();
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    25
back();
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    26
back();
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    27
back();
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    28
back();
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    29
back();
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    30
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    31
goalw NatClass.thy [add_def] "0+n = n";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    32
by (resolve_tac [rec_0] 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    33
qed "add_0";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    34
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    35
goalw NatClass.thy [add_def] "Suc(m)+n = Suc(m+n)";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    36
by (resolve_tac [rec_Suc] 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    37
qed "add_Suc";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    38
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    39
val add_ss = FOL_ss addsimps [add_0, add_Suc];
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    40
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    41
goal NatClass.thy "(k+m)+n = k+(m+n)";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    42
by (res_inst_tac [("n","k")] induct 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    43
by (simp_tac add_ss 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    44
by (asm_simp_tac add_ss 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    45
qed "add_assoc";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    46
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    47
goal NatClass.thy "m+0 = m";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    48
by (res_inst_tac [("n","m")] induct 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    49
by (simp_tac add_ss 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    50
by (asm_simp_tac add_ss 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    51
qed "add_0_right";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    52
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    53
goal NatClass.thy "m+Suc(n) = Suc(m+n)";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    54
by (res_inst_tac [("n","m")] induct 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    55
by (ALLGOALS (asm_simp_tac add_ss));
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    56
qed "add_Suc_right";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    57
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    58
val [prem] = goal NatClass.thy "(!!n. f(Suc(n)) = Suc(f(n))) ==> f(i+j) = i+f(j)";
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    59
by (res_inst_tac [("n","i")] induct 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    60
by (simp_tac add_ss 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    61
by (asm_simp_tac (add_ss addsimps [prem]) 1);
706cfddca75c the NatClass demo of the axclass tutorial;
wenzelm
parents:
diff changeset
    62
result();