New class "order" and accompanying changes.
authornipkow
Wed, 12 Feb 1997 18:53:59 +0100
changeset 2608 450c9b682a92
parent 2607 a224a2865e05
child 2609 4370e5f0fa3f
New class "order" and accompanying changes. In particular reflexivity of <= is now one rewrite rule.
src/HOL/List.ML
src/HOL/List.thy
src/HOL/Nat.ML
src/HOL/Nat.thy
src/HOL/NatDef.ML
src/HOL/NatDef.thy
src/HOL/Ord.ML
src/HOL/Ord.thy
src/HOL/Set.ML
--- a/src/HOL/List.ML	Wed Feb 12 15:43:50 1997 +0100
+++ b/src/HOL/List.ML	Wed Feb 12 18:53:59 1997 +0100
@@ -16,7 +16,7 @@
 goal List.thy "!x. xs ~= x#xs";
 by (list.induct_tac "xs" 1);
 by (ALLGOALS Asm_simp_tac);
-qed "not_Cons_self";
+qed_spec_mp "not_Cons_self";
 Addsimps [not_Cons_self];
 
 goal List.thy "(xs ~= []) = (? y ys. xs = y#ys)";
@@ -27,6 +27,29 @@
 qed "neq_Nil_conv";
 
 
+(** list_case **)
+
+goal List.thy
+ "P(list_case a f xs) = ((xs=[] --> P(a)) & \
+\                         (!y ys. xs=y#ys --> P(f y ys)))";
+by (list.induct_tac "xs" 1);
+by (ALLGOALS Asm_simp_tac);
+by (Fast_tac 1);
+qed "expand_list_case";
+
+val prems = goal List.thy "[| P([]); !!x xs. P(x#xs) |] ==> P(xs)";
+by(list.induct_tac "xs" 1);
+by(REPEAT(resolve_tac prems 1));
+qed "list_cases";
+
+goal List.thy  "(xs=[] --> P([])) & (!y ys. xs=y#ys --> P(y#ys)) --> P(xs)";
+by (list.induct_tac "xs" 1);
+by (Fast_tac 1);
+by (Fast_tac 1);
+bind_thm("list_eq_cases",
+  impI RSN (2,allI RSN (2,allI RSN (2,impI RS (conjI RS (result() RS mp))))));
+
+
 (** @ - append **)
 
 goal List.thy "(xs@ys)@zs = xs@(ys@zs)";
@@ -44,20 +67,80 @@
 goal List.thy "(xs@ys = []) = (xs=[] & ys=[])";
 by (list.induct_tac "xs" 1);
 by (ALLGOALS Asm_simp_tac);
-qed "append_is_Nil";
-Addsimps [append_is_Nil];
+qed "append_is_Nil_conv";
+AddIffs [append_is_Nil_conv];
+
+goal List.thy "([] = xs@ys) = (xs=[] & ys=[])";
+by (list.induct_tac "xs" 1);
+by (ALLGOALS Asm_simp_tac);
+by(Fast_tac 1);
+qed "Nil_is_append_conv";
+AddIffs [Nil_is_append_conv];
 
 goal List.thy "(xs @ ys = xs @ zs) = (ys=zs)";
 by (list.induct_tac "xs" 1);
 by (ALLGOALS Asm_simp_tac);
 qed "same_append_eq";
-Addsimps [same_append_eq];
+AddIffs [same_append_eq];
+
+goal List.thy "!ys. (xs @ [x] = ys @ [y]) = (xs = ys & x = y)"; 
+by(list.induct_tac "xs" 1);
+ br allI 1;
+ by(list.induct_tac "ys" 1);
+  by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(list.induct_tac "ys" 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "append1_eq_conv";
+AddIffs [append1_eq_conv];
+
+goal List.thy "xs ~= [] --> hd xs # tl xs = xs";
+by(list.induct_tac "xs" 1);
+by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "hd_Cons_tl";
+Addsimps [hd_Cons_tl];
 
 goal List.thy "hd(xs@ys) = (if xs=[] then hd ys else hd xs)";
 by (list.induct_tac "xs" 1);
 by (ALLGOALS Asm_simp_tac);
 qed "hd_append";
 
+goal List.thy "tl(xs@ys) = (case xs of [] => tl(ys) | z#zs => zs@ys)";
+by(simp_tac (!simpset setloop(split_tac[expand_list_case])) 1);
+qed "tl_append";
+
+(** map **)
+
+goal List.thy
+  "(!x. x : set_of_list xs --> f x = g x) --> map f xs = map g xs";
+by(list.induct_tac "xs" 1);
+by(ALLGOALS Asm_simp_tac);
+bind_thm("map_ext", impI RS (allI RS (result() RS mp)));
+
+goal List.thy "map (%x.x) = (%xs.xs)";
+by (rtac ext 1);
+by (list.induct_tac "xs" 1);
+by (ALLGOALS Asm_simp_tac);
+qed "map_ident";
+Addsimps[map_ident];
+
+goal List.thy "map f (xs@ys) = map f xs @ map f ys";
+by (list.induct_tac "xs" 1);
+by (ALLGOALS Asm_simp_tac);
+qed "map_append";
+Addsimps[map_append];
+
+goalw List.thy [o_def] "map (f o g) xs = map f (map g xs)";
+by (list.induct_tac "xs" 1);
+by (ALLGOALS Asm_simp_tac);
+qed "map_compose";
+Addsimps[map_compose];
+
+goal List.thy "rev(map f xs) = map f (rev xs)";
+by (list.induct_tac "xs" 1);
+by (ALLGOALS Asm_simp_tac);
+qed "rev_map";
+
 (** rev **)
 
 goal List.thy "rev(xs@ys) = rev(ys) @ rev(xs)";
@@ -72,6 +155,7 @@
 qed "rev_rev_ident";
 Addsimps[rev_rev_ident];
 
+
 (** mem **)
 
 goal List.thy "x mem (xs@ys) = (x mem xs | x mem ys)";
@@ -106,6 +190,25 @@
 by (Fast_tac 1);
 qed "set_of_list_subset_Cons";
 
+goal List.thy "(set_of_list xs = {}) = (xs = [])";
+by(list.induct_tac "xs" 1);
+by(ALLGOALS Asm_simp_tac);
+qed "set_of_list_empty";
+Addsimps [set_of_list_empty];
+
+goal List.thy "set_of_list(rev xs) = set_of_list(xs)";
+by(list.induct_tac "xs" 1);
+by(ALLGOALS Asm_simp_tac);
+by(Fast_tac 1);
+qed "set_of_list_rev";
+Addsimps [set_of_list_rev];
+
+goal List.thy "set_of_list(map f xs) = f``(set_of_list xs)";
+by(list.induct_tac "xs" 1);
+by(ALLGOALS Asm_simp_tac);
+qed "set_of_list_map";
+Addsimps [set_of_list_map];
+
 
 (** list_all **)
 
@@ -128,35 +231,27 @@
 qed "list_all_mem_conv";
 
 
-(** list_case **)
+(** filter **)
 
-goal List.thy
- "P(list_case a f xs) = ((xs=[] --> P(a)) & \
-\                         (!y ys. xs=y#ys --> P(f y ys)))";
+goal List.thy "[x:xs@ys . P] = [x:xs . P] @ [y:ys . P]";
+by(list.induct_tac "xs" 1);
+ by(ALLGOALS (asm_simp_tac (!simpset setloop (split_tac [expand_if]))));
+qed "filter_append";
+Addsimps [filter_append];
+
+
+(** concat **)
+
+goal List.thy  "concat(xs@ys) = concat(xs)@concat(ys)";
 by (list.induct_tac "xs" 1);
 by (ALLGOALS Asm_simp_tac);
-by (Fast_tac 1);
-qed "expand_list_case";
-
-val prems = goal List.thy "[| P([]); !!x xs. P(x#xs) |] ==> P(xs)";
-by(list.induct_tac "xs" 1);
-by(REPEAT(resolve_tac prems 1));
-qed "list_cases";
+qed"concat_append";
+Addsimps [concat_append];
 
-goal List.thy  "(xs=[] --> P([])) & (!y ys. xs=y#ys --> P(y#ys)) --> P(xs)";
-by (list.induct_tac "xs" 1);
-by (Fast_tac 1);
-by (Fast_tac 1);
-bind_thm("list_eq_cases",
-  impI RSN (2,allI RSN (2,allI RSN (2,impI RS (conjI RS (result() RS mp))))));
-
-(** flat **)
-
-goal List.thy  "flat(xs@ys) = flat(xs)@flat(ys)";
-by (list.induct_tac "xs" 1);
+goal List.thy "rev(concat ls) = concat (map rev (rev ls))";
+by (list.induct_tac "ls" 1);
 by (ALLGOALS Asm_simp_tac);
-qed"flat_append";
-Addsimps [flat_append];
+qed "rev_concat";
 
 (** length **)
 
@@ -178,6 +273,19 @@
 qed "length_rev";
 Addsimps [length_rev];
 
+goal List.thy "(length xs = 0) = (xs = [])";
+by(list.induct_tac "xs" 1);
+by(ALLGOALS Asm_simp_tac);
+qed "length_0_conv";
+AddIffs [length_0_conv];
+
+goal List.thy "(0 < length xs) = (xs ~= [])";
+by(list.induct_tac "xs" 1);
+by(ALLGOALS Asm_simp_tac);
+qed "length_greater_0_conv";
+AddIffs [length_greater_0_conv];
+
+
 (** nth **)
 
 val [nth_0,nth_Suc] = nat_recs nth_def; 
@@ -185,6 +293,19 @@
 store_thm("nth_Suc",nth_Suc);
 Addsimps [nth_0,nth_Suc];
 
+goal List.thy
+  "!xs. nth n (xs@ys) = \
+\          (if n < length xs then nth n xs else nth (n - length xs) ys)";
+by(nat_ind_tac "n" 1);
+ by(Asm_simp_tac 1);
+ br allI 1;
+ by(res_inst_tac [("xs","xs")]list_cases 1);
+  by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "nth_append";
+
 goal List.thy "!n. n < length xs --> nth n (map f xs) = f (nth n xs)";
 by (list.induct_tac "xs" 1);
 (* case [] *)
@@ -220,61 +341,188 @@
 qed_spec_mp "nth_mem";
 Addsimps [nth_mem];
 
-(** drop **)
 
-goal thy "drop 0 xs = xs";
-by (list.induct_tac "xs" 1);
-by (ALLGOALS Asm_simp_tac);
-qed "drop_0";
-
-goal thy "drop (Suc n) (x#xs) = drop n xs";
-by (Simp_tac 1);
-qed "drop_Suc_Cons";
-
-Delsimps [drop_Cons];
-Addsimps [drop_0,drop_Suc_Cons];
-
-(** take **)
+(** take  & drop **)
+section "take & drop";
 
 goal thy "take 0 xs = []";
 by (list.induct_tac "xs" 1);
 by (ALLGOALS Asm_simp_tac);
 qed "take_0";
 
+goal thy "drop 0 xs = xs";
+by (list.induct_tac "xs" 1);
+by (ALLGOALS Asm_simp_tac);
+qed "drop_0";
+
 goal thy "take (Suc n) (x#xs) = x # take n xs";
 by (Simp_tac 1);
 qed "take_Suc_Cons";
 
-Delsimps [take_Cons];
-Addsimps [take_0,take_Suc_Cons];
+goal thy "drop (Suc n) (x#xs) = drop n xs";
+by (Simp_tac 1);
+qed "drop_Suc_Cons";
+
+Delsimps [take_Cons,drop_Cons];
+Addsimps [take_0,take_Suc_Cons,drop_0,drop_Suc_Cons];
+
+goal List.thy "!xs. length(take n xs) = min (length xs) n";
+by(nat_ind_tac "n" 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "length_take";
+Addsimps [length_take];
 
-(** Additional mapping lemmas **)
+goal List.thy "!xs. length(drop n xs) = (length xs - n)";
+by(nat_ind_tac "n" 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "length_drop";
+Addsimps [length_drop];
+
+goal List.thy "!xs. length xs <= n --> take n xs = xs";
+by(nat_ind_tac "n" 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "take_all";
 
-goal List.thy "map (%x.x) = (%xs.xs)";
-by (rtac ext 1);
-by (list.induct_tac "xs" 1);
-by (ALLGOALS Asm_simp_tac);
-qed "map_ident";
-Addsimps[map_ident];
+goal List.thy "!xs. length xs <= n --> drop n xs = []";
+by(nat_ind_tac "n" 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "drop_all";
+
+goal List.thy 
+  "!xs. take n (xs @ ys) = (take n xs @ take (n - length xs) ys)";
+by(nat_ind_tac "n" 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "take_append";
+Addsimps [take_append];
+
+goal List.thy "!xs. drop n (xs@ys) = drop n xs @ drop (n - length xs) ys"; 
+by(nat_ind_tac "n" 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "drop_append";
+Addsimps [drop_append];
+
+goal List.thy "!xs n. take n (take m xs) = take (min n m) xs"; 
+by(nat_ind_tac "m" 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("n","n")]natE 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "take_take";
+
+goal List.thy "!xs. drop n (drop m xs) = drop (n + m) xs"; 
+by(nat_ind_tac "m" 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "drop_drop";
 
-goal List.thy "map f (xs@ys) = map f xs @ map f ys";
-by (list.induct_tac "xs" 1);
-by (ALLGOALS Asm_simp_tac);
-qed "map_append";
-Addsimps[map_append];
+goal List.thy "!xs n. take n (drop m xs) = drop m (take (n + m) xs)"; 
+by(nat_ind_tac "m" 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "take_drop";
+
+goal List.thy "!xs. take n (map f xs) = map f (take n xs)"; 
+by(nat_ind_tac "n" 1);
+by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "take_map"; 
+
+goal List.thy "!xs. drop n (map f xs) = map f (drop n xs)"; 
+by(nat_ind_tac "n" 1);
+by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "drop_map";
+
+goal List.thy
+  "!n i. i < n --> nth i (take n xs) = nth i xs";
+by(list.induct_tac "xs" 1);
+ by(ALLGOALS Asm_simp_tac);
+by(strip_tac 1);
+by(res_inst_tac [("n","n")] natE 1);
+ by(Fast_tac 1);
+by(res_inst_tac [("n","i")] natE 1);
+by(ALLGOALS (hyp_subst_tac THEN' Asm_full_simp_tac));
+qed_spec_mp "nth_take";
+Addsimps [nth_take];
 
-goalw List.thy [o_def] "map (f o g) xs = map f (map g xs)";
-by (list.induct_tac "xs" 1);
-by (ALLGOALS Asm_simp_tac);
-qed "map_compose";
-Addsimps[map_compose];
+goal List.thy
+  "!xs i. n + i < length xs --> nth i (drop n xs) = nth (n + i) xs";
+by(nat_ind_tac "n" 1);
+ by(ALLGOALS Asm_simp_tac);
+br allI 1;
+by(res_inst_tac [("xs","xs")]list_cases 1);
+ by(ALLGOALS Asm_simp_tac);
+qed_spec_mp "nth_drop";
+Addsimps [nth_drop];
+
+(** takeWhile & dropWhile **)
+
+goal List.thy
+  "x:set_of_list xs & ~P(x) --> takeWhile P (xs @ ys) = takeWhile P xs";
+by(list.induct_tac "xs" 1);
+ by(Simp_tac 1);
+by(asm_full_simp_tac (!simpset setloop (split_tac[expand_if])) 1);
+by(Fast_tac 1);
+bind_thm("takeWhile_append1", conjI RS (result() RS mp));
+Addsimps [takeWhile_append1];
 
-goal List.thy "rev(map f l) = map f (rev l)";
-by (list.induct_tac "l" 1);
-by (ALLGOALS Asm_simp_tac);
-qed "rev_map_distrib";
+goal List.thy
+  "(!x:set_of_list xs.P(x)) --> takeWhile P (xs @ ys) = xs @ takeWhile P ys";
+by(list.induct_tac "xs" 1);
+ by(Simp_tac 1);
+by(asm_full_simp_tac (!simpset setloop (split_tac[expand_if])) 1);
+bind_thm("takeWhile_append2", ballI RS (result() RS mp));
+Addsimps [takeWhile_append2];
 
-goal List.thy "rev(flat ls) = flat (map rev (rev ls))";
-by (list.induct_tac "ls" 1);
-by (ALLGOALS Asm_simp_tac);
-qed "rev_flat";
+goal List.thy
+  "x:set_of_list xs & ~P(x) --> dropWhile P (xs @ ys) = (dropWhile P xs)@ys";
+by(list.induct_tac "xs" 1);
+ by(Simp_tac 1);
+by(asm_full_simp_tac (!simpset setloop (split_tac[expand_if])) 1);
+by(Fast_tac 1);
+bind_thm("dropWhile_append1", conjI RS (result() RS mp));
+Addsimps [dropWhile_append1];
+
+goal List.thy
+  "(!x:set_of_list xs.P(x)) --> dropWhile P (xs @ ys) = dropWhile P ys";
+by(list.induct_tac "xs" 1);
+ by(Simp_tac 1);
+by(asm_full_simp_tac (!simpset setloop (split_tac[expand_if])) 1);
+bind_thm("dropWhile_append2", ballI RS (result() RS mp));
+Addsimps [dropWhile_append2];
+
+goal List.thy "x:set_of_list(takeWhile P xs) --> x:set_of_list xs & P x";
+by(list.induct_tac "xs" 1);
+ by(Simp_tac 1);
+by(asm_full_simp_tac (!simpset setloop (split_tac[expand_if])) 1);
+qed_spec_mp"set_of_list_take_whileD";
+
--- a/src/HOL/List.thy	Wed Feb 12 15:43:50 1997 +0100
+++ b/src/HOL/List.thy	Wed Feb 12 18:53:59 1997 +0100
@@ -13,18 +13,19 @@
 consts
 
   "@"         :: ['a list, 'a list] => 'a list            (infixr 65)
-  drop        :: [nat, 'a list] => 'a list
   filter      :: ['a => bool, 'a list] => 'a list
-  flat        :: 'a list list => 'a list
+  concat      :: 'a list list => 'a list
   foldl       :: [['b,'a] => 'b, 'b, 'a list] => 'b
   hd          :: 'a list => 'a
   length      :: 'a list => nat
-  set_of_list :: ('a list => 'a set)
+  set_of_list :: 'a list => 'a set
   list_all    :: ('a => bool) => ('a list => bool)
   map         :: ('a=>'b) => ('a list => 'b list)
   mem         :: ['a, 'a list] => bool                    (infixl 55)
   nth         :: [nat, 'a list] => 'a
-  take        :: [nat, 'a list] => 'a list
+  take, drop  :: [nat, 'a list] => 'a list
+  takeWhile,
+  dropWhile   :: ('a => bool) => 'a list => 'a list
   tl,ttl      :: 'a list => 'a list
   rev         :: 'a list => 'a list
 
@@ -81,9 +82,9 @@
 primrec length list
   "length([]) = 0"
   "length(x#xs) = Suc(length(xs))"
-primrec flat list
-  "flat([]) = []"
-  "flat(x#xs) = x @ flat(xs)"
+primrec concat list
+  "concat([]) = []"
+  "concat(x#xs) = x @ concat(xs)"
 primrec drop list
   drop_Nil  "drop n [] = []"
   drop_Cons "drop n (x#xs) = (case n of 0 => x#xs | Suc(m) => drop m xs)"
@@ -92,4 +93,10 @@
   take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)"
 defs
   nth_def  "nth(n) == nat_rec hd (%m r xs. r(tl(xs))) n"
+primrec takeWhile list
+  "takeWhile P [] = []"
+  "takeWhile P (x#xs) = (if P x then x#takeWhile P xs else [])"
+primrec dropWhile list
+  "dropWhile P [] = []"
+  "dropWhile P (x#xs) = (if P x then dropWhile P xs else xs)"
 end
--- a/src/HOL/Nat.ML	Wed Feb 12 15:43:50 1997 +0100
+++ b/src/HOL/Nat.ML	Wed Feb 12 18:53:59 1997 +0100
@@ -1,678 +1,22 @@
 (*  Title:      HOL/Nat.ML
     ID:         $Id$
-    Author:     Tobias Nipkow, Cambridge University Computer Laboratory
-    Copyright   1991  University of Cambridge
-
-For Nat.thy.  Type nat is defined as a set (Nat) over the type ind.
+    Author:     Tobias Nipkow
+    Copyright   1997 TU Muenchen
 *)
 
-open Nat;
-
-goal Nat.thy "mono(%X. {Zero_Rep} Un (Suc_Rep``X))";
-by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1));
-qed "Nat_fun_mono";
-
-val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski);
-
-(* Zero is a natural number -- this also justifies the type definition*)
-goal Nat.thy "Zero_Rep: Nat";
-by (stac Nat_unfold 1);
-by (rtac (singletonI RS UnI1) 1);
-qed "Zero_RepI";
-
-val prems = goal Nat.thy "i: Nat ==> Suc_Rep(i) : Nat";
-by (stac Nat_unfold 1);
-by (rtac (imageI RS UnI2) 1);
-by (resolve_tac prems 1);
-qed "Suc_RepI";
-
-(*** Induction ***)
-
-val major::prems = goal Nat.thy
-    "[| i: Nat;  P(Zero_Rep);   \
-\       !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |]  ==> P(i)";
-by (rtac ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1);
-by (fast_tac (!claset addIs prems) 1);
-qed "Nat_induct";
-
-val prems = goalw Nat.thy [Zero_def,Suc_def]
-    "[| P(0);   \
-\       !!k. P(k) ==> P(Suc(k)) |]  ==> P(n)";
-by (rtac (Rep_Nat_inverse RS subst) 1);   (*types force good instantiation*)
-by (rtac (Rep_Nat RS Nat_induct) 1);
-by (REPEAT (ares_tac prems 1
-     ORELSE eresolve_tac [Abs_Nat_inverse RS subst] 1));
-qed "nat_induct";
-
-(*Perform induction on n. *)
-fun nat_ind_tac a i = 
-    EVERY [res_inst_tac [("n",a)] nat_induct i,
-           rename_last_tac a ["1"] (i+1)];
-
-(*A special form of induction for reasoning about m<n and m-n*)
-val prems = goal Nat.thy
-    "[| !!x. P x 0;  \
-\       !!y. P 0 (Suc y);  \
-\       !!x y. [| P x y |] ==> P (Suc x) (Suc y)  \
-\    |] ==> P m n";
-by (res_inst_tac [("x","m")] spec 1);
-by (nat_ind_tac "n" 1);
-by (rtac allI 2);
-by (nat_ind_tac "x" 2);
-by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1));
-qed "diff_induct";
-
-(*Case analysis on the natural numbers*)
-val prems = goal Nat.thy 
-    "[| n=0 ==> P;  !!x. n = Suc(x) ==> P |] ==> P";
-by (subgoal_tac "n=0 | (EX x. n = Suc(x))" 1);
-by (fast_tac (!claset addSEs prems) 1);
-by (nat_ind_tac "n" 1);
-by (rtac (refl RS disjI1) 1);
-by (Fast_tac 1);
-qed "natE";
-
-(*** Isomorphisms: Abs_Nat and Rep_Nat ***)
-
-(*We can't take these properties as axioms, or take Abs_Nat==Inv(Rep_Nat),
-  since we assume the isomorphism equations will one day be given by Isabelle*)
-
-goal Nat.thy "inj(Rep_Nat)";
-by (rtac inj_inverseI 1);
-by (rtac Rep_Nat_inverse 1);
-qed "inj_Rep_Nat";
-
-goal Nat.thy "inj_onto Abs_Nat Nat";
-by (rtac inj_onto_inverseI 1);
-by (etac Abs_Nat_inverse 1);
-qed "inj_onto_Abs_Nat";
-
-(*** Distinctness of constructors ***)
-
-goalw Nat.thy [Zero_def,Suc_def] "Suc(m) ~= 0";
-by (rtac (inj_onto_Abs_Nat RS inj_onto_contraD) 1);
-by (rtac Suc_Rep_not_Zero_Rep 1);
-by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1));
-qed "Suc_not_Zero";
-
-bind_thm ("Zero_not_Suc", Suc_not_Zero RS not_sym);
-
-AddIffs [Suc_not_Zero,Zero_not_Suc];
-
-bind_thm ("Suc_neq_Zero", (Suc_not_Zero RS notE));
-val Zero_neq_Suc = sym RS Suc_neq_Zero;
-
-(** Injectiveness of Suc **)
-
-goalw Nat.thy [Suc_def] "inj(Suc)";
-by (rtac injI 1);
-by (dtac (inj_onto_Abs_Nat RS inj_ontoD) 1);
-by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1));
-by (dtac (inj_Suc_Rep RS injD) 1);
-by (etac (inj_Rep_Nat RS injD) 1);
-qed "inj_Suc";
-
-val Suc_inject = inj_Suc RS injD;
-
-goal Nat.thy "(Suc(m)=Suc(n)) = (m=n)";
-by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); 
-qed "Suc_Suc_eq";
-
-AddIffs [Suc_Suc_eq];
-
-goal Nat.thy "n ~= Suc(n)";
-by (nat_ind_tac "n" 1);
-by (ALLGOALS Asm_simp_tac);
-qed "n_not_Suc_n";
-
-bind_thm ("Suc_n_not_n", n_not_Suc_n RS not_sym);
-
-(*** nat_case -- the selection operator for nat ***)
-
-goalw Nat.thy [nat_case_def] "nat_case a f 0 = a";
-by (fast_tac (!claset addIs [select_equality]) 1);
-qed "nat_case_0";
-
-goalw Nat.thy [nat_case_def] "nat_case a f (Suc k) = f(k)";
-by (fast_tac (!claset addIs [select_equality]) 1);
-qed "nat_case_Suc";
-
-(** Introduction rules for 'pred_nat' **)
-
-goalw Nat.thy [pred_nat_def] "(n, Suc(n)) : pred_nat";
-by (Fast_tac 1);
-qed "pred_natI";
-
-val major::prems = goalw Nat.thy [pred_nat_def]
-    "[| p : pred_nat;  !!x n. [| p = (n, Suc(n)) |] ==> R \
-\    |] ==> R";
-by (rtac (major RS CollectE) 1);
-by (REPEAT (eresolve_tac ([asm_rl,exE]@prems) 1));
-qed "pred_natE";
-
-goalw Nat.thy [wf_def] "wf(pred_nat)";
-by (strip_tac 1);
-by (nat_ind_tac "x" 1);
-by (fast_tac (!claset addSEs [mp, pred_natE]) 2);
-by (fast_tac (!claset addSEs [mp, pred_natE]) 1);
-qed "wf_pred_nat";
-
-
-(*** nat_rec -- by wf recursion on pred_nat ***)
-
-(* The unrolling rule for nat_rec *)
-goal Nat.thy
-   "(%n. nat_rec c d n) = wfrec pred_nat (%f. nat_case ?c (%m. ?d m (f m)))";
-  by (simp_tac (HOL_ss addsimps [nat_rec_def]) 1);
-bind_thm("nat_rec_unfold", wf_pred_nat RS 
-                            ((result() RS eq_reflection) RS def_wfrec));
-
-(*---------------------------------------------------------------------------
- * Old:
- * bind_thm ("nat_rec_unfold", (wf_pred_nat RS (nat_rec_def RS def_wfrec))); 
- *---------------------------------------------------------------------------*)
-
-(** conversion rules **)
-
-goal Nat.thy "nat_rec c h 0 = c";
-by (rtac (nat_rec_unfold RS trans) 1);
-by (simp_tac (!simpset addsimps [nat_case_0]) 1);
-qed "nat_rec_0";
-
-goal Nat.thy "nat_rec c h (Suc n) = h n (nat_rec c h n)";
-by (rtac (nat_rec_unfold RS trans) 1);
-by (simp_tac (!simpset addsimps [nat_case_Suc, pred_natI, cut_apply]) 1);
-qed "nat_rec_Suc";
-
-(*These 2 rules ease the use of primitive recursion.  NOTE USE OF == *)
-val [rew] = goal Nat.thy
-    "[| !!n. f(n) == nat_rec c h n |] ==> f(0) = c";
-by (rewtac rew);
-by (rtac nat_rec_0 1);
-qed "def_nat_rec_0";
-
-val [rew] = goal Nat.thy
-    "[| !!n. f(n) == nat_rec c h n |] ==> f(Suc(n)) = h n (f n)";
-by (rewtac rew);
-by (rtac nat_rec_Suc 1);
-qed "def_nat_rec_Suc";
-
-fun nat_recs def =
-      [standard (def RS def_nat_rec_0),
-       standard (def RS def_nat_rec_Suc)];
-
-
-(*** Basic properties of "less than" ***)
-
-(** Introduction properties **)
-
-val prems = goalw Nat.thy [less_def] "[| i<j;  j<k |] ==> i<(k::nat)";
-by (rtac (trans_trancl RS transD) 1);
-by (resolve_tac prems 1);
-by (resolve_tac prems 1);
-qed "less_trans";
-
-goalw Nat.thy [less_def] "n < Suc(n)";
-by (rtac (pred_natI RS r_into_trancl) 1);
-qed "lessI";
-AddIffs [lessI];
-
-(* i<j ==> i<Suc(j) *)
-val less_SucI = lessI RSN (2, less_trans);
-
-goal Nat.thy "0 < Suc(n)";
-by (nat_ind_tac "n" 1);
-by (rtac lessI 1);
-by (etac less_trans 1);
-by (rtac lessI 1);
-qed "zero_less_Suc";
-AddIffs [zero_less_Suc];
-
-(** Elimination properties **)
-
-val prems = goalw Nat.thy [less_def] "n<m ==> ~ m<(n::nat)";
-by (fast_tac (!claset addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1);
-qed "less_not_sym";
-
-(* [| n<m; m<n |] ==> R *)
-bind_thm ("less_asym", (less_not_sym RS notE));
-
-goalw Nat.thy [less_def] "~ n<(n::nat)";
-by (rtac notI 1);
-by (etac (wf_pred_nat RS wf_trancl RS wf_irrefl) 1);
-qed "less_not_refl";
-
-(* n<n ==> R *)
-bind_thm ("less_irrefl", (less_not_refl RS notE));
-
-goal Nat.thy "!!m. n<m ==> m ~= (n::nat)";
-by (fast_tac (!claset addEs [less_irrefl]) 1);
-qed "less_not_refl2";
-
-
-val major::prems = goalw Nat.thy [less_def]
-    "[| i<k;  k=Suc(i) ==> P;  !!j. [| i<j;  k=Suc(j) |] ==> P \
-\    |] ==> P";
-by (rtac (major RS tranclE) 1);
-by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE'
-                  eresolve_tac (prems@[pred_natE, Pair_inject])));
-by (rtac refl 1);
-qed "lessE";
-
-goal Nat.thy "~ n<0";
-by (rtac notI 1);
-by (etac lessE 1);
-by (etac Zero_neq_Suc 1);
-by (etac Zero_neq_Suc 1);
-qed "not_less0";
-
-AddIffs [not_less0];
-
-(* n<0 ==> R *)
-bind_thm ("less_zeroE", not_less0 RS notE);
-
-val [major,less,eq] = goal Nat.thy
-    "[| m < Suc(n);  m<n ==> P;  m=n ==> P |] ==> P";
-by (rtac (major RS lessE) 1);
-by (rtac eq 1);
-by (Fast_tac 1);
-by (rtac less 1);
-by (Fast_tac 1);
-qed "less_SucE";
-
-goal Nat.thy "(m < Suc(n)) = (m < n | m = n)";
-by (fast_tac (!claset addSIs [lessI]
-                      addEs  [less_trans, less_SucE]) 1);
-qed "less_Suc_eq";
-
-val prems = goal Nat.thy "m<n ==> n ~= 0";
-by (res_inst_tac [("n","n")] natE 1);
-by (cut_facts_tac prems 1);
-by (ALLGOALS Asm_full_simp_tac);
-qed "gr_implies_not0";
-Addsimps [gr_implies_not0];
-
-qed_goal "zero_less_eq" Nat.thy "0 < n = (n ~= 0)" (fn _ => [
-        rtac iffI 1,
-        etac gr_implies_not0 1,
-        rtac natE 1,
-        contr_tac 1,
-        etac ssubst 1,
-        rtac zero_less_Suc 1]);
-
-(** Inductive (?) properties **)
-
-val [prem] = goal Nat.thy "Suc(m) < n ==> m<n";
-by (rtac (prem RS rev_mp) 1);
-by (nat_ind_tac "n" 1);
-by (ALLGOALS (fast_tac (!claset addSIs [lessI RS less_SucI]
-                                addEs  [less_trans, lessE])));
-qed "Suc_lessD";
-
-val [major,minor] = goal Nat.thy 
-    "[| Suc(i)<k;  !!j. [| i<j;  k=Suc(j) |] ==> P \
-\    |] ==> P";
-by (rtac (major RS lessE) 1);
-by (etac (lessI RS minor) 1);
-by (etac (Suc_lessD RS minor) 1);
-by (assume_tac 1);
-qed "Suc_lessE";
-
-goal Nat.thy "!!m n. Suc(m) < Suc(n) ==> m<n";
-by (fast_tac (!claset addEs [lessE, Suc_lessD] addIs [lessI]) 1);
-qed "Suc_less_SucD";
-
-goal Nat.thy "!!m n. m<n ==> Suc(m) < Suc(n)";
-by (etac rev_mp 1);
-by (nat_ind_tac "n" 1);
-by (ALLGOALS (fast_tac (!claset addSIs [lessI]
-                                addEs  [less_trans, lessE])));
-qed "Suc_mono";
-
-
-goal Nat.thy "(Suc(m) < Suc(n)) = (m<n)";
-by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]);
-qed "Suc_less_eq";
-Addsimps [Suc_less_eq];
-
-goal Nat.thy "~(Suc(n) < n)";
-by (fast_tac (!claset addEs [Suc_lessD RS less_irrefl]) 1);
-qed "not_Suc_n_less_n";
-Addsimps [not_Suc_n_less_n];
+goal thy "min 0 n = 0";
+br min_leastL 1;
+by(trans_tac 1);
+qed "min_0L";
 
-goal Nat.thy "!!i. i<j ==> j<k --> Suc i < k";
-by (nat_ind_tac "k" 1);
-by (ALLGOALS (asm_simp_tac (!simpset)));
-by (asm_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
-by (fast_tac (!claset addDs [Suc_lessD]) 1);
-qed_spec_mp "less_trans_Suc";
-
-(*"Less than" is a linear ordering*)
-goal Nat.thy "m<n | m=n | n<(m::nat)";
-by (nat_ind_tac "m" 1);
-by (nat_ind_tac "n" 1);
-by (rtac (refl RS disjI1 RS disjI2) 1);
-by (rtac (zero_less_Suc RS disjI1) 1);
-by (fast_tac (!claset addIs [lessI, Suc_mono, less_SucI] addEs [lessE]) 1);
-qed "less_linear";
-
-qed_goal "nat_less_cases" Nat.thy 
-   "[| (m::nat)<n ==> P n m; m=n ==> P n m; n<m ==> P n m |] ==> P n m"
-( fn prems =>
-        [
-        (res_inst_tac [("m1","n"),("n1","m")] (less_linear RS disjE) 1),
-        (etac disjE 2),
-        (etac (hd (tl (tl prems))) 1),
-        (etac (sym RS hd (tl prems)) 1),
-        (etac (hd prems) 1)
-        ]);
-
-(*Can be used with less_Suc_eq to get n=m | n<m *)
-goal Nat.thy "(~ m < n) = (n < Suc(m))";
-by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
-by (ALLGOALS Asm_simp_tac);
-qed "not_less_eq";
-
-(*Complete induction, aka course-of-values induction*)
-val prems = goalw Nat.thy [less_def]
-    "[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |]  ==>  P(n)";
-by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1);
-by (eresolve_tac prems 1);
-qed "less_induct";
-
-qed_goal "nat_induct2" Nat.thy 
-"[| P 0; P 1; !!k. P k ==> P (Suc (Suc k)) |] ==> P n" (fn prems => [
-	cut_facts_tac prems 1,
-	rtac less_induct 1,
-	res_inst_tac [("n","n")] natE 1,
-	 hyp_subst_tac 1,
-	 atac 1,
-	hyp_subst_tac 1,
-	res_inst_tac [("n","x")] natE 1,
-	 hyp_subst_tac 1,
-	 atac 1,
-	hyp_subst_tac 1,
-	resolve_tac prems 1,
-	dtac spec 1,
-	etac mp 1,
-	rtac (lessI RS less_trans) 1,
-	rtac (lessI RS Suc_mono) 1]);
-
-(*** Properties of <= ***)
-
-goalw Nat.thy [le_def] "(m <= n) = (m < Suc n)";
-by (rtac not_less_eq 1);
-qed "le_eq_less_Suc";
-
-goalw Nat.thy [le_def] "0 <= n";
-by (rtac not_less0 1);
-qed "le0";
-
-goalw Nat.thy [le_def] "~ Suc n <= n";
-by (Simp_tac 1);
-qed "Suc_n_not_le_n";
-
-goalw Nat.thy [le_def] "(i <= 0) = (i = 0)";
-by (nat_ind_tac "i" 1);
-by (ALLGOALS Asm_simp_tac);
-qed "le_0_eq";
-
-Addsimps [less_not_refl,
-          (*less_Suc_eq, makes simpset non-confluent*) le0, le_0_eq,
-          Suc_n_not_le_n,
-          n_not_Suc_n, Suc_n_not_n,
-          nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc];
-
-(*
-goal Nat.thy "(Suc m < n | Suc m = n) = (m < n)";
-by (stac (less_Suc_eq RS sym) 1);
-by (rtac Suc_less_eq 1);
-qed "Suc_le_eq";
-
-this could make the simpset (with less_Suc_eq added again) more confluent,
-but less_Suc_eq makes additional problems with terms of the form 0 < Suc (...)
-*)
-
-(*Prevents simplification of f and g: much faster*)
-qed_goal "nat_case_weak_cong" Nat.thy
-  "m=n ==> nat_case a f m = nat_case a f n"
-  (fn [prem] => [rtac (prem RS arg_cong) 1]);
-
-qed_goal "nat_rec_weak_cong" Nat.thy
-  "m=n ==> nat_rec a f m = nat_rec a f n"
-  (fn [prem] => [rtac (prem RS arg_cong) 1]);
-
-qed_goal "expand_nat_case" Nat.thy
-  "P(nat_case z s n) = ((n=0 --> P z) & (!m. n = Suc m --> P(s m)))"
-  (fn _ => [nat_ind_tac "n" 1, ALLGOALS Asm_simp_tac]);
-
-val prems = goalw Nat.thy [le_def] "~n<m ==> m<=(n::nat)";
-by (resolve_tac prems 1);
-qed "leI";
-
-val prems = goalw Nat.thy [le_def] "m<=n ==> ~ n < (m::nat)";
-by (resolve_tac prems 1);
-qed "leD";
-
-val leE = make_elim leD;
-
-goal Nat.thy "(~n<m) = (m<=(n::nat))";
-by (fast_tac (!claset addIs [leI] addEs [leE]) 1);
-qed "not_less_iff_le";
-
-goalw Nat.thy [le_def] "!!m. ~ m <= n ==> n<(m::nat)";
-by (Fast_tac 1);
-qed "not_leE";
-
-goalw Nat.thy [le_def] "!!m. m < n ==> Suc(m) <= n";
-by (simp_tac (!simpset addsimps [less_Suc_eq]) 1);
-by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
-qed "lessD";
-
-goalw Nat.thy [le_def] "!!m. Suc(m) <= n ==> m <= n";
-by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
-qed "Suc_leD";
-
-(* stronger version of Suc_leD *)
-goalw Nat.thy [le_def] 
-        "!!m. Suc m <= n ==> m < n";
-by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
-by (cut_facts_tac [less_linear] 1);
-by (Fast_tac 1);
-qed "Suc_le_lessD";
-
-goal Nat.thy "(Suc m <= n) = (m < n)";
-by (fast_tac (!claset addIs [lessD, Suc_le_lessD]) 1);
-qed "Suc_le_eq";
-
-goalw Nat.thy [le_def] "!!m. m <= n ==> m <= Suc n";
-by (fast_tac (!claset addDs [Suc_lessD]) 1);
-qed "le_SucI";
-Addsimps[le_SucI];
-
-bind_thm ("le_Suc", not_Suc_n_less_n RS leI);
-
-goalw Nat.thy [le_def] "!!m. m < n ==> m <= (n::nat)";
-by (fast_tac (!claset addEs [less_asym]) 1);
-qed "less_imp_le";
-
-goalw Nat.thy [le_def] "!!m. m <= n ==> m < n | m=(n::nat)";
-by (cut_facts_tac [less_linear] 1);
-by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
-qed "le_imp_less_or_eq";
-
-goalw Nat.thy [le_def] "!!m. m<n | m=n ==> m <=(n::nat)";
-by (cut_facts_tac [less_linear] 1);
-by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
-by (flexflex_tac);
-qed "less_or_eq_imp_le";
+goal thy "min n 0 = 0";
+br min_leastR 1;
+by(trans_tac 1);
+qed "min_0R";
 
-goal Nat.thy "(x <= (y::nat)) = (x < y | x=y)";
-by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1));
-qed "le_eq_less_or_eq";
-
-goal Nat.thy "n <= (n::nat)";
-by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1);
-qed "le_refl";
-
-val prems = goal Nat.thy "!!i. [| i <= j; j < k |] ==> i < (k::nat)";
-by (dtac le_imp_less_or_eq 1);
-by (fast_tac (!claset addIs [less_trans]) 1);
-qed "le_less_trans";
-
-goal Nat.thy "!!i. [| i < j; j <= k |] ==> i < (k::nat)";
-by (dtac le_imp_less_or_eq 1);
-by (fast_tac (!claset addIs [less_trans]) 1);
-qed "less_le_trans";
-
-goal Nat.thy "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)";
-by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
-          rtac less_or_eq_imp_le, fast_tac (!claset addIs [less_trans])]);
-qed "le_trans";
-
-val prems = goal Nat.thy "!!m. [| m <= n; n <= m |] ==> m = (n::nat)";
-by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
-          fast_tac (!claset addEs [less_irrefl,less_asym])]);
-qed "le_anti_sym";
-
-goal Nat.thy "(Suc(n) <= Suc(m)) = (n <= m)";
-by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1);
-qed "Suc_le_mono";
-
-AddIffs [le_refl,Suc_le_mono];
-
-
-(** LEAST -- the least number operator **)
-
-val [prem1,prem2] = goalw Nat.thy [Least_def]
-    "[| P(k);  !!x. x<k ==> ~P(x) |] ==> (LEAST x.P(x)) = k";
-by (rtac select_equality 1);
-by (fast_tac (!claset addSIs [prem1,prem2]) 1);
-by (cut_facts_tac [less_linear] 1);
-by (fast_tac (!claset addSIs [prem1] addSDs [prem2]) 1);
-qed "Least_equality";
-
-val [prem] = goal Nat.thy "P(k) ==> P(LEAST x.P(x))";
-by (rtac (prem RS rev_mp) 1);
-by (res_inst_tac [("n","k")] less_induct 1);
-by (rtac impI 1);
-by (rtac classical 1);
-by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
-by (assume_tac 1);
-by (assume_tac 2);
-by (Fast_tac 1);
-qed "LeastI";
-
-(*Proof is almost identical to the one above!*)
-val [prem] = goal Nat.thy "P(k) ==> (LEAST x.P(x)) <= k";
-by (rtac (prem RS rev_mp) 1);
-by (res_inst_tac [("n","k")] less_induct 1);
-by (rtac impI 1);
-by (rtac classical 1);
-by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
-by (assume_tac 1);
-by (rtac le_refl 2);
-by (fast_tac (!claset addIs [less_imp_le,le_trans]) 1);
-qed "Least_le";
-
-val [prem] = goal Nat.thy "k < (LEAST x.P(x)) ==> ~P(k)";
-by (rtac notI 1);
-by (etac (rewrite_rule [le_def] Least_le RS notE) 1);
-by (rtac prem 1);
-qed "not_less_Least";
+goalw thy [min_def] "min (Suc m) (Suc n) = Suc(min m n)";
+by(split_tac [expand_if] 1);
+by(Simp_tac 1);
+qed "min_Suc_Suc";
 
-qed_goalw "Least_Suc" Nat.thy [Least_def]
- "!!P. [| ? n. P(Suc n); ~ P 0 |] ==> (LEAST n. P n) = Suc (LEAST m. P(Suc m))"
- (fn _ => [
-        rtac select_equality 1,
-        fold_goals_tac [Least_def],
-        safe_tac (!claset addSEs [LeastI]),
-        res_inst_tac [("n","j")] natE 1,
-        Fast_tac 1,
-        fast_tac (!claset addDs [Suc_less_SucD, not_less_Least]) 1,
-        res_inst_tac [("n","k")] natE 1,
-        Fast_tac 1,
-        hyp_subst_tac 1,
-        rewtac Least_def,
-        rtac (select_equality RS arg_cong RS sym) 1,
-        safe_tac (!claset),
-        dtac Suc_mono 1,
-        Fast_tac 1,
-        cut_facts_tac [less_linear] 1,
-        safe_tac (!claset),
-        atac 2,
-        Fast_tac 2,
-        dtac Suc_mono 1,
-        Fast_tac 1]);
-
-
-(*** Instantiation of transitivity prover ***)
-
-structure Less_Arith =
-struct
-val nat_leI = leI;
-val nat_leD = leD;
-val lessI = lessI;
-val zero_less_Suc = zero_less_Suc;
-val less_reflE = less_irrefl;
-val less_zeroE = less_zeroE;
-val less_incr = Suc_mono;
-val less_decr = Suc_less_SucD;
-val less_incr_rhs = Suc_mono RS Suc_lessD;
-val less_decr_lhs = Suc_lessD;
-val less_trans_Suc = less_trans_Suc;
-val leI = lessD RS (Suc_le_mono RS iffD1);
-val not_lessI = leI RS leD
-val not_leI = prove_goal Nat.thy "!!m::nat. n < m ==> ~ m <= n"
-  (fn _ => [etac swap2 1, etac leD 1]);
-val eqI = prove_goal Nat.thy "!!m. [| m < Suc n; n < Suc m |] ==> m=n"
-  (fn _ => [etac less_SucE 1,
-            fast_tac (HOL_cs addSDs [Suc_less_SucD] addSEs [less_irrefl]
-                             addDs [less_trans_Suc]) 1,
-            atac 1]);
-val leD = le_eq_less_Suc RS iffD1;
-val not_lessD = nat_leI RS leD;
-val not_leD = not_leE
-val eqD1 = prove_goal Nat.thy  "!!n. m = n ==> m < Suc n"
- (fn _ => [etac subst 1, rtac lessI 1]);
-val eqD2 = sym RS eqD1;
-
-fun is_zero(t) =  t = Const("0",Type("nat",[]));
-
-fun nnb T = T = Type("fun",[Type("nat",[]),
-                            Type("fun",[Type("nat",[]),
-                                        Type("bool",[])])])
-
-fun decomp_Suc(Const("Suc",_)$t) = let val (a,i) = decomp_Suc t in (a,i+1) end
-  | decomp_Suc t = (t,0);
-
-fun decomp2(rel,T,lhs,rhs) =
-  if not(nnb T) then None else
-  let val (x,i) = decomp_Suc lhs
-      val (y,j) = decomp_Suc rhs
-  in case rel of
-       "op <"  => Some(x,i,"<",y,j)
-     | "op <=" => Some(x,i,"<=",y,j)
-     | "op ="  => Some(x,i,"=",y,j)
-     | _       => None
-  end;
-
-fun negate(Some(x,i,rel,y,j)) = Some(x,i,"~"^rel,y,j)
-  | negate None = None;
-
-fun decomp(_$(Const(rel,T)$lhs$rhs)) = decomp2(rel,T,lhs,rhs)
-  | decomp(_$(Const("not",_)$(Const(rel,T)$lhs$rhs))) =
-      negate(decomp2(rel,T,lhs,rhs))
-  | decomp _ = None
-
-end;
-
-structure Trans_Tac = Trans_Tac_Fun(Less_Arith);
-
-open Trans_Tac;
-
-(*** eliminates ~= in premises, which trans_tac cannot deal with ***)
-qed_goal "nat_neqE" Nat.thy
-  "[| (m::nat) ~= n; m < n ==> P; n < m ==> P |] ==> P"
-  (fn major::prems => [cut_facts_tac [less_linear] 1,
-                       REPEAT(eresolve_tac ([disjE,major RS notE]@prems) 1)]);
+Addsimps [min_0L,min_0R,min_Suc_Suc];
--- a/src/HOL/Nat.thy	Wed Feb 12 15:43:50 1997 +0100
+++ b/src/HOL/Nat.thy	Wed Feb 12 18:53:59 1997 +0100
@@ -1,82 +1,13 @@
 (*  Title:      HOL/Nat.thy
     ID:         $Id$
-    Author:     Tobias Nipkow, Cambridge University Computer Laboratory
-    Copyright   1991  University of Cambridge
+    Author:     Tobias Nipkow
+    Copyright   1997 TU Muenchen
 
-Definition of types ind and nat.
-
-Type nat is defined as a set Nat over type ind.
+Nat is a partial order
 *)
 
-Nat = WF +
-
-(** type ind **)
-
-types
-  ind
-
-arities
-  ind :: term
-
-consts
-  Zero_Rep      :: ind
-  Suc_Rep       :: ind => ind
-
-rules
-  (*the axiom of infinity in 2 parts*)
-  inj_Suc_Rep           "inj(Suc_Rep)"
-  Suc_Rep_not_Zero_Rep  "Suc_Rep(x) ~= Zero_Rep"
-
-
-
-(** type nat **)
-
-(* type definition *)
-
-typedef (Nat)
-  nat = "lfp(%X. {Zero_Rep} Un (Suc_Rep``X))"   (lfp_def)
-
-instance
-  nat :: ord
-
-
-(* abstract constants and syntax *)
+Nat = NatDef +
 
-consts
-  "0"       :: nat                ("0")
-  Suc       :: nat => nat
-  nat_case  :: ['a, nat => 'a, nat] => 'a
-  pred_nat  :: "(nat * nat) set"
-  nat_rec   :: ['a, [nat, 'a] => 'a, nat] => 'a
-
-  Least     :: (nat=>bool) => nat    (binder "LEAST " 10)
-
-syntax
-  "1"       :: nat                ("1")
-  "2"       :: nat                ("2")
-
-translations
-   "1"  == "Suc 0"
-   "2"  == "Suc 1"
-  "case p of 0 => a | Suc y => b" == "nat_case a (%y.b) p"
-
-
-defs
-  Zero_def      "0 == Abs_Nat(Zero_Rep)"
-  Suc_def       "Suc == (%n. Abs_Nat(Suc_Rep(Rep_Nat(n))))"
-
-  (*nat operations and recursion*)
-  nat_case_def  "nat_case a f n == @z.  (n=0 --> z=a)  
-                                        & (!x. n=Suc x --> z=f(x))"
-  pred_nat_def  "pred_nat == {p. ? n. p = (n, Suc n)}"
-
-  less_def      "m<n == (m,n):trancl(pred_nat)"
-
-  le_def        "m<=(n::nat) == ~(n<m)"
-
-  nat_rec_def   "nat_rec c d ==
-                 wfrec pred_nat (%f. nat_case c (%m. d m (f m)))"
-  (*least number operator*)
-  Least_def     "Least P == @k. P(k) & (ALL j. j<k --> ~P(j))"
+instance nat :: order (le_refl,le_trans,le_anti_sym,nat_less_le)
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NatDef.ML	Wed Feb 12 18:53:59 1997 +0100
@@ -0,0 +1,684 @@
+(*  Title:      HOL/NatDef.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow, Cambridge University Computer Laboratory
+    Copyright   1991  University of Cambridge
+*)
+
+goal thy "mono(%X. {Zero_Rep} Un (Suc_Rep``X))";
+by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1));
+qed "Nat_fun_mono";
+
+val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski);
+
+(* Zero is a natural number -- this also justifies the type definition*)
+goal thy "Zero_Rep: Nat";
+by (stac Nat_unfold 1);
+by (rtac (singletonI RS UnI1) 1);
+qed "Zero_RepI";
+
+val prems = goal thy "i: Nat ==> Suc_Rep(i) : Nat";
+by (stac Nat_unfold 1);
+by (rtac (imageI RS UnI2) 1);
+by (resolve_tac prems 1);
+qed "Suc_RepI";
+
+(*** Induction ***)
+
+val major::prems = goal thy
+    "[| i: Nat;  P(Zero_Rep);   \
+\       !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |]  ==> P(i)";
+by (rtac ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1);
+by (fast_tac (!claset addIs prems) 1);
+qed "Nat_induct";
+
+val prems = goalw thy [Zero_def,Suc_def]
+    "[| P(0);   \
+\       !!k. P(k) ==> P(Suc(k)) |]  ==> P(n)";
+by (rtac (Rep_Nat_inverse RS subst) 1);   (*types force good instantiation*)
+by (rtac (Rep_Nat RS Nat_induct) 1);
+by (REPEAT (ares_tac prems 1
+     ORELSE eresolve_tac [Abs_Nat_inverse RS subst] 1));
+qed "nat_induct";
+
+(*Perform induction on n. *)
+fun nat_ind_tac a i = 
+    EVERY [res_inst_tac [("n",a)] nat_induct i,
+           rename_last_tac a ["1"] (i+1)];
+
+(*A special form of induction for reasoning about m<n and m-n*)
+val prems = goal thy
+    "[| !!x. P x 0;  \
+\       !!y. P 0 (Suc y);  \
+\       !!x y. [| P x y |] ==> P (Suc x) (Suc y)  \
+\    |] ==> P m n";
+by (res_inst_tac [("x","m")] spec 1);
+by (nat_ind_tac "n" 1);
+by (rtac allI 2);
+by (nat_ind_tac "x" 2);
+by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1));
+qed "diff_induct";
+
+(*Case analysis on the natural numbers*)
+val prems = goal thy 
+    "[| n=0 ==> P;  !!x. n = Suc(x) ==> P |] ==> P";
+by (subgoal_tac "n=0 | (EX x. n = Suc(x))" 1);
+by (fast_tac (!claset addSEs prems) 1);
+by (nat_ind_tac "n" 1);
+by (rtac (refl RS disjI1) 1);
+by (Fast_tac 1);
+qed "natE";
+
+(*** Isomorphisms: Abs_Nat and Rep_Nat ***)
+
+(*We can't take these properties as axioms, or take Abs_Nat==Inv(Rep_Nat),
+  since we assume the isomorphism equations will one day be given by Isabelle*)
+
+goal thy "inj(Rep_Nat)";
+by (rtac inj_inverseI 1);
+by (rtac Rep_Nat_inverse 1);
+qed "inj_Rep_Nat";
+
+goal thy "inj_onto Abs_Nat Nat";
+by (rtac inj_onto_inverseI 1);
+by (etac Abs_Nat_inverse 1);
+qed "inj_onto_Abs_Nat";
+
+(*** Distinctness of constructors ***)
+
+goalw thy [Zero_def,Suc_def] "Suc(m) ~= 0";
+by (rtac (inj_onto_Abs_Nat RS inj_onto_contraD) 1);
+by (rtac Suc_Rep_not_Zero_Rep 1);
+by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1));
+qed "Suc_not_Zero";
+
+bind_thm ("Zero_not_Suc", Suc_not_Zero RS not_sym);
+
+AddIffs [Suc_not_Zero,Zero_not_Suc];
+
+bind_thm ("Suc_neq_Zero", (Suc_not_Zero RS notE));
+val Zero_neq_Suc = sym RS Suc_neq_Zero;
+
+(** Injectiveness of Suc **)
+
+goalw thy [Suc_def] "inj(Suc)";
+by (rtac injI 1);
+by (dtac (inj_onto_Abs_Nat RS inj_ontoD) 1);
+by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1));
+by (dtac (inj_Suc_Rep RS injD) 1);
+by (etac (inj_Rep_Nat RS injD) 1);
+qed "inj_Suc";
+
+val Suc_inject = inj_Suc RS injD;
+
+goal thy "(Suc(m)=Suc(n)) = (m=n)";
+by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); 
+qed "Suc_Suc_eq";
+
+AddIffs [Suc_Suc_eq];
+
+goal thy "n ~= Suc(n)";
+by (nat_ind_tac "n" 1);
+by (ALLGOALS Asm_simp_tac);
+qed "n_not_Suc_n";
+
+bind_thm ("Suc_n_not_n", n_not_Suc_n RS not_sym);
+
+(*** nat_case -- the selection operator for nat ***)
+
+goalw thy [nat_case_def] "nat_case a f 0 = a";
+by (fast_tac (!claset addIs [select_equality]) 1);
+qed "nat_case_0";
+
+goalw thy [nat_case_def] "nat_case a f (Suc k) = f(k)";
+by (fast_tac (!claset addIs [select_equality]) 1);
+qed "nat_case_Suc";
+
+(** Introduction rules for 'pred_nat' **)
+
+goalw thy [pred_nat_def] "(n, Suc(n)) : pred_nat";
+by (Fast_tac 1);
+qed "pred_natI";
+
+val major::prems = goalw thy [pred_nat_def]
+    "[| p : pred_nat;  !!x n. [| p = (n, Suc(n)) |] ==> R \
+\    |] ==> R";
+by (rtac (major RS CollectE) 1);
+by (REPEAT (eresolve_tac ([asm_rl,exE]@prems) 1));
+qed "pred_natE";
+
+goalw thy [wf_def] "wf(pred_nat)";
+by (strip_tac 1);
+by (nat_ind_tac "x" 1);
+by (fast_tac (!claset addSEs [mp, pred_natE]) 2);
+by (fast_tac (!claset addSEs [mp, pred_natE]) 1);
+qed "wf_pred_nat";
+
+
+(*** nat_rec -- by wf recursion on pred_nat ***)
+
+(* The unrolling rule for nat_rec *)
+goal thy
+   "(%n. nat_rec c d n) = wfrec pred_nat (%f. nat_case ?c (%m. ?d m (f m)))";
+  by (simp_tac (HOL_ss addsimps [nat_rec_def]) 1);
+bind_thm("nat_rec_unfold", wf_pred_nat RS 
+                            ((result() RS eq_reflection) RS def_wfrec));
+
+(*---------------------------------------------------------------------------
+ * Old:
+ * bind_thm ("nat_rec_unfold", (wf_pred_nat RS (nat_rec_def RS def_wfrec))); 
+ *---------------------------------------------------------------------------*)
+
+(** conversion rules **)
+
+goal thy "nat_rec c h 0 = c";
+by (rtac (nat_rec_unfold RS trans) 1);
+by (simp_tac (!simpset addsimps [nat_case_0]) 1);
+qed "nat_rec_0";
+
+goal thy "nat_rec c h (Suc n) = h n (nat_rec c h n)";
+by (rtac (nat_rec_unfold RS trans) 1);
+by (simp_tac (!simpset addsimps [nat_case_Suc, pred_natI, cut_apply]) 1);
+qed "nat_rec_Suc";
+
+(*These 2 rules ease the use of primitive recursion.  NOTE USE OF == *)
+val [rew] = goal thy
+    "[| !!n. f(n) == nat_rec c h n |] ==> f(0) = c";
+by (rewtac rew);
+by (rtac nat_rec_0 1);
+qed "def_nat_rec_0";
+
+val [rew] = goal thy
+    "[| !!n. f(n) == nat_rec c h n |] ==> f(Suc(n)) = h n (f n)";
+by (rewtac rew);
+by (rtac nat_rec_Suc 1);
+qed "def_nat_rec_Suc";
+
+fun nat_recs def =
+      [standard (def RS def_nat_rec_0),
+       standard (def RS def_nat_rec_Suc)];
+
+
+(*** Basic properties of "less than" ***)
+
+(** Introduction properties **)
+
+val prems = goalw thy [less_def] "[| i<j;  j<k |] ==> i<(k::nat)";
+by (rtac (trans_trancl RS transD) 1);
+by (resolve_tac prems 1);
+by (resolve_tac prems 1);
+qed "less_trans";
+
+goalw thy [less_def] "n < Suc(n)";
+by (rtac (pred_natI RS r_into_trancl) 1);
+qed "lessI";
+AddIffs [lessI];
+
+(* i<j ==> i<Suc(j) *)
+bind_thm("less_SucI", lessI RSN (2, less_trans));
+Addsimps [less_SucI];
+
+goal thy "0 < Suc(n)";
+by (nat_ind_tac "n" 1);
+by (rtac lessI 1);
+by (etac less_trans 1);
+by (rtac lessI 1);
+qed "zero_less_Suc";
+AddIffs [zero_less_Suc];
+
+(** Elimination properties **)
+
+val prems = goalw thy [less_def] "n<m ==> ~ m<(n::nat)";
+by (fast_tac (!claset addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1);
+qed "less_not_sym";
+
+(* [| n<m; m<n |] ==> R *)
+bind_thm ("less_asym", (less_not_sym RS notE));
+
+goalw thy [less_def] "~ n<(n::nat)";
+by (rtac notI 1);
+by (etac (wf_pred_nat RS wf_trancl RS wf_irrefl) 1);
+qed "less_not_refl";
+
+(* n<n ==> R *)
+bind_thm ("less_irrefl", (less_not_refl RS notE));
+
+goal thy "!!m. n<m ==> m ~= (n::nat)";
+by (fast_tac (!claset addEs [less_irrefl]) 1);
+qed "less_not_refl2";
+
+
+val major::prems = goalw thy [less_def]
+    "[| i<k;  k=Suc(i) ==> P;  !!j. [| i<j;  k=Suc(j) |] ==> P \
+\    |] ==> P";
+by (rtac (major RS tranclE) 1);
+by (REPEAT_FIRST (bound_hyp_subst_tac ORELSE'
+                  eresolve_tac (prems@[pred_natE, Pair_inject])));
+by (rtac refl 1);
+qed "lessE";
+
+goal thy "~ n<0";
+by (rtac notI 1);
+by (etac lessE 1);
+by (etac Zero_neq_Suc 1);
+by (etac Zero_neq_Suc 1);
+qed "not_less0";
+
+AddIffs [not_less0];
+
+(* n<0 ==> R *)
+bind_thm ("less_zeroE", not_less0 RS notE);
+
+val [major,less,eq] = goal thy
+    "[| m < Suc(n);  m<n ==> P;  m=n ==> P |] ==> P";
+by (rtac (major RS lessE) 1);
+by (rtac eq 1);
+by (Fast_tac 1);
+by (rtac less 1);
+by (Fast_tac 1);
+qed "less_SucE";
+
+goal thy "(m < Suc(n)) = (m < n | m = n)";
+by (fast_tac (!claset addSIs [lessI]
+                      addEs  [less_trans, less_SucE]) 1);
+qed "less_Suc_eq";
+
+val prems = goal thy "m<n ==> n ~= 0";
+by (res_inst_tac [("n","n")] natE 1);
+by (cut_facts_tac prems 1);
+by (ALLGOALS Asm_full_simp_tac);
+qed "gr_implies_not0";
+Addsimps [gr_implies_not0];
+
+qed_goal "zero_less_eq" thy "0 < n = (n ~= 0)" (fn _ => [
+        rtac iffI 1,
+        etac gr_implies_not0 1,
+        rtac natE 1,
+        contr_tac 1,
+        etac ssubst 1,
+        rtac zero_less_Suc 1]);
+
+(** Inductive (?) properties **)
+
+val [prem] = goal thy "Suc(m) < n ==> m<n";
+by (rtac (prem RS rev_mp) 1);
+by (nat_ind_tac "n" 1);
+by (ALLGOALS (fast_tac (!claset addSIs [lessI RS less_SucI]
+                                addEs  [less_trans, lessE])));
+qed "Suc_lessD";
+
+val [major,minor] = goal thy 
+    "[| Suc(i)<k;  !!j. [| i<j;  k=Suc(j) |] ==> P \
+\    |] ==> P";
+by (rtac (major RS lessE) 1);
+by (etac (lessI RS minor) 1);
+by (etac (Suc_lessD RS minor) 1);
+by (assume_tac 1);
+qed "Suc_lessE";
+
+goal thy "!!m n. Suc(m) < Suc(n) ==> m<n";
+by (fast_tac (!claset addEs [lessE, Suc_lessD] addIs [lessI]) 1);
+qed "Suc_less_SucD";
+
+goal thy "!!m n. m<n ==> Suc(m) < Suc(n)";
+by (etac rev_mp 1);
+by (nat_ind_tac "n" 1);
+by (ALLGOALS (fast_tac (!claset addSIs [lessI]
+                                addEs  [less_trans, lessE])));
+qed "Suc_mono";
+
+
+goal thy "(Suc(m) < Suc(n)) = (m<n)";
+by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]);
+qed "Suc_less_eq";
+Addsimps [Suc_less_eq];
+
+goal thy "~(Suc(n) < n)";
+by (fast_tac (!claset addEs [Suc_lessD RS less_irrefl]) 1);
+qed "not_Suc_n_less_n";
+Addsimps [not_Suc_n_less_n];
+
+goal thy "!!i. i<j ==> j<k --> Suc i < k";
+by (nat_ind_tac "k" 1);
+by (ALLGOALS (asm_simp_tac (!simpset)));
+by (asm_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
+by (fast_tac (!claset addDs [Suc_lessD]) 1);
+qed_spec_mp "less_trans_Suc";
+
+(*"Less than" is a linear ordering*)
+goal thy "m<n | m=n | n<(m::nat)";
+by (nat_ind_tac "m" 1);
+by (nat_ind_tac "n" 1);
+by (rtac (refl RS disjI1 RS disjI2) 1);
+by (rtac (zero_less_Suc RS disjI1) 1);
+by (fast_tac (!claset addIs [lessI, Suc_mono, less_SucI] addEs [lessE]) 1);
+qed "less_linear";
+
+qed_goal "nat_less_cases" thy 
+   "[| (m::nat)<n ==> P n m; m=n ==> P n m; n<m ==> P n m |] ==> P n m"
+( fn prems =>
+        [
+        (res_inst_tac [("m1","n"),("n1","m")] (less_linear RS disjE) 1),
+        (etac disjE 2),
+        (etac (hd (tl (tl prems))) 1),
+        (etac (sym RS hd (tl prems)) 1),
+        (etac (hd prems) 1)
+        ]);
+
+(*Can be used with less_Suc_eq to get n=m | n<m *)
+goal thy "(~ m < n) = (n < Suc(m))";
+by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
+by (ALLGOALS Asm_simp_tac);
+qed "not_less_eq";
+
+(*Complete induction, aka course-of-values induction*)
+val prems = goalw thy [less_def]
+    "[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |]  ==>  P(n)";
+by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1);
+by (eresolve_tac prems 1);
+qed "less_induct";
+
+qed_goal "nat_induct2" thy 
+"[| P 0; P 1; !!k. P k ==> P (Suc (Suc k)) |] ==> P n" (fn prems => [
+	cut_facts_tac prems 1,
+	rtac less_induct 1,
+	res_inst_tac [("n","n")] natE 1,
+	 hyp_subst_tac 1,
+	 atac 1,
+	hyp_subst_tac 1,
+	res_inst_tac [("n","x")] natE 1,
+	 hyp_subst_tac 1,
+	 atac 1,
+	hyp_subst_tac 1,
+	resolve_tac prems 1,
+	dtac spec 1,
+	etac mp 1,
+	rtac (lessI RS less_trans) 1,
+	rtac (lessI RS Suc_mono) 1]);
+
+(*** Properties of <= ***)
+
+goalw thy [le_def] "(m <= n) = (m < Suc n)";
+by (rtac not_less_eq 1);
+qed "le_eq_less_Suc";
+
+goalw thy [le_def] "0 <= n";
+by (rtac not_less0 1);
+qed "le0";
+
+goalw thy [le_def] "~ Suc n <= n";
+by (Simp_tac 1);
+qed "Suc_n_not_le_n";
+
+goalw thy [le_def] "(i <= 0) = (i = 0)";
+by (nat_ind_tac "i" 1);
+by (ALLGOALS Asm_simp_tac);
+qed "le_0_eq";
+
+Addsimps [(*less_Suc_eq, makes simpset non-confluent*) le0, le_0_eq,
+          Suc_n_not_le_n,
+          n_not_Suc_n, Suc_n_not_n,
+          nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc];
+
+(*
+goal thy "(Suc m < n | Suc m = n) = (m < n)";
+by (stac (less_Suc_eq RS sym) 1);
+by (rtac Suc_less_eq 1);
+qed "Suc_le_eq";
+
+this could make the simpset (with less_Suc_eq added again) more confluent,
+but less_Suc_eq makes additional problems with terms of the form 0 < Suc (...)
+*)
+
+(*Prevents simplification of f and g: much faster*)
+qed_goal "nat_case_weak_cong" thy
+  "m=n ==> nat_case a f m = nat_case a f n"
+  (fn [prem] => [rtac (prem RS arg_cong) 1]);
+
+qed_goal "nat_rec_weak_cong" thy
+  "m=n ==> nat_rec a f m = nat_rec a f n"
+  (fn [prem] => [rtac (prem RS arg_cong) 1]);
+
+qed_goal "expand_nat_case" thy
+  "P(nat_case z s n) = ((n=0 --> P z) & (!m. n = Suc m --> P(s m)))"
+  (fn _ => [nat_ind_tac "n" 1, ALLGOALS Asm_simp_tac]);
+
+val prems = goalw thy [le_def] "~n<m ==> m<=(n::nat)";
+by (resolve_tac prems 1);
+qed "leI";
+
+val prems = goalw thy [le_def] "m<=n ==> ~ n < (m::nat)";
+by (resolve_tac prems 1);
+qed "leD";
+
+val leE = make_elim leD;
+
+goal thy "(~n<m) = (m<=(n::nat))";
+by (fast_tac (!claset addIs [leI] addEs [leE]) 1);
+qed "not_less_iff_le";
+
+goalw thy [le_def] "!!m. ~ m <= n ==> n<(m::nat)";
+by (Fast_tac 1);
+qed "not_leE";
+
+goalw thy [le_def] "!!m. m < n ==> Suc(m) <= n";
+by (simp_tac (!simpset addsimps [less_Suc_eq]) 1);
+by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
+qed "lessD";
+
+goalw thy [le_def] "!!m. Suc(m) <= n ==> m <= n";
+by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
+qed "Suc_leD";
+
+(* stronger version of Suc_leD *)
+goalw thy [le_def] 
+        "!!m. Suc m <= n ==> m < n";
+by (asm_full_simp_tac (!simpset addsimps [less_Suc_eq]) 1);
+by (cut_facts_tac [less_linear] 1);
+by (Fast_tac 1);
+qed "Suc_le_lessD";
+
+goal thy "(Suc m <= n) = (m < n)";
+by (fast_tac (!claset addIs [lessD, Suc_le_lessD]) 1);
+qed "Suc_le_eq";
+
+goalw thy [le_def] "!!m. m <= n ==> m <= Suc n";
+by (fast_tac (!claset addDs [Suc_lessD]) 1);
+qed "le_SucI";
+Addsimps[le_SucI];
+
+bind_thm ("le_Suc", not_Suc_n_less_n RS leI);
+
+goalw thy [le_def] "!!m. m < n ==> m <= (n::nat)";
+by (fast_tac (!claset addEs [less_asym]) 1);
+qed "less_imp_le";
+
+goalw thy [le_def] "!!m. m <= n ==> m < n | m=(n::nat)";
+by (cut_facts_tac [less_linear] 1);
+by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
+qed "le_imp_less_or_eq";
+
+goalw thy [le_def] "!!m. m<n | m=n ==> m <=(n::nat)";
+by (cut_facts_tac [less_linear] 1);
+by (fast_tac (!claset addEs [less_irrefl,less_asym]) 1);
+by (flexflex_tac);
+qed "less_or_eq_imp_le";
+
+goal thy "(x <= (y::nat)) = (x < y | x=y)";
+by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1));
+qed "le_eq_less_or_eq";
+
+goal thy "n <= (n::nat)";
+by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1);
+qed "le_refl";
+
+val prems = goal thy "!!i. [| i <= j; j < k |] ==> i < (k::nat)";
+by (dtac le_imp_less_or_eq 1);
+by (fast_tac (!claset addIs [less_trans]) 1);
+qed "le_less_trans";
+
+goal thy "!!i. [| i < j; j <= k |] ==> i < (k::nat)";
+by (dtac le_imp_less_or_eq 1);
+by (fast_tac (!claset addIs [less_trans]) 1);
+qed "less_le_trans";
+
+goal thy "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)";
+by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
+          rtac less_or_eq_imp_le, fast_tac (!claset addIs [less_trans])]);
+qed "le_trans";
+
+val prems = goal thy "!!m. [| m <= n; n <= m |] ==> m = (n::nat)";
+by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
+          fast_tac (!claset addEs [less_irrefl,less_asym])]);
+qed "le_anti_sym";
+
+goal thy "(Suc(n) <= Suc(m)) = (n <= m)";
+by (simp_tac (!simpset addsimps [le_eq_less_or_eq]) 1);
+qed "Suc_le_mono";
+
+AddIffs [Suc_le_mono];
+
+(* Axiom 'order_le_less' of class 'order': *)
+goal thy "(m::nat) < n = (m <= n & m ~= n)";
+br iffI 1;
+ br conjI 1;
+  be less_imp_le 1;
+ be (less_not_refl2 RS not_sym) 1;
+by(fast_tac (!claset addSDs [le_imp_less_or_eq]) 1);
+qed "nat_less_le";
+
+(** LEAST -- the least number operator **)
+
+val [prem1,prem2] = goalw thy [Least_def]
+    "[| P(k::nat);  !!x. x<k ==> ~P(x) |] ==> (LEAST x.P(x)) = k";
+by (rtac select_equality 1);
+by (fast_tac (!claset addSIs [prem1,prem2]) 1);
+by (cut_facts_tac [less_linear] 1);
+by (fast_tac (!claset addSIs [prem1] addSDs [prem2]) 1);
+qed "Least_equality";
+
+val [prem] = goal thy "P(k::nat) ==> P(LEAST x.P(x))";
+by (rtac (prem RS rev_mp) 1);
+by (res_inst_tac [("n","k")] less_induct 1);
+by (rtac impI 1);
+by (rtac classical 1);
+by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
+by (assume_tac 1);
+by (assume_tac 2);
+by (Fast_tac 1);
+qed "LeastI";
+
+(*Proof is almost identical to the one above!*)
+val [prem] = goal thy "P(k::nat) ==> (LEAST x.P(x)) <= k";
+by (rtac (prem RS rev_mp) 1);
+by (res_inst_tac [("n","k")] less_induct 1);
+by (rtac impI 1);
+by (rtac classical 1);
+by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1);
+by (assume_tac 1);
+by (rtac le_refl 2);
+by (fast_tac (!claset addIs [less_imp_le,le_trans]) 1);
+qed "Least_le";
+
+val [prem] = goal thy "k < (LEAST x.P(x)) ==> ~P(k::nat)";
+by (rtac notI 1);
+by (etac (rewrite_rule [le_def] Least_le RS notE) 1);
+by (rtac prem 1);
+qed "not_less_Least";
+
+qed_goalw "Least_Suc" thy [Least_def]
+ "!!P. [| ? n. P(Suc n); ~ P 0 |] ==> (LEAST n. P n) = Suc (LEAST m. P(Suc m))"
+ (fn _ => [
+        rtac select_equality 1,
+        fold_goals_tac [Least_def],
+        safe_tac (!claset addSEs [LeastI]),
+        rename_tac "j" 1,
+        res_inst_tac [("n","j")] natE 1,
+        Fast_tac 1,
+        fast_tac (!claset addDs [Suc_less_SucD, not_less_Least]) 1,
+        rename_tac "k n" 1,
+        res_inst_tac [("n","k")] natE 1,
+        Fast_tac 1,
+        hyp_subst_tac 1,
+        rewtac Least_def,
+        rtac (select_equality RS arg_cong RS sym) 1,
+        safe_tac (!claset),
+        dtac Suc_mono 1,
+        Fast_tac 1,
+        cut_facts_tac [less_linear] 1,
+        safe_tac (!claset),
+        atac 2,
+        Fast_tac 2,
+        dtac Suc_mono 1,
+        Fast_tac 1]);
+
+
+(*** Instantiation of transitivity prover ***)
+
+structure Less_Arith =
+struct
+val nat_leI = leI;
+val nat_leD = leD;
+val lessI = lessI;
+val zero_less_Suc = zero_less_Suc;
+val less_reflE = less_irrefl;
+val less_zeroE = less_zeroE;
+val less_incr = Suc_mono;
+val less_decr = Suc_less_SucD;
+val less_incr_rhs = Suc_mono RS Suc_lessD;
+val less_decr_lhs = Suc_lessD;
+val less_trans_Suc = less_trans_Suc;
+val leI = lessD RS (Suc_le_mono RS iffD1);
+val not_lessI = leI RS leD
+val not_leI = prove_goal thy "!!m::nat. n < m ==> ~ m <= n"
+  (fn _ => [etac swap2 1, etac leD 1]);
+val eqI = prove_goal thy "!!m. [| m < Suc n; n < Suc m |] ==> m=n"
+  (fn _ => [etac less_SucE 1,
+            fast_tac (HOL_cs addSDs [Suc_less_SucD] addSEs [less_irrefl]
+                             addDs [less_trans_Suc]) 1,
+            atac 1]);
+val leD = le_eq_less_Suc RS iffD1;
+val not_lessD = nat_leI RS leD;
+val not_leD = not_leE
+val eqD1 = prove_goal thy  "!!n. m = n ==> m < Suc n"
+ (fn _ => [etac subst 1, rtac lessI 1]);
+val eqD2 = sym RS eqD1;
+
+fun is_zero(t) =  t = Const("0",Type("nat",[]));
+
+fun nnb T = T = Type("fun",[Type("nat",[]),
+                            Type("fun",[Type("nat",[]),
+                                        Type("bool",[])])])
+
+fun decomp_Suc(Const("Suc",_)$t) = let val (a,i) = decomp_Suc t in (a,i+1) end
+  | decomp_Suc t = (t,0);
+
+fun decomp2(rel,T,lhs,rhs) =
+  if not(nnb T) then None else
+  let val (x,i) = decomp_Suc lhs
+      val (y,j) = decomp_Suc rhs
+  in case rel of
+       "op <"  => Some(x,i,"<",y,j)
+     | "op <=" => Some(x,i,"<=",y,j)
+     | "op ="  => Some(x,i,"=",y,j)
+     | _       => None
+  end;
+
+fun negate(Some(x,i,rel,y,j)) = Some(x,i,"~"^rel,y,j)
+  | negate None = None;
+
+fun decomp(_$(Const(rel,T)$lhs$rhs)) = decomp2(rel,T,lhs,rhs)
+  | decomp(_$(Const("not",_)$(Const(rel,T)$lhs$rhs))) =
+      negate(decomp2(rel,T,lhs,rhs))
+  | decomp _ = None
+
+end;
+
+structure Trans_Tac = Trans_Tac_Fun(Less_Arith);
+
+open Trans_Tac;
+
+(*** eliminates ~= in premises, which trans_tac cannot deal with ***)
+qed_goal "nat_neqE" thy
+  "[| (m::nat) ~= n; m < n ==> P; n < m ==> P |] ==> P"
+  (fn major::prems => [cut_facts_tac [less_linear] 1,
+                       REPEAT(eresolve_tac ([disjE,major RS notE]@prems) 1)]);
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NatDef.thy	Wed Feb 12 18:53:59 1997 +0100
@@ -0,0 +1,77 @@
+(*  Title:      HOL/NatDef.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow, Cambridge University Computer Laboratory
+    Copyright   1991  University of Cambridge
+
+Definition of types ind and nat.
+
+Type nat is defined as a set Nat over type ind.
+*)
+
+NatDef = WF +
+
+(** type ind **)
+
+types
+  ind
+
+arities
+  ind :: term
+
+consts
+  Zero_Rep      :: ind
+  Suc_Rep       :: ind => ind
+
+rules
+  (*the axiom of infinity in 2 parts*)
+  inj_Suc_Rep           "inj(Suc_Rep)"
+  Suc_Rep_not_Zero_Rep  "Suc_Rep(x) ~= Zero_Rep"
+
+
+
+(** type nat **)
+
+(* type definition *)
+
+typedef (Nat)
+  nat = "lfp(%X. {Zero_Rep} Un (Suc_Rep``X))"   (lfp_def)
+
+instance
+  nat :: ord
+
+
+(* abstract constants and syntax *)
+
+consts
+  "0"       :: nat                ("0")
+  Suc       :: nat => nat
+  nat_case  :: ['a, nat => 'a, nat] => 'a
+  pred_nat  :: "(nat * nat) set"
+  nat_rec   :: ['a, [nat, 'a] => 'a, nat] => 'a
+
+syntax
+  "1"       :: nat                ("1")
+  "2"       :: nat                ("2")
+
+translations
+   "1"  == "Suc 0"
+   "2"  == "Suc 1"
+  "case p of 0 => a | Suc y => b" == "nat_case a (%y.b) p"
+
+
+defs
+  Zero_def      "0 == Abs_Nat(Zero_Rep)"
+  Suc_def       "Suc == (%n. Abs_Nat(Suc_Rep(Rep_Nat(n))))"
+
+  (*nat operations and recursion*)
+  nat_case_def  "nat_case a f n == @z.  (n=0 --> z=a)  
+                                        & (!x. n=Suc x --> z=f(x))"
+  pred_nat_def  "pred_nat == {p. ? n. p = (n, Suc n)}"
+
+  less_def      "m<n == (m,n):trancl(pred_nat)"
+
+  le_def        "m<=(n::nat) == ~(n<m)"
+
+  nat_rec_def   "nat_rec c d ==
+                 wfrec pred_nat (%f. nat_case c (%m. d m (f m)))"
+end
--- a/src/HOL/Ord.ML	Wed Feb 12 15:43:50 1997 +0100
+++ b/src/HOL/Ord.ML	Wed Feb 12 18:53:59 1997 +0100
@@ -6,7 +6,7 @@
 The type class for ordered types
 *)
 
-open Ord;
+(** mono **)
 
 val [prem] = goalw Ord.thy [mono_def]
     "[| !!A B. A <= B ==> f(A) <= f(B) |] ==> mono(f)";
@@ -19,3 +19,32 @@
 by (rtac minor 1);
 qed "monoD";
 
+
+section "Orders";
+
+AddIffs [order_refl];
+
+goal Ord.thy "~ x < (x::'a::order)";
+by(simp_tac (!simpset addsimps [order_less_le]) 1);
+qed "order_less_irrefl";
+AddIffs [order_less_irrefl];
+
+goal thy "(x::'a::order) <= y = (x < y | x = y)";
+by(simp_tac (!simpset addsimps [order_less_le]) 1);
+by(Fast_tac 1);
+qed "order_le_less";
+
+(** min **)
+
+goalw thy [min_def] "!!least. (!!x. least <= x) ==> min least x = least";
+by(split_tac [expand_if] 1);
+by(Asm_simp_tac 1);
+qed "min_leastL";
+
+val prems = goalw thy [min_def]
+ "(!!x::'a::order. least <= x) ==> min x least = least";
+by(cut_facts_tac prems 1);
+by(split_tac [expand_if] 1);
+by(Asm_simp_tac 1);
+by(fast_tac (!claset addEs [order_antisym]) 1);
+qed "min_leastR";
--- a/src/HOL/Ord.thy	Wed Feb 12 15:43:50 1997 +0100
+++ b/src/HOL/Ord.thy	Wed Feb 12 18:53:59 1997 +0100
@@ -17,6 +17,8 @@
   mono          :: ['a::ord => 'b::ord] => bool       (*monotonicity*)
   min, max      :: ['a::ord, 'a] => 'a
 
+  Least         :: ('a::ord=>bool) => 'a             (binder "LEAST " 10)
+
 syntax
   "op <"        :: ['a::ord, 'a] => bool             ("op <")
   "op <="       :: ['a::ord, 'a] => bool             ("op <=")
@@ -29,5 +31,13 @@
   mono_def      "mono(f) == (!A B. A <= B --> f(A) <= f(B))"
   min_def       "min a b == (if a <= b then a else b)"
   max_def       "max a b == (if a <= b then b else a)"
+  Least_def     "Least P == @x. P(x) & (ALL y. y<x --> ~P(y))"
+
+
+axclass order < ord
+  order_refl    "x <= x"
+  order_trans   "[| x <= y; y <= z |] ==> x <= z"
+  order_antisym "[| x <= y; y <= x |] ==> x = y"
+  order_less_le "x < y = (x <= y & x ~= y)"
 
 end
--- a/src/HOL/Set.ML	Wed Feb 12 15:43:50 1997 +0100
+++ b/src/HOL/Set.ML	Wed Feb 12 18:53:59 1997 +0100
@@ -341,6 +341,18 @@
 qed "bex_empty";
 Addsimps [ball_empty, bex_empty];
 
+goalw Set.thy [Ball_def] "(!x:A.False) = (A = {})";
+by(Fast_tac 1);
+qed "ball_False";
+Addsimps [ball_False];
+
+(* The dual is probably not helpful:
+goal Set.thy "(? x:A.True) = (A ~= {})";
+by(Fast_tac 1);
+qed "bex_True";
+Addsimps [bex_True];
+*)
+
 
 section "Augmenting a set -- insert";