rename Nat_Infinity (inat) to Extended_Nat (enat)
authorhoelzl
Tue, 19 Jul 2011 14:35:44 +0200
changeset 43919 a7e4fb1a0502
parent 43903 1e2aa420c660
child 43920 cedb5cb948fd
rename Nat_Infinity (inat) to Extended_Nat (enat)
src/HOL/HOLCF/FOCUS/Fstreams.thy
src/HOL/HOLCF/FOCUS/Stream_adm.thy
src/HOL/HOLCF/IsaMakefile
src/HOL/HOLCF/Library/Stream.thy
src/HOL/IsaMakefile
src/HOL/Library/Extended_Nat.thy
src/HOL/Library/Library.thy
src/HOL/Library/Nat_Infinity.thy
--- a/src/HOL/HOLCF/FOCUS/Fstreams.thy	Tue Jul 19 11:15:38 2011 +0200
+++ b/src/HOL/HOLCF/FOCUS/Fstreams.thy	Tue Jul 19 14:35:44 2011 +0200
@@ -55,7 +55,7 @@
 by (simp add: fsingleton_def2)
 
 lemma slen_fsingleton[simp]: "#(<a>) = Fin 1"
-by (simp add: fsingleton_def2 inat_defs)
+by (simp add: fsingleton_def2 enat_defs)
 
 lemma slen_fstreams[simp]: "#(<a> ooo s) = iSuc (#s)"
 by (simp add: fsingleton_def2)
@@ -83,11 +83,11 @@
 lemma jth_n[simp]: "Fin n = #s ==> jth n (s ooo <a>) = a"
 apply (simp add: jth_def, auto)
 apply (simp add: i_th_def rt_sconc1)
-by (simp add: inat_defs split: inat.splits)
+by (simp add: enat_defs split: enat.splits)
 
 lemma last_sconc[simp]: "Fin n = #s ==> last (s ooo <a>) = a"
 apply (simp add: last_def)
-apply (simp add: inat_defs split:inat.splits)
+apply (simp add: enat_defs split:enat.splits)
 by (drule sym, auto)
 
 lemma last_fsingleton[simp]: "last (<a>) = a"
@@ -97,13 +97,13 @@
 by (simp add: first_def jth_def)
 
 lemma last_UU[simp]:"last UU = undefined"
-by (simp add: last_def jth_def inat_defs)
+by (simp add: last_def jth_def enat_defs)
 
 lemma last_infinite[simp]:"#s = Infty ==> last s = undefined"
 by (simp add: last_def)
 
 lemma jth_slen_lemma1:"n <= k & Fin n = #s ==> jth k s = undefined"
-by (simp add: jth_def inat_defs split:inat.splits, auto)
+by (simp add: jth_def enat_defs split:enat.splits, auto)
 
 lemma jth_UU[simp]:"jth n UU = undefined" 
 by (simp add: jth_def)
--- a/src/HOL/HOLCF/FOCUS/Stream_adm.thy	Tue Jul 19 11:15:38 2011 +0200
+++ b/src/HOL/HOLCF/FOCUS/Stream_adm.thy	Tue Jul 19 14:35:44 2011 +0200
@@ -74,7 +74,7 @@
 apply (   erule spec)
 apply (  assumption)
 apply ( assumption)
-apply (metis inat_ord_code(4) slen_infinite)
+apply (metis enat_ord_code(4) slen_infinite)
 done
 
 (* should be without reference to stream length? *)
--- a/src/HOL/HOLCF/IsaMakefile	Tue Jul 19 11:15:38 2011 +0200
+++ b/src/HOL/HOLCF/IsaMakefile	Tue Jul 19 14:35:44 2011 +0200
@@ -134,7 +134,7 @@
 HOLCF-ex: HOLCF $(LOG)/HOLCF-ex.gz
 
 $(LOG)/HOLCF-ex.gz: $(OUT)/HOLCF \
-  ../Library/Nat_Infinity.thy \
+  ../Library/Extended_Nat.thy \
   ex/Concurrency_Monad.thy \
   ex/Dagstuhl.thy \
   ex/Dnat.thy \
--- a/src/HOL/HOLCF/Library/Stream.thy	Tue Jul 19 11:15:38 2011 +0200
+++ b/src/HOL/HOLCF/Library/Stream.thy	Tue Jul 19 14:35:44 2011 +0200
@@ -5,7 +5,7 @@
 header {* General Stream domain *}
 
 theory Stream
-imports HOLCF "~~/src/HOL/Library/Nat_Infinity"
+imports HOLCF "~~/src/HOL/Library/Extended_Nat"
 begin
 
 default_sort pcpo
@@ -22,7 +22,7 @@
                                      If p\<cdot>x then x && h\<cdot>p\<cdot>xs else h\<cdot>p\<cdot>xs)"
 
 definition
-  slen :: "'a stream \<Rightarrow> inat"  ("#_" [1000] 1000) where
+  slen :: "'a stream \<Rightarrow> enat"  ("#_" [1000] 1000) where
   "#s = (if stream_finite s then Fin (LEAST n. stream_take n\<cdot>s = s) else \<infinity>)"
 
 
@@ -327,7 +327,7 @@
 section "slen"
 
 lemma slen_empty [simp]: "#\<bottom> = 0"
-by (simp add: slen_def stream.finite_def zero_inat_def Least_equality)
+by (simp add: slen_def stream.finite_def zero_enat_def Least_equality)
 
 lemma slen_scons [simp]: "x ~= \<bottom> ==> #(x&&xs) = iSuc (#xs)"
 apply (case_tac "stream_finite (x && xs)")
@@ -361,7 +361,7 @@
 
 lemma slen_scons_eq_rev: "(#x < Fin (Suc (Suc n))) = (!a y. x ~= a && y |  a = \<bottom> |  #y < Fin (Suc n))"
  apply (cases x, auto)
-   apply (simp add: zero_inat_def)
+   apply (simp add: zero_enat_def)
   apply (case_tac "#stream") apply (simp_all add: iSuc_Fin)
  apply (case_tac "#stream") apply (simp_all add: iSuc_Fin)
 done
@@ -445,7 +445,7 @@
 
 lemma slen_rt_mult [rule_format]: "!x. Fin (i + j) <= #x --> Fin j <= #(iterate i$rt$x)"
 apply (induct i, auto)
-apply (case_tac "x=UU", auto simp add: zero_inat_def)
+apply (case_tac "x=UU", auto simp add: zero_enat_def)
 apply (drule stream_exhaust_eq [THEN iffD1], auto)
 apply (erule_tac x="y" in allE, auto)
 apply (simp add: not_le) apply (case_tac "#y") apply (simp_all add: iSuc_Fin)
@@ -455,7 +455,7 @@
   "!(x::'a::flat stream) y. Fin n <= #x --> x << y --> stream_take n\<cdot>x = stream_take n\<cdot>y"
 apply (induct_tac n, auto)
 apply (case_tac "x=UU", auto)
-apply (simp add: zero_inat_def)
+apply (simp add: zero_enat_def)
 apply (simp add: Suc_ile_eq)
 apply (case_tac "y=UU", clarsimp)
 apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)+
@@ -566,11 +566,11 @@
  apply (subgoal_tac "#(i_rt n s)=0")
   apply (case_tac "stream_take n$s = s",simp+)
   apply (insert slen_take_eq [rule_format,of n s],simp)
-  apply (cases "#s") apply (simp_all add: zero_inat_def)
+  apply (cases "#s") apply (simp_all add: zero_enat_def)
   apply (simp add: slen_take_eq)
   apply (cases "#s")
   using i_rt_take_lemma1 [of n s]
-  apply (simp_all add: zero_inat_def)
+  apply (simp_all add: zero_enat_def)
   done
 
 lemma i_rt_lemma_slen: "#s=Fin n ==> i_rt n s = UU"
@@ -585,20 +585,20 @@
                             #(stream_take n$x) = Fin t & #(i_rt n x)= Fin j
                                               --> Fin (j + t) = #x"
 apply (induct n, auto)
- apply (simp add: zero_inat_def)
+ apply (simp add: zero_enat_def)
 apply (case_tac "x=UU",auto)
- apply (simp add: zero_inat_def)
+ apply (simp add: zero_enat_def)
 apply (drule stream_exhaust_eq [THEN iffD1],clarsimp)
 apply (subgoal_tac "EX k. Fin k = #y",clarify)
  apply (erule_tac x="k" in allE)
  apply (erule_tac x="y" in allE,auto)
  apply (erule_tac x="THE p. Suc p = t" in allE,auto)
-   apply (simp add: iSuc_def split: inat.splits)
-  apply (simp add: iSuc_def split: inat.splits)
+   apply (simp add: iSuc_def split: enat.splits)
+  apply (simp add: iSuc_def split: enat.splits)
   apply (simp only: the_equality)
- apply (simp add: iSuc_def split: inat.splits)
+ apply (simp add: iSuc_def split: enat.splits)
  apply force
-apply (simp add: iSuc_def split: inat.splits)
+apply (simp add: iSuc_def split: enat.splits)
 done
 
 lemma take_i_rt_len:
@@ -690,13 +690,13 @@
 (* ----------------------------------------------------------------------- *)
 
 lemma UU_sconc [simp]: " UU ooo s = s "
-by (simp add: sconc_def zero_inat_def)
+by (simp add: sconc_def zero_enat_def)
 
 lemma scons_neq_UU: "a~=UU ==> a && s ~=UU"
 by auto
 
 lemma singleton_sconc [rule_format, simp]: "x~=UU --> (x && UU) ooo y = x && y"
-apply (simp add: sconc_def zero_inat_def iSuc_def split: inat.splits, auto)
+apply (simp add: sconc_def zero_enat_def iSuc_def split: enat.splits, auto)
 apply (rule someI2_ex,auto)
  apply (rule_tac x="x && y" in exI,auto)
 apply (simp add: i_rt_Suc_forw)
@@ -709,12 +709,12 @@
  apply (rule stream_finite_ind [of x],auto)
   apply (simp add: stream.finite_def)
   apply (drule slen_take_lemma1,blast)
- apply (simp_all add: zero_inat_def iSuc_def split: inat.splits)
+ apply (simp_all add: zero_enat_def iSuc_def split: enat.splits)
 apply (erule_tac x="y" in allE,auto)
 by (rule_tac x="a && w" in exI,auto)
 
 lemma rt_sconc1: "Fin n = #x ==> i_rt n (x ooo y) = y"
-apply (simp add: sconc_def split: inat.splits, arith?,auto)
+apply (simp add: sconc_def split: enat.splits, arith?,auto)
 apply (rule someI2_ex,auto)
 by (drule ex_sconc,simp)
 
@@ -948,7 +948,7 @@
 (* ----------------------------------------------------------------------- *)
 
 lemma constr_sconc_UUs [simp]: "constr_sconc UU s = s"
-by (simp add: constr_sconc_def zero_inat_def)
+by (simp add: constr_sconc_def zero_enat_def)
 
 lemma "x ooo y = constr_sconc x y"
 apply (case_tac "#x")
--- a/src/HOL/IsaMakefile	Tue Jul 19 11:15:38 2011 +0200
+++ b/src/HOL/IsaMakefile	Tue Jul 19 14:35:44 2011 +0200
@@ -444,25 +444,25 @@
   Library/AssocList.thy Library/BigO.thy Library/Binomial.thy		\
   Library/Bit.thy Library/Boolean_Algebra.thy Library/Cardinality.thy	\
   Library/Char_nat.thy Library/Code_Char.thy Library/Code_Char_chr.thy	\
-  Library/Code_Char_ord.thy Library/Code_Integer.thy Library/Code_Natural.thy			\
-  Library/Code_Prolog.thy Tools/Predicate_Compile/code_prolog.ML	\
-  Library/ContNotDenum.thy Library/Continuity.thy Library/Convex.thy	\
-  Library/Countable.thy Library/Diagonalize.thy Library/Dlist.thy	\
-  Library/Dlist_Cset.thy 						\
+  Library/Code_Char_ord.thy Library/Code_Integer.thy			\
+  Library/Code_Natural.thy Library/Code_Prolog.thy			\
+  Tools/Predicate_Compile/code_prolog.ML Library/ContNotDenum.thy	\
+  Library/Continuity.thy Library/Convex.thy Library/Countable.thy	\
+  Library/Diagonalize.thy Library/Dlist.thy Library/Dlist_Cset.thy 	\
   Library/Efficient_Nat.thy Library/Eval_Witness.thy 			\
   Library/Executable_Set.thy Library/Extended_Reals.thy			\
-  Library/Float.thy Library/Formal_Power_Series.thy			\
-  Library/Fraction_Field.thy Library/FrechetDeriv.thy Library/Cset.thy	\
-  Library/FuncSet.thy Library/Function_Algebras.thy			\
-  Library/Fundamental_Theorem_Algebra.thy Library/Glbs.thy		\
-  Library/Indicator_Function.thy Library/Infinite_Set.thy		\
-  Library/Inner_Product.thy Library/Kleene_Algebra.thy			\
-  Library/LaTeXsugar.thy Library/Lattice_Algebras.thy			\
-  Library/Lattice_Syntax.thy Library/Library.thy Library/List_Cset.thy 	\
-  Library/List_Prefix.thy Library/List_lexord.thy Library/Mapping.thy	\
-  Library/Monad_Syntax.thy Library/More_List.thy Library/More_Set.thy	\
-  Library/Multiset.thy Library/Nat_Bijection.thy			\
-  Library/Nat_Infinity.thy Library/Nested_Environment.thy		\
+  Library/Extended_Nat.thy Library/Float.thy				\
+  Library/Formal_Power_Series.thy Library/Fraction_Field.thy		\
+  Library/FrechetDeriv.thy Library/Cset.thy Library/FuncSet.thy		\
+  Library/Function_Algebras.thy Library/Fundamental_Theorem_Algebra.thy	\
+  Library/Glbs.thy Library/Indicator_Function.thy			\
+  Library/Infinite_Set.thy Library/Inner_Product.thy			\
+  Library/Kleene_Algebra.thy Library/LaTeXsugar.thy			\
+  Library/Lattice_Algebras.thy Library/Lattice_Syntax.thy		\
+  Library/Library.thy Library/List_Cset.thy Library/List_Prefix.thy	\
+  Library/List_lexord.thy Library/Mapping.thy Library/Monad_Syntax.thy	\
+  Library/More_List.thy Library/More_Set.thy Library/Multiset.thy	\
+  Library/Nat_Bijection.thy Library/Nested_Environment.thy		\
   Library/Numeral_Type.thy Library/OptionalSugar.thy			\
   Library/Order_Relation.thy Library/Permutation.thy			\
   Library/Permutations.thy Library/Poly_Deriv.thy			\
@@ -481,7 +481,7 @@
   Library/Sum_of_Squares/sos_wrapper.ML					\
   Library/Sum_of_Squares/sum_of_squares.ML				\
   Library/Transitive_Closure_Table.thy Library/Univ_Poly.thy		\
-  Library/While_Combinator.thy Library/Zorn.thy	\
+  Library/While_Combinator.thy Library/Zorn.thy				\
   $(SRC)/Tools/adhoc_overloading.ML Library/positivstellensatz.ML	\
   Library/reflection.ML Library/reify_data.ML				\
   Library/document/root.bib Library/document/root.tex
@@ -1576,7 +1576,7 @@
 HOLCF-Library: HOLCF $(LOG)/HOLCF-Library.gz
 
 $(LOG)/HOLCF-Library.gz: $(OUT)/HOLCF \
-  Library/Nat_Infinity.thy \
+  Library/Extended_Nat.thy \
   HOLCF/Library/Defl_Bifinite.thy \
   HOLCF/Library/Bool_Discrete.thy \
   HOLCF/Library/Char_Discrete.thy \
@@ -1630,7 +1630,7 @@
 HOLCF-FOCUS: HOLCF $(LOG)/HOLCF-FOCUS.gz
 
 $(LOG)/HOLCF-FOCUS.gz: $(OUT)/HOLCF HOLCF/FOCUS/ROOT.ML \
-  Library/Nat_Infinity.thy \
+  Library/Extended_Nat.thy \
   HOLCF/Library/Stream.thy \
   HOLCF/FOCUS/Fstreams.thy \
   HOLCF/FOCUS/Fstream.thy \
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Extended_Nat.thy	Tue Jul 19 14:35:44 2011 +0200
@@ -0,0 +1,551 @@
+(*  Title:      HOL/Library/Extended_Nat.thy
+    Author:     David von Oheimb, TU Muenchen;  Florian Haftmann, TU Muenchen
+    Contributions: David Trachtenherz, TU Muenchen
+*)
+
+header {* Extended natural numbers (i.e. with infinity) *}
+
+theory Extended_Nat
+imports Main
+begin
+
+subsection {* Type definition *}
+
+text {*
+  We extend the standard natural numbers by a special value indicating
+  infinity.
+*}
+
+datatype enat = Fin nat | Infty
+
+notation (xsymbols)
+  Infty  ("\<infinity>")
+
+notation (HTML output)
+  Infty  ("\<infinity>")
+
+
+lemma not_Infty_eq[iff]: "(x ~= Infty) = (EX i. x = Fin i)"
+by (cases x) auto
+
+lemma not_Fin_eq [iff]: "(ALL y. x ~= Fin y) = (x = Infty)"
+by (cases x) auto
+
+
+primrec the_Fin :: "enat \<Rightarrow> nat"
+where "the_Fin (Fin n) = n"
+
+
+subsection {* Constructors and numbers *}
+
+instantiation enat :: "{zero, one, number}"
+begin
+
+definition
+  "0 = Fin 0"
+
+definition
+  [code_unfold]: "1 = Fin 1"
+
+definition
+  [code_unfold, code del]: "number_of k = Fin (number_of k)"
+
+instance ..
+
+end
+
+definition iSuc :: "enat \<Rightarrow> enat" where
+  "iSuc i = (case i of Fin n \<Rightarrow> Fin (Suc n) | \<infinity> \<Rightarrow> \<infinity>)"
+
+lemma Fin_0: "Fin 0 = 0"
+  by (simp add: zero_enat_def)
+
+lemma Fin_1: "Fin 1 = 1"
+  by (simp add: one_enat_def)
+
+lemma Fin_number: "Fin (number_of k) = number_of k"
+  by (simp add: number_of_enat_def)
+
+lemma one_iSuc: "1 = iSuc 0"
+  by (simp add: zero_enat_def one_enat_def iSuc_def)
+
+lemma Infty_ne_i0 [simp]: "\<infinity> \<noteq> 0"
+  by (simp add: zero_enat_def)
+
+lemma i0_ne_Infty [simp]: "0 \<noteq> \<infinity>"
+  by (simp add: zero_enat_def)
+
+lemma zero_enat_eq [simp]:
+  "number_of k = (0\<Colon>enat) \<longleftrightarrow> number_of k = (0\<Colon>nat)"
+  "(0\<Colon>enat) = number_of k \<longleftrightarrow> number_of k = (0\<Colon>nat)"
+  unfolding zero_enat_def number_of_enat_def by simp_all
+
+lemma one_enat_eq [simp]:
+  "number_of k = (1\<Colon>enat) \<longleftrightarrow> number_of k = (1\<Colon>nat)"
+  "(1\<Colon>enat) = number_of k \<longleftrightarrow> number_of k = (1\<Colon>nat)"
+  unfolding one_enat_def number_of_enat_def by simp_all
+
+lemma zero_one_enat_neq [simp]:
+  "\<not> 0 = (1\<Colon>enat)"
+  "\<not> 1 = (0\<Colon>enat)"
+  unfolding zero_enat_def one_enat_def by simp_all
+
+lemma Infty_ne_i1 [simp]: "\<infinity> \<noteq> 1"
+  by (simp add: one_enat_def)
+
+lemma i1_ne_Infty [simp]: "1 \<noteq> \<infinity>"
+  by (simp add: one_enat_def)
+
+lemma Infty_ne_number [simp]: "\<infinity> \<noteq> number_of k"
+  by (simp add: number_of_enat_def)
+
+lemma number_ne_Infty [simp]: "number_of k \<noteq> \<infinity>"
+  by (simp add: number_of_enat_def)
+
+lemma iSuc_Fin: "iSuc (Fin n) = Fin (Suc n)"
+  by (simp add: iSuc_def)
+
+lemma iSuc_number_of: "iSuc (number_of k) = Fin (Suc (number_of k))"
+  by (simp add: iSuc_Fin number_of_enat_def)
+
+lemma iSuc_Infty [simp]: "iSuc \<infinity> = \<infinity>"
+  by (simp add: iSuc_def)
+
+lemma iSuc_ne_0 [simp]: "iSuc n \<noteq> 0"
+  by (simp add: iSuc_def zero_enat_def split: enat.splits)
+
+lemma zero_ne_iSuc [simp]: "0 \<noteq> iSuc n"
+  by (rule iSuc_ne_0 [symmetric])
+
+lemma iSuc_inject [simp]: "iSuc m = iSuc n \<longleftrightarrow> m = n"
+  by (simp add: iSuc_def split: enat.splits)
+
+lemma number_of_enat_inject [simp]:
+  "(number_of k \<Colon> enat) = number_of l \<longleftrightarrow> (number_of k \<Colon> nat) = number_of l"
+  by (simp add: number_of_enat_def)
+
+
+subsection {* Addition *}
+
+instantiation enat :: comm_monoid_add
+begin
+
+definition [nitpick_simp]:
+  "m + n = (case m of \<infinity> \<Rightarrow> \<infinity> | Fin m \<Rightarrow> (case n of \<infinity> \<Rightarrow> \<infinity> | Fin n \<Rightarrow> Fin (m + n)))"
+
+lemma plus_enat_simps [simp, code]:
+  "Fin m + Fin n = Fin (m + n)"
+  "\<infinity> + q = \<infinity>"
+  "q + \<infinity> = \<infinity>"
+  by (simp_all add: plus_enat_def split: enat.splits)
+
+instance proof
+  fix n m q :: enat
+  show "n + m + q = n + (m + q)"
+    by (cases n, auto, cases m, auto, cases q, auto)
+  show "n + m = m + n"
+    by (cases n, auto, cases m, auto)
+  show "0 + n = n"
+    by (cases n) (simp_all add: zero_enat_def)
+qed
+
+end
+
+lemma plus_enat_0 [simp]:
+  "0 + (q\<Colon>enat) = q"
+  "(q\<Colon>enat) + 0 = q"
+  by (simp_all add: plus_enat_def zero_enat_def split: enat.splits)
+
+lemma plus_enat_number [simp]:
+  "(number_of k \<Colon> enat) + number_of l = (if k < Int.Pls then number_of l
+    else if l < Int.Pls then number_of k else number_of (k + l))"
+  unfolding number_of_enat_def plus_enat_simps nat_arith(1) if_distrib [symmetric, of _ Fin] ..
+
+lemma iSuc_number [simp]:
+  "iSuc (number_of k) = (if neg (number_of k \<Colon> int) then 1 else number_of (Int.succ k))"
+  unfolding iSuc_number_of
+  unfolding one_enat_def number_of_enat_def Suc_nat_number_of if_distrib [symmetric] ..
+
+lemma iSuc_plus_1:
+  "iSuc n = n + 1"
+  by (cases n) (simp_all add: iSuc_Fin one_enat_def)
+  
+lemma plus_1_iSuc:
+  "1 + q = iSuc q"
+  "q + 1 = iSuc q"
+by (simp_all add: iSuc_plus_1 add_ac)
+
+lemma iadd_Suc: "iSuc m + n = iSuc (m + n)"
+by (simp_all add: iSuc_plus_1 add_ac)
+
+lemma iadd_Suc_right: "m + iSuc n = iSuc (m + n)"
+by (simp only: add_commute[of m] iadd_Suc)
+
+lemma iadd_is_0: "(m + n = (0::enat)) = (m = 0 \<and> n = 0)"
+by (cases m, cases n, simp_all add: zero_enat_def)
+
+subsection {* Multiplication *}
+
+instantiation enat :: comm_semiring_1
+begin
+
+definition times_enat_def [nitpick_simp]:
+  "m * n = (case m of \<infinity> \<Rightarrow> if n = 0 then 0 else \<infinity> | Fin m \<Rightarrow>
+    (case n of \<infinity> \<Rightarrow> if m = 0 then 0 else \<infinity> | Fin n \<Rightarrow> Fin (m * n)))"
+
+lemma times_enat_simps [simp, code]:
+  "Fin m * Fin n = Fin (m * n)"
+  "\<infinity> * \<infinity> = \<infinity>"
+  "\<infinity> * Fin n = (if n = 0 then 0 else \<infinity>)"
+  "Fin m * \<infinity> = (if m = 0 then 0 else \<infinity>)"
+  unfolding times_enat_def zero_enat_def
+  by (simp_all split: enat.split)
+
+instance proof
+  fix a b c :: enat
+  show "(a * b) * c = a * (b * c)"
+    unfolding times_enat_def zero_enat_def
+    by (simp split: enat.split)
+  show "a * b = b * a"
+    unfolding times_enat_def zero_enat_def
+    by (simp split: enat.split)
+  show "1 * a = a"
+    unfolding times_enat_def zero_enat_def one_enat_def
+    by (simp split: enat.split)
+  show "(a + b) * c = a * c + b * c"
+    unfolding times_enat_def zero_enat_def
+    by (simp split: enat.split add: left_distrib)
+  show "0 * a = 0"
+    unfolding times_enat_def zero_enat_def
+    by (simp split: enat.split)
+  show "a * 0 = 0"
+    unfolding times_enat_def zero_enat_def
+    by (simp split: enat.split)
+  show "(0::enat) \<noteq> 1"
+    unfolding zero_enat_def one_enat_def
+    by simp
+qed
+
+end
+
+lemma mult_iSuc: "iSuc m * n = n + m * n"
+  unfolding iSuc_plus_1 by (simp add: algebra_simps)
+
+lemma mult_iSuc_right: "m * iSuc n = m + m * n"
+  unfolding iSuc_plus_1 by (simp add: algebra_simps)
+
+lemma of_nat_eq_Fin: "of_nat n = Fin n"
+  apply (induct n)
+  apply (simp add: Fin_0)
+  apply (simp add: plus_1_iSuc iSuc_Fin)
+  done
+
+instance enat :: number_semiring
+proof
+  fix n show "number_of (int n) = (of_nat n :: enat)"
+    unfolding number_of_enat_def number_of_int of_nat_id of_nat_eq_Fin ..
+qed
+
+instance enat :: semiring_char_0 proof
+  have "inj Fin" by (rule injI) simp
+  then show "inj (\<lambda>n. of_nat n :: enat)" by (simp add: of_nat_eq_Fin)
+qed
+
+lemma imult_is_0[simp]: "((m::enat) * n = 0) = (m = 0 \<or> n = 0)"
+by(auto simp add: times_enat_def zero_enat_def split: enat.split)
+
+lemma imult_is_Infty: "((a::enat) * b = \<infinity>) = (a = \<infinity> \<and> b \<noteq> 0 \<or> b = \<infinity> \<and> a \<noteq> 0)"
+by(auto simp add: times_enat_def zero_enat_def split: enat.split)
+
+
+subsection {* Subtraction *}
+
+instantiation enat :: minus
+begin
+
+definition diff_enat_def:
+"a - b = (case a of (Fin x) \<Rightarrow> (case b of (Fin y) \<Rightarrow> Fin (x - y) | \<infinity> \<Rightarrow> 0)
+          | \<infinity> \<Rightarrow> \<infinity>)"
+
+instance ..
+
+end
+
+lemma idiff_Fin_Fin[simp,code]: "Fin a - Fin b = Fin (a - b)"
+by(simp add: diff_enat_def)
+
+lemma idiff_Infty[simp,code]: "\<infinity> - n = \<infinity>"
+by(simp add: diff_enat_def)
+
+lemma idiff_Infty_right[simp,code]: "Fin a - \<infinity> = 0"
+by(simp add: diff_enat_def)
+
+lemma idiff_0[simp]: "(0::enat) - n = 0"
+by (cases n, simp_all add: zero_enat_def)
+
+lemmas idiff_Fin_0[simp] = idiff_0[unfolded zero_enat_def]
+
+lemma idiff_0_right[simp]: "(n::enat) - 0 = n"
+by (cases n) (simp_all add: zero_enat_def)
+
+lemmas idiff_Fin_0_right[simp] = idiff_0_right[unfolded zero_enat_def]
+
+lemma idiff_self[simp]: "n \<noteq> \<infinity> \<Longrightarrow> (n::enat) - n = 0"
+by(auto simp: zero_enat_def)
+
+lemma iSuc_minus_iSuc [simp]: "iSuc n - iSuc m = n - m"
+by(simp add: iSuc_def split: enat.split)
+
+lemma iSuc_minus_1 [simp]: "iSuc n - 1 = n"
+by(simp add: one_enat_def iSuc_Fin[symmetric] zero_enat_def[symmetric])
+
+(*lemmas idiff_self_eq_0_Fin = idiff_self_eq_0[unfolded zero_enat_def]*)
+
+
+subsection {* Ordering *}
+
+instantiation enat :: linordered_ab_semigroup_add
+begin
+
+definition [nitpick_simp]:
+  "m \<le> n = (case n of Fin n1 \<Rightarrow> (case m of Fin m1 \<Rightarrow> m1 \<le> n1 | \<infinity> \<Rightarrow> False)
+    | \<infinity> \<Rightarrow> True)"
+
+definition [nitpick_simp]:
+  "m < n = (case m of Fin m1 \<Rightarrow> (case n of Fin n1 \<Rightarrow> m1 < n1 | \<infinity> \<Rightarrow> True)
+    | \<infinity> \<Rightarrow> False)"
+
+lemma enat_ord_simps [simp]:
+  "Fin m \<le> Fin n \<longleftrightarrow> m \<le> n"
+  "Fin m < Fin n \<longleftrightarrow> m < n"
+  "q \<le> \<infinity>"
+  "q < \<infinity> \<longleftrightarrow> q \<noteq> \<infinity>"
+  "\<infinity> \<le> q \<longleftrightarrow> q = \<infinity>"
+  "\<infinity> < q \<longleftrightarrow> False"
+  by (simp_all add: less_eq_enat_def less_enat_def split: enat.splits)
+
+lemma enat_ord_code [code]:
+  "Fin m \<le> Fin n \<longleftrightarrow> m \<le> n"
+  "Fin m < Fin n \<longleftrightarrow> m < n"
+  "q \<le> \<infinity> \<longleftrightarrow> True"
+  "Fin m < \<infinity> \<longleftrightarrow> True"
+  "\<infinity> \<le> Fin n \<longleftrightarrow> False"
+  "\<infinity> < q \<longleftrightarrow> False"
+  by simp_all
+
+instance by default
+  (auto simp add: less_eq_enat_def less_enat_def plus_enat_def split: enat.splits)
+
+end
+
+instance enat :: ordered_comm_semiring
+proof
+  fix a b c :: enat
+  assume "a \<le> b" and "0 \<le> c"
+  thus "c * a \<le> c * b"
+    unfolding times_enat_def less_eq_enat_def zero_enat_def
+    by (simp split: enat.splits)
+qed
+
+lemma enat_ord_number [simp]:
+  "(number_of m \<Colon> enat) \<le> number_of n \<longleftrightarrow> (number_of m \<Colon> nat) \<le> number_of n"
+  "(number_of m \<Colon> enat) < number_of n \<longleftrightarrow> (number_of m \<Colon> nat) < number_of n"
+  by (simp_all add: number_of_enat_def)
+
+lemma i0_lb [simp]: "(0\<Colon>enat) \<le> n"
+  by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
+
+lemma ile0_eq [simp]: "n \<le> (0\<Colon>enat) \<longleftrightarrow> n = 0"
+by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
+
+lemma Infty_ileE [elim!]: "\<infinity> \<le> Fin m \<Longrightarrow> R"
+  by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
+
+lemma Infty_ilessE [elim!]: "\<infinity> < Fin m \<Longrightarrow> R"
+  by simp
+
+lemma not_iless0 [simp]: "\<not> n < (0\<Colon>enat)"
+  by (simp add: zero_enat_def less_enat_def split: enat.splits)
+
+lemma i0_less [simp]: "(0\<Colon>enat) < n \<longleftrightarrow> n \<noteq> 0"
+by (simp add: zero_enat_def less_enat_def split: enat.splits)
+
+lemma iSuc_ile_mono [simp]: "iSuc n \<le> iSuc m \<longleftrightarrow> n \<le> m"
+  by (simp add: iSuc_def less_eq_enat_def split: enat.splits)
+ 
+lemma iSuc_mono [simp]: "iSuc n < iSuc m \<longleftrightarrow> n < m"
+  by (simp add: iSuc_def less_enat_def split: enat.splits)
+
+lemma ile_iSuc [simp]: "n \<le> iSuc n"
+  by (simp add: iSuc_def less_eq_enat_def split: enat.splits)
+
+lemma not_iSuc_ilei0 [simp]: "\<not> iSuc n \<le> 0"
+  by (simp add: zero_enat_def iSuc_def less_eq_enat_def split: enat.splits)
+
+lemma i0_iless_iSuc [simp]: "0 < iSuc n"
+  by (simp add: zero_enat_def iSuc_def less_enat_def split: enat.splits)
+
+lemma iless_iSuc0[simp]: "(n < iSuc 0) = (n = 0)"
+by (simp add: zero_enat_def iSuc_def less_enat_def split: enat.split)
+
+lemma ileI1: "m < n \<Longrightarrow> iSuc m \<le> n"
+  by (simp add: iSuc_def less_eq_enat_def less_enat_def split: enat.splits)
+
+lemma Suc_ile_eq: "Fin (Suc m) \<le> n \<longleftrightarrow> Fin m < n"
+  by (cases n) auto
+
+lemma iless_Suc_eq [simp]: "Fin m < iSuc n \<longleftrightarrow> Fin m \<le> n"
+  by (auto simp add: iSuc_def less_enat_def split: enat.splits)
+
+lemma imult_Infty: "(0::enat) < n \<Longrightarrow> \<infinity> * n = \<infinity>"
+by (simp add: zero_enat_def less_enat_def split: enat.splits)
+
+lemma imult_Infty_right: "(0::enat) < n \<Longrightarrow> n * \<infinity> = \<infinity>"
+by (simp add: zero_enat_def less_enat_def split: enat.splits)
+
+lemma enat_0_less_mult_iff: "(0 < (m::enat) * n) = (0 < m \<and> 0 < n)"
+by (simp only: i0_less imult_is_0, simp)
+
+lemma mono_iSuc: "mono iSuc"
+by(simp add: mono_def)
+
+
+lemma min_enat_simps [simp]:
+  "min (Fin m) (Fin n) = Fin (min m n)"
+  "min q 0 = 0"
+  "min 0 q = 0"
+  "min q \<infinity> = q"
+  "min \<infinity> q = q"
+  by (auto simp add: min_def)
+
+lemma max_enat_simps [simp]:
+  "max (Fin m) (Fin n) = Fin (max m n)"
+  "max q 0 = q"
+  "max 0 q = q"
+  "max q \<infinity> = \<infinity>"
+  "max \<infinity> q = \<infinity>"
+  by (simp_all add: max_def)
+
+lemma Fin_ile: "n \<le> Fin m \<Longrightarrow> \<exists>k. n = Fin k"
+  by (cases n) simp_all
+
+lemma Fin_iless: "n < Fin m \<Longrightarrow> \<exists>k. n = Fin k"
+  by (cases n) simp_all
+
+lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. Fin k < Y j"
+apply (induct_tac k)
+ apply (simp (no_asm) only: Fin_0)
+ apply (fast intro: le_less_trans [OF i0_lb])
+apply (erule exE)
+apply (drule spec)
+apply (erule exE)
+apply (drule ileI1)
+apply (rule iSuc_Fin [THEN subst])
+apply (rule exI)
+apply (erule (1) le_less_trans)
+done
+
+instantiation enat :: "{bot, top}"
+begin
+
+definition bot_enat :: enat where
+  "bot_enat = 0"
+
+definition top_enat :: enat where
+  "top_enat = \<infinity>"
+
+instance proof
+qed (simp_all add: bot_enat_def top_enat_def)
+
+end
+
+lemma finite_Fin_bounded:
+  assumes le_fin: "\<And>y. y \<in> A \<Longrightarrow> y \<le> Fin n"
+  shows "finite A"
+proof (rule finite_subset)
+  show "finite (Fin ` {..n})" by blast
+
+  have "A \<subseteq> {..Fin n}" using le_fin by fastsimp
+  also have "\<dots> \<subseteq> Fin ` {..n}"
+    by (rule subsetI) (case_tac x, auto)
+  finally show "A \<subseteq> Fin ` {..n}" .
+qed
+
+
+subsection {* Well-ordering *}
+
+lemma less_FinE:
+  "[| n < Fin m; !!k. n = Fin k ==> k < m ==> P |] ==> P"
+by (induct n) auto
+
+lemma less_InftyE:
+  "[| n < Infty; !!k. n = Fin k ==> P |] ==> P"
+by (induct n) auto
+
+lemma enat_less_induct:
+  assumes prem: "!!n. \<forall>m::enat. m < n --> P m ==> P n" shows "P n"
+proof -
+  have P_Fin: "!!k. P (Fin k)"
+    apply (rule nat_less_induct)
+    apply (rule prem, clarify)
+    apply (erule less_FinE, simp)
+    done
+  show ?thesis
+  proof (induct n)
+    fix nat
+    show "P (Fin nat)" by (rule P_Fin)
+  next
+    show "P Infty"
+      apply (rule prem, clarify)
+      apply (erule less_InftyE)
+      apply (simp add: P_Fin)
+      done
+  qed
+qed
+
+instance enat :: wellorder
+proof
+  fix P and n
+  assume hyp: "(\<And>n\<Colon>enat. (\<And>m\<Colon>enat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
+  show "P n" by (blast intro: enat_less_induct hyp)
+qed
+
+subsection {* Complete Lattice *}
+
+instantiation enat :: complete_lattice
+begin
+
+definition inf_enat :: "enat \<Rightarrow> enat \<Rightarrow> enat" where
+  "inf_enat \<equiv> min"
+
+definition sup_enat :: "enat \<Rightarrow> enat \<Rightarrow> enat" where
+  "sup_enat \<equiv> max"
+
+definition Inf_enat :: "enat set \<Rightarrow> enat" where
+  "Inf_enat A \<equiv> if A = {} then \<infinity> else (LEAST x. x \<in> A)"
+
+definition Sup_enat :: "enat set \<Rightarrow> enat" where
+  "Sup_enat A \<equiv> if A = {} then 0
+    else if finite A then Max A
+                     else \<infinity>"
+instance proof
+  fix x :: "enat" and A :: "enat set"
+  { assume "x \<in> A" then show "Inf A \<le> x"
+      unfolding Inf_enat_def by (auto intro: Least_le) }
+  { assume "\<And>y. y \<in> A \<Longrightarrow> x \<le> y" then show "x \<le> Inf A"
+      unfolding Inf_enat_def
+      by (cases "A = {}") (auto intro: LeastI2_ex) }
+  { assume "x \<in> A" then show "x \<le> Sup A"
+      unfolding Sup_enat_def by (cases "finite A") auto }
+  { assume "\<And>y. y \<in> A \<Longrightarrow> y \<le> x" then show "Sup A \<le> x"
+      unfolding Sup_enat_def using finite_Fin_bounded by auto }
+qed (simp_all add: inf_enat_def sup_enat_def)
+end
+
+
+subsection {* Traditional theorem names *}
+
+lemmas enat_defs = zero_enat_def one_enat_def number_of_enat_def iSuc_def
+  plus_enat_def less_eq_enat_def less_enat_def
+
+end
--- a/src/HOL/Library/Library.thy	Tue Jul 19 11:15:38 2011 +0200
+++ b/src/HOL/Library/Library.thy	Tue Jul 19 14:35:44 2011 +0200
@@ -15,6 +15,7 @@
   Diagonalize
   Dlist_Cset
   Eval_Witness
+  Extended_Nat
   Float
   Formal_Power_Series
   Fraction_Field
@@ -35,7 +36,6 @@
   Monad_Syntax
   More_List
   Multiset
-  Nat_Infinity
   Nested_Environment
   Numeral_Type
   OptionalSugar
--- a/src/HOL/Library/Nat_Infinity.thy	Tue Jul 19 11:15:38 2011 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,551 +0,0 @@
-(*  Title:      HOL/Library/Nat_Infinity.thy
-    Author:     David von Oheimb, TU Muenchen;  Florian Haftmann, TU Muenchen
-    Contributions: David Trachtenherz, TU Muenchen
-*)
-
-header {* Natural numbers with infinity *}
-
-theory Nat_Infinity
-imports Main
-begin
-
-subsection {* Type definition *}
-
-text {*
-  We extend the standard natural numbers by a special value indicating
-  infinity.
-*}
-
-datatype inat = Fin nat | Infty
-
-notation (xsymbols)
-  Infty  ("\<infinity>")
-
-notation (HTML output)
-  Infty  ("\<infinity>")
-
-
-lemma not_Infty_eq[iff]: "(x ~= Infty) = (EX i. x = Fin i)"
-by (cases x) auto
-
-lemma not_Fin_eq [iff]: "(ALL y. x ~= Fin y) = (x = Infty)"
-by (cases x) auto
-
-
-primrec the_Fin :: "inat \<Rightarrow> nat"
-where "the_Fin (Fin n) = n"
-
-
-subsection {* Constructors and numbers *}
-
-instantiation inat :: "{zero, one, number}"
-begin
-
-definition
-  "0 = Fin 0"
-
-definition
-  [code_unfold]: "1 = Fin 1"
-
-definition
-  [code_unfold, code del]: "number_of k = Fin (number_of k)"
-
-instance ..
-
-end
-
-definition iSuc :: "inat \<Rightarrow> inat" where
-  "iSuc i = (case i of Fin n \<Rightarrow> Fin (Suc n) | \<infinity> \<Rightarrow> \<infinity>)"
-
-lemma Fin_0: "Fin 0 = 0"
-  by (simp add: zero_inat_def)
-
-lemma Fin_1: "Fin 1 = 1"
-  by (simp add: one_inat_def)
-
-lemma Fin_number: "Fin (number_of k) = number_of k"
-  by (simp add: number_of_inat_def)
-
-lemma one_iSuc: "1 = iSuc 0"
-  by (simp add: zero_inat_def one_inat_def iSuc_def)
-
-lemma Infty_ne_i0 [simp]: "\<infinity> \<noteq> 0"
-  by (simp add: zero_inat_def)
-
-lemma i0_ne_Infty [simp]: "0 \<noteq> \<infinity>"
-  by (simp add: zero_inat_def)
-
-lemma zero_inat_eq [simp]:
-  "number_of k = (0\<Colon>inat) \<longleftrightarrow> number_of k = (0\<Colon>nat)"
-  "(0\<Colon>inat) = number_of k \<longleftrightarrow> number_of k = (0\<Colon>nat)"
-  unfolding zero_inat_def number_of_inat_def by simp_all
-
-lemma one_inat_eq [simp]:
-  "number_of k = (1\<Colon>inat) \<longleftrightarrow> number_of k = (1\<Colon>nat)"
-  "(1\<Colon>inat) = number_of k \<longleftrightarrow> number_of k = (1\<Colon>nat)"
-  unfolding one_inat_def number_of_inat_def by simp_all
-
-lemma zero_one_inat_neq [simp]:
-  "\<not> 0 = (1\<Colon>inat)"
-  "\<not> 1 = (0\<Colon>inat)"
-  unfolding zero_inat_def one_inat_def by simp_all
-
-lemma Infty_ne_i1 [simp]: "\<infinity> \<noteq> 1"
-  by (simp add: one_inat_def)
-
-lemma i1_ne_Infty [simp]: "1 \<noteq> \<infinity>"
-  by (simp add: one_inat_def)
-
-lemma Infty_ne_number [simp]: "\<infinity> \<noteq> number_of k"
-  by (simp add: number_of_inat_def)
-
-lemma number_ne_Infty [simp]: "number_of k \<noteq> \<infinity>"
-  by (simp add: number_of_inat_def)
-
-lemma iSuc_Fin: "iSuc (Fin n) = Fin (Suc n)"
-  by (simp add: iSuc_def)
-
-lemma iSuc_number_of: "iSuc (number_of k) = Fin (Suc (number_of k))"
-  by (simp add: iSuc_Fin number_of_inat_def)
-
-lemma iSuc_Infty [simp]: "iSuc \<infinity> = \<infinity>"
-  by (simp add: iSuc_def)
-
-lemma iSuc_ne_0 [simp]: "iSuc n \<noteq> 0"
-  by (simp add: iSuc_def zero_inat_def split: inat.splits)
-
-lemma zero_ne_iSuc [simp]: "0 \<noteq> iSuc n"
-  by (rule iSuc_ne_0 [symmetric])
-
-lemma iSuc_inject [simp]: "iSuc m = iSuc n \<longleftrightarrow> m = n"
-  by (simp add: iSuc_def split: inat.splits)
-
-lemma number_of_inat_inject [simp]:
-  "(number_of k \<Colon> inat) = number_of l \<longleftrightarrow> (number_of k \<Colon> nat) = number_of l"
-  by (simp add: number_of_inat_def)
-
-
-subsection {* Addition *}
-
-instantiation inat :: comm_monoid_add
-begin
-
-definition [nitpick_simp]:
-  "m + n = (case m of \<infinity> \<Rightarrow> \<infinity> | Fin m \<Rightarrow> (case n of \<infinity> \<Rightarrow> \<infinity> | Fin n \<Rightarrow> Fin (m + n)))"
-
-lemma plus_inat_simps [simp, code]:
-  "Fin m + Fin n = Fin (m + n)"
-  "\<infinity> + q = \<infinity>"
-  "q + \<infinity> = \<infinity>"
-  by (simp_all add: plus_inat_def split: inat.splits)
-
-instance proof
-  fix n m q :: inat
-  show "n + m + q = n + (m + q)"
-    by (cases n, auto, cases m, auto, cases q, auto)
-  show "n + m = m + n"
-    by (cases n, auto, cases m, auto)
-  show "0 + n = n"
-    by (cases n) (simp_all add: zero_inat_def)
-qed
-
-end
-
-lemma plus_inat_0 [simp]:
-  "0 + (q\<Colon>inat) = q"
-  "(q\<Colon>inat) + 0 = q"
-  by (simp_all add: plus_inat_def zero_inat_def split: inat.splits)
-
-lemma plus_inat_number [simp]:
-  "(number_of k \<Colon> inat) + number_of l = (if k < Int.Pls then number_of l
-    else if l < Int.Pls then number_of k else number_of (k + l))"
-  unfolding number_of_inat_def plus_inat_simps nat_arith(1) if_distrib [symmetric, of _ Fin] ..
-
-lemma iSuc_number [simp]:
-  "iSuc (number_of k) = (if neg (number_of k \<Colon> int) then 1 else number_of (Int.succ k))"
-  unfolding iSuc_number_of
-  unfolding one_inat_def number_of_inat_def Suc_nat_number_of if_distrib [symmetric] ..
-
-lemma iSuc_plus_1:
-  "iSuc n = n + 1"
-  by (cases n) (simp_all add: iSuc_Fin one_inat_def)
-  
-lemma plus_1_iSuc:
-  "1 + q = iSuc q"
-  "q + 1 = iSuc q"
-by (simp_all add: iSuc_plus_1 add_ac)
-
-lemma iadd_Suc: "iSuc m + n = iSuc (m + n)"
-by (simp_all add: iSuc_plus_1 add_ac)
-
-lemma iadd_Suc_right: "m + iSuc n = iSuc (m + n)"
-by (simp only: add_commute[of m] iadd_Suc)
-
-lemma iadd_is_0: "(m + n = (0::inat)) = (m = 0 \<and> n = 0)"
-by (cases m, cases n, simp_all add: zero_inat_def)
-
-subsection {* Multiplication *}
-
-instantiation inat :: comm_semiring_1
-begin
-
-definition times_inat_def [nitpick_simp]:
-  "m * n = (case m of \<infinity> \<Rightarrow> if n = 0 then 0 else \<infinity> | Fin m \<Rightarrow>
-    (case n of \<infinity> \<Rightarrow> if m = 0 then 0 else \<infinity> | Fin n \<Rightarrow> Fin (m * n)))"
-
-lemma times_inat_simps [simp, code]:
-  "Fin m * Fin n = Fin (m * n)"
-  "\<infinity> * \<infinity> = \<infinity>"
-  "\<infinity> * Fin n = (if n = 0 then 0 else \<infinity>)"
-  "Fin m * \<infinity> = (if m = 0 then 0 else \<infinity>)"
-  unfolding times_inat_def zero_inat_def
-  by (simp_all split: inat.split)
-
-instance proof
-  fix a b c :: inat
-  show "(a * b) * c = a * (b * c)"
-    unfolding times_inat_def zero_inat_def
-    by (simp split: inat.split)
-  show "a * b = b * a"
-    unfolding times_inat_def zero_inat_def
-    by (simp split: inat.split)
-  show "1 * a = a"
-    unfolding times_inat_def zero_inat_def one_inat_def
-    by (simp split: inat.split)
-  show "(a + b) * c = a * c + b * c"
-    unfolding times_inat_def zero_inat_def
-    by (simp split: inat.split add: left_distrib)
-  show "0 * a = 0"
-    unfolding times_inat_def zero_inat_def
-    by (simp split: inat.split)
-  show "a * 0 = 0"
-    unfolding times_inat_def zero_inat_def
-    by (simp split: inat.split)
-  show "(0::inat) \<noteq> 1"
-    unfolding zero_inat_def one_inat_def
-    by simp
-qed
-
-end
-
-lemma mult_iSuc: "iSuc m * n = n + m * n"
-  unfolding iSuc_plus_1 by (simp add: algebra_simps)
-
-lemma mult_iSuc_right: "m * iSuc n = m + m * n"
-  unfolding iSuc_plus_1 by (simp add: algebra_simps)
-
-lemma of_nat_eq_Fin: "of_nat n = Fin n"
-  apply (induct n)
-  apply (simp add: Fin_0)
-  apply (simp add: plus_1_iSuc iSuc_Fin)
-  done
-
-instance inat :: number_semiring
-proof
-  fix n show "number_of (int n) = (of_nat n :: inat)"
-    unfolding number_of_inat_def number_of_int of_nat_id of_nat_eq_Fin ..
-qed
-
-instance inat :: semiring_char_0 proof
-  have "inj Fin" by (rule injI) simp
-  then show "inj (\<lambda>n. of_nat n :: inat)" by (simp add: of_nat_eq_Fin)
-qed
-
-lemma imult_is_0[simp]: "((m::inat) * n = 0) = (m = 0 \<or> n = 0)"
-by(auto simp add: times_inat_def zero_inat_def split: inat.split)
-
-lemma imult_is_Infty: "((a::inat) * b = \<infinity>) = (a = \<infinity> \<and> b \<noteq> 0 \<or> b = \<infinity> \<and> a \<noteq> 0)"
-by(auto simp add: times_inat_def zero_inat_def split: inat.split)
-
-
-subsection {* Subtraction *}
-
-instantiation inat :: minus
-begin
-
-definition diff_inat_def:
-"a - b = (case a of (Fin x) \<Rightarrow> (case b of (Fin y) \<Rightarrow> Fin (x - y) | \<infinity> \<Rightarrow> 0)
-          | \<infinity> \<Rightarrow> \<infinity>)"
-
-instance ..
-
-end
-
-lemma idiff_Fin_Fin[simp,code]: "Fin a - Fin b = Fin (a - b)"
-by(simp add: diff_inat_def)
-
-lemma idiff_Infty[simp,code]: "\<infinity> - n = \<infinity>"
-by(simp add: diff_inat_def)
-
-lemma idiff_Infty_right[simp,code]: "Fin a - \<infinity> = 0"
-by(simp add: diff_inat_def)
-
-lemma idiff_0[simp]: "(0::inat) - n = 0"
-by (cases n, simp_all add: zero_inat_def)
-
-lemmas idiff_Fin_0[simp] = idiff_0[unfolded zero_inat_def]
-
-lemma idiff_0_right[simp]: "(n::inat) - 0 = n"
-by (cases n) (simp_all add: zero_inat_def)
-
-lemmas idiff_Fin_0_right[simp] = idiff_0_right[unfolded zero_inat_def]
-
-lemma idiff_self[simp]: "n \<noteq> \<infinity> \<Longrightarrow> (n::inat) - n = 0"
-by(auto simp: zero_inat_def)
-
-lemma iSuc_minus_iSuc [simp]: "iSuc n - iSuc m = n - m"
-by(simp add: iSuc_def split: inat.split)
-
-lemma iSuc_minus_1 [simp]: "iSuc n - 1 = n"
-by(simp add: one_inat_def iSuc_Fin[symmetric] zero_inat_def[symmetric])
-
-(*lemmas idiff_self_eq_0_Fin = idiff_self_eq_0[unfolded zero_inat_def]*)
-
-
-subsection {* Ordering *}
-
-instantiation inat :: linordered_ab_semigroup_add
-begin
-
-definition [nitpick_simp]:
-  "m \<le> n = (case n of Fin n1 \<Rightarrow> (case m of Fin m1 \<Rightarrow> m1 \<le> n1 | \<infinity> \<Rightarrow> False)
-    | \<infinity> \<Rightarrow> True)"
-
-definition [nitpick_simp]:
-  "m < n = (case m of Fin m1 \<Rightarrow> (case n of Fin n1 \<Rightarrow> m1 < n1 | \<infinity> \<Rightarrow> True)
-    | \<infinity> \<Rightarrow> False)"
-
-lemma inat_ord_simps [simp]:
-  "Fin m \<le> Fin n \<longleftrightarrow> m \<le> n"
-  "Fin m < Fin n \<longleftrightarrow> m < n"
-  "q \<le> \<infinity>"
-  "q < \<infinity> \<longleftrightarrow> q \<noteq> \<infinity>"
-  "\<infinity> \<le> q \<longleftrightarrow> q = \<infinity>"
-  "\<infinity> < q \<longleftrightarrow> False"
-  by (simp_all add: less_eq_inat_def less_inat_def split: inat.splits)
-
-lemma inat_ord_code [code]:
-  "Fin m \<le> Fin n \<longleftrightarrow> m \<le> n"
-  "Fin m < Fin n \<longleftrightarrow> m < n"
-  "q \<le> \<infinity> \<longleftrightarrow> True"
-  "Fin m < \<infinity> \<longleftrightarrow> True"
-  "\<infinity> \<le> Fin n \<longleftrightarrow> False"
-  "\<infinity> < q \<longleftrightarrow> False"
-  by simp_all
-
-instance by default
-  (auto simp add: less_eq_inat_def less_inat_def plus_inat_def split: inat.splits)
-
-end
-
-instance inat :: ordered_comm_semiring
-proof
-  fix a b c :: inat
-  assume "a \<le> b" and "0 \<le> c"
-  thus "c * a \<le> c * b"
-    unfolding times_inat_def less_eq_inat_def zero_inat_def
-    by (simp split: inat.splits)
-qed
-
-lemma inat_ord_number [simp]:
-  "(number_of m \<Colon> inat) \<le> number_of n \<longleftrightarrow> (number_of m \<Colon> nat) \<le> number_of n"
-  "(number_of m \<Colon> inat) < number_of n \<longleftrightarrow> (number_of m \<Colon> nat) < number_of n"
-  by (simp_all add: number_of_inat_def)
-
-lemma i0_lb [simp]: "(0\<Colon>inat) \<le> n"
-  by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
-
-lemma ile0_eq [simp]: "n \<le> (0\<Colon>inat) \<longleftrightarrow> n = 0"
-by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
-
-lemma Infty_ileE [elim!]: "\<infinity> \<le> Fin m \<Longrightarrow> R"
-  by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
-
-lemma Infty_ilessE [elim!]: "\<infinity> < Fin m \<Longrightarrow> R"
-  by simp
-
-lemma not_iless0 [simp]: "\<not> n < (0\<Colon>inat)"
-  by (simp add: zero_inat_def less_inat_def split: inat.splits)
-
-lemma i0_less [simp]: "(0\<Colon>inat) < n \<longleftrightarrow> n \<noteq> 0"
-by (simp add: zero_inat_def less_inat_def split: inat.splits)
-
-lemma iSuc_ile_mono [simp]: "iSuc n \<le> iSuc m \<longleftrightarrow> n \<le> m"
-  by (simp add: iSuc_def less_eq_inat_def split: inat.splits)
- 
-lemma iSuc_mono [simp]: "iSuc n < iSuc m \<longleftrightarrow> n < m"
-  by (simp add: iSuc_def less_inat_def split: inat.splits)
-
-lemma ile_iSuc [simp]: "n \<le> iSuc n"
-  by (simp add: iSuc_def less_eq_inat_def split: inat.splits)
-
-lemma not_iSuc_ilei0 [simp]: "\<not> iSuc n \<le> 0"
-  by (simp add: zero_inat_def iSuc_def less_eq_inat_def split: inat.splits)
-
-lemma i0_iless_iSuc [simp]: "0 < iSuc n"
-  by (simp add: zero_inat_def iSuc_def less_inat_def split: inat.splits)
-
-lemma iless_iSuc0[simp]: "(n < iSuc 0) = (n = 0)"
-by (simp add: zero_inat_def iSuc_def less_inat_def split: inat.split)
-
-lemma ileI1: "m < n \<Longrightarrow> iSuc m \<le> n"
-  by (simp add: iSuc_def less_eq_inat_def less_inat_def split: inat.splits)
-
-lemma Suc_ile_eq: "Fin (Suc m) \<le> n \<longleftrightarrow> Fin m < n"
-  by (cases n) auto
-
-lemma iless_Suc_eq [simp]: "Fin m < iSuc n \<longleftrightarrow> Fin m \<le> n"
-  by (auto simp add: iSuc_def less_inat_def split: inat.splits)
-
-lemma imult_Infty: "(0::inat) < n \<Longrightarrow> \<infinity> * n = \<infinity>"
-by (simp add: zero_inat_def less_inat_def split: inat.splits)
-
-lemma imult_Infty_right: "(0::inat) < n \<Longrightarrow> n * \<infinity> = \<infinity>"
-by (simp add: zero_inat_def less_inat_def split: inat.splits)
-
-lemma inat_0_less_mult_iff: "(0 < (m::inat) * n) = (0 < m \<and> 0 < n)"
-by (simp only: i0_less imult_is_0, simp)
-
-lemma mono_iSuc: "mono iSuc"
-by(simp add: mono_def)
-
-
-lemma min_inat_simps [simp]:
-  "min (Fin m) (Fin n) = Fin (min m n)"
-  "min q 0 = 0"
-  "min 0 q = 0"
-  "min q \<infinity> = q"
-  "min \<infinity> q = q"
-  by (auto simp add: min_def)
-
-lemma max_inat_simps [simp]:
-  "max (Fin m) (Fin n) = Fin (max m n)"
-  "max q 0 = q"
-  "max 0 q = q"
-  "max q \<infinity> = \<infinity>"
-  "max \<infinity> q = \<infinity>"
-  by (simp_all add: max_def)
-
-lemma Fin_ile: "n \<le> Fin m \<Longrightarrow> \<exists>k. n = Fin k"
-  by (cases n) simp_all
-
-lemma Fin_iless: "n < Fin m \<Longrightarrow> \<exists>k. n = Fin k"
-  by (cases n) simp_all
-
-lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. Fin k < Y j"
-apply (induct_tac k)
- apply (simp (no_asm) only: Fin_0)
- apply (fast intro: le_less_trans [OF i0_lb])
-apply (erule exE)
-apply (drule spec)
-apply (erule exE)
-apply (drule ileI1)
-apply (rule iSuc_Fin [THEN subst])
-apply (rule exI)
-apply (erule (1) le_less_trans)
-done
-
-instantiation inat :: "{bot, top}"
-begin
-
-definition bot_inat :: inat where
-  "bot_inat = 0"
-
-definition top_inat :: inat where
-  "top_inat = \<infinity>"
-
-instance proof
-qed (simp_all add: bot_inat_def top_inat_def)
-
-end
-
-lemma finite_Fin_bounded:
-  assumes le_fin: "\<And>y. y \<in> A \<Longrightarrow> y \<le> Fin n"
-  shows "finite A"
-proof (rule finite_subset)
-  show "finite (Fin ` {..n})" by blast
-
-  have "A \<subseteq> {..Fin n}" using le_fin by fastsimp
-  also have "\<dots> \<subseteq> Fin ` {..n}"
-    by (rule subsetI) (case_tac x, auto)
-  finally show "A \<subseteq> Fin ` {..n}" .
-qed
-
-
-subsection {* Well-ordering *}
-
-lemma less_FinE:
-  "[| n < Fin m; !!k. n = Fin k ==> k < m ==> P |] ==> P"
-by (induct n) auto
-
-lemma less_InftyE:
-  "[| n < Infty; !!k. n = Fin k ==> P |] ==> P"
-by (induct n) auto
-
-lemma inat_less_induct:
-  assumes prem: "!!n. \<forall>m::inat. m < n --> P m ==> P n" shows "P n"
-proof -
-  have P_Fin: "!!k. P (Fin k)"
-    apply (rule nat_less_induct)
-    apply (rule prem, clarify)
-    apply (erule less_FinE, simp)
-    done
-  show ?thesis
-  proof (induct n)
-    fix nat
-    show "P (Fin nat)" by (rule P_Fin)
-  next
-    show "P Infty"
-      apply (rule prem, clarify)
-      apply (erule less_InftyE)
-      apply (simp add: P_Fin)
-      done
-  qed
-qed
-
-instance inat :: wellorder
-proof
-  fix P and n
-  assume hyp: "(\<And>n\<Colon>inat. (\<And>m\<Colon>inat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
-  show "P n" by (blast intro: inat_less_induct hyp)
-qed
-
-subsection {* Complete Lattice *}
-
-instantiation inat :: complete_lattice
-begin
-
-definition inf_inat :: "inat \<Rightarrow> inat \<Rightarrow> inat" where
-  "inf_inat \<equiv> min"
-
-definition sup_inat :: "inat \<Rightarrow> inat \<Rightarrow> inat" where
-  "sup_inat \<equiv> max"
-
-definition Inf_inat :: "inat set \<Rightarrow> inat" where
-  "Inf_inat A \<equiv> if A = {} then \<infinity> else (LEAST x. x \<in> A)"
-
-definition Sup_inat :: "inat set \<Rightarrow> inat" where
-  "Sup_inat A \<equiv> if A = {} then 0
-    else if finite A then Max A
-                     else \<infinity>"
-instance proof
-  fix x :: "inat" and A :: "inat set"
-  { assume "x \<in> A" then show "Inf A \<le> x"
-      unfolding Inf_inat_def by (auto intro: Least_le) }
-  { assume "\<And>y. y \<in> A \<Longrightarrow> x \<le> y" then show "x \<le> Inf A"
-      unfolding Inf_inat_def
-      by (cases "A = {}") (auto intro: LeastI2_ex) }
-  { assume "x \<in> A" then show "x \<le> Sup A"
-      unfolding Sup_inat_def by (cases "finite A") auto }
-  { assume "\<And>y. y \<in> A \<Longrightarrow> y \<le> x" then show "Sup A \<le> x"
-      unfolding Sup_inat_def using finite_Fin_bounded by auto }
-qed (simp_all add: inf_inat_def sup_inat_def)
-end
-
-
-subsection {* Traditional theorem names *}
-
-lemmas inat_defs = zero_inat_def one_inat_def number_of_inat_def iSuc_def
-  plus_inat_def less_eq_inat_def less_inat_def
-
-end