1459
|
1 |
(* Title: FOLP/ex/nat.ML
|
0
|
2 |
ID: $Id$
|
1459
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
0
|
4 |
Copyright 1992 University of Cambridge
|
|
5 |
|
|
6 |
Examples for the manual "Introduction to Isabelle"
|
|
7 |
|
|
8 |
Proofs about the natural numbers
|
|
9 |
|
|
10 |
To generate similar output to manual, execute these commands:
|
|
11 |
Pretty.setmargin 72; print_depth 0;
|
|
12 |
*)
|
|
13 |
|
|
14 |
open Nat;
|
|
15 |
|
5061
|
16 |
Goal "?p : ~ (Suc(k) = k)";
|
0
|
17 |
by (res_inst_tac [("n","k")] induct 1);
|
|
18 |
by (rtac notI 1);
|
|
19 |
by (etac Suc_neq_0 1);
|
|
20 |
by (rtac notI 1);
|
|
21 |
by (etac notE 1);
|
|
22 |
by (etac Suc_inject 1);
|
|
23 |
val Suc_n_not_n = result();
|
|
24 |
|
|
25 |
|
5061
|
26 |
Goal "?p : (k+m)+n = k+(m+n)";
|
0
|
27 |
prths ([induct] RL [topthm()]); (*prints all 14 next states!*)
|
|
28 |
by (rtac induct 1);
|
|
29 |
back();
|
|
30 |
back();
|
|
31 |
back();
|
|
32 |
back();
|
|
33 |
back();
|
|
34 |
back();
|
|
35 |
|
5061
|
36 |
Goalw [add_def] "?p : 0+n = n";
|
0
|
37 |
by (rtac rec_0 1);
|
|
38 |
val add_0 = result();
|
|
39 |
|
5061
|
40 |
Goalw [add_def] "?p : Suc(m)+n = Suc(m+n)";
|
0
|
41 |
by (rtac rec_Suc 1);
|
|
42 |
val add_Suc = result();
|
|
43 |
|
|
44 |
(*
|
|
45 |
val nat_congs = mk_congs Nat.thy ["Suc", "op +"];
|
|
46 |
prths nat_congs;
|
|
47 |
*)
|
|
48 |
val prems = goal Nat.thy "p: x=y ==> ?p : Suc(x) = Suc(y)";
|
|
49 |
by (resolve_tac (prems RL [subst]) 1);
|
|
50 |
by (rtac refl 1);
|
|
51 |
val Suc_cong = result();
|
|
52 |
|
|
53 |
val prems = goal Nat.thy "[| p : a=x; q: b=y |] ==> ?p : a+b=x+y";
|
|
54 |
by (resolve_tac (prems RL [subst]) 1 THEN resolve_tac (prems RL [subst]) 1 THEN
|
|
55 |
rtac refl 1);
|
|
56 |
val Plus_cong = result();
|
|
57 |
|
|
58 |
val nat_congs = [Suc_cong,Plus_cong];
|
|
59 |
|
|
60 |
|
|
61 |
val add_ss = FOLP_ss addcongs nat_congs
|
1459
|
62 |
addrews [add_0, add_Suc];
|
0
|
63 |
|
5061
|
64 |
Goal "?p : (k+m)+n = k+(m+n)";
|
0
|
65 |
by (res_inst_tac [("n","k")] induct 1);
|
|
66 |
by (SIMP_TAC add_ss 1);
|
|
67 |
by (ASM_SIMP_TAC add_ss 1);
|
|
68 |
val add_assoc = result();
|
|
69 |
|
5061
|
70 |
Goal "?p : m+0 = m";
|
0
|
71 |
by (res_inst_tac [("n","m")] induct 1);
|
|
72 |
by (SIMP_TAC add_ss 1);
|
|
73 |
by (ASM_SIMP_TAC add_ss 1);
|
|
74 |
val add_0_right = result();
|
|
75 |
|
5061
|
76 |
Goal "?p : m+Suc(n) = Suc(m+n)";
|
0
|
77 |
by (res_inst_tac [("n","m")] induct 1);
|
|
78 |
by (ALLGOALS (ASM_SIMP_TAC add_ss));
|
|
79 |
val add_Suc_right = result();
|
|
80 |
|
|
81 |
(*mk_typed_congs appears not to work with FOLP's version of subst*)
|
|
82 |
|