wenzelm@11355: (* Title: HOL/Library/Nat_Infinity.thy wenzelm@11355: ID: $Id$ wenzelm@11355: Author: David von Oheimb, TU Muenchen oheimb@11351: License: GPL (GNU GENERAL PUBLIC LICENSE) oheimb@11351: *) oheimb@11351: oheimb@11351: header {* oheimb@11351: \title{Natural numbers with infinity} oheimb@11351: \author{David von Oheimb} oheimb@11351: *} oheimb@11351: wenzelm@11355: theory Nat_Infinity = Main: oheimb@11351: oheimb@11351: subsection "Definitions" oheimb@11351: oheimb@11351: text {* wenzelm@11355: We extend the standard natural numbers by a special value indicating wenzelm@11355: infinity. This includes extending the ordering relations @{term "op wenzelm@11355: <"} and @{term "op \"}. oheimb@11351: *} oheimb@11351: oheimb@11351: datatype inat = Fin nat | Infty oheimb@11351: oheimb@11351: instance inat :: ord .. oheimb@11351: instance inat :: zero .. oheimb@11351: oheimb@11351: consts wenzelm@11355: iSuc :: "inat => inat" oheimb@11351: oheimb@11351: syntax (xsymbols) wenzelm@11355: Infty :: inat ("\") oheimb@11351: kleing@14565: syntax (HTML output) kleing@14565: Infty :: inat ("\") kleing@14565: oheimb@11351: defs wenzelm@11701: Zero_inat_def: "0 == Fin 0" wenzelm@11355: iSuc_def: "iSuc i == case i of Fin n => Fin (Suc n) | \ => \" wenzelm@11355: iless_def: "m < n == wenzelm@11355: case m of Fin m1 => (case n of Fin n1 => m1 < n1 | \ => True) wenzelm@11355: | \ => False" wenzelm@11355: ile_def: "(m::inat) \ n == \ (n < m)" oheimb@11351: wenzelm@11701: lemmas inat_defs = Zero_inat_def iSuc_def iless_def ile_def oheimb@11351: lemmas inat_splits = inat.split inat.split_asm oheimb@11351: wenzelm@11355: text {* wenzelm@11357: Below is a not quite complete set of theorems. Use the method wenzelm@11357: @{text "(simp add: inat_defs split:inat_splits, arith?)"} to prove wenzelm@11357: new theorems or solve arithmetic subgoals involving @{typ inat} on wenzelm@11357: the fly. oheimb@11351: *} oheimb@11351: oheimb@11351: subsection "Constructors" oheimb@11351: oheimb@11351: lemma Fin_0: "Fin 0 = 0" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma Infty_ne_i0 [simp]: "\ \ 0" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma i0_ne_Infty [simp]: "0 \ \" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma iSuc_Fin [simp]: "iSuc (Fin n) = Fin (Suc n)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma iSuc_Infty [simp]: "iSuc \ = \" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma iSuc_ne_0 [simp]: "iSuc n \ 0" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma iSuc_inject [simp]: "(iSuc x = iSuc y) = (x = y)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: oheimb@11351: subsection "Ordering relations" oheimb@11351: oheimb@11351: lemma Infty_ilessE [elim!]: "\ < Fin m ==> R" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma iless_linear: "m < n \ m = n \ n < (m::inat)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma iless_not_refl [simp]: "\ n < (n::inat)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma iless_trans: "i < j ==> j < k ==> i < (k::inat)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma iless_not_sym: "n < m ==> \ m < (n::inat)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma Fin_iless_mono [simp]: "(Fin n < Fin m) = (n < m)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma Fin_iless_Infty [simp]: "Fin n < \" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11655: lemma Infty_eq [simp]: "(n < \) = (n \ \)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma i0_eq [simp]: "((0::inat) < n) = (n \ 0)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma i0_iless_iSuc [simp]: "0 < iSuc n" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma not_ilessi0 [simp]: "\ n < (0::inat)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma Fin_iless: "n < Fin m ==> \k. n = Fin k" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11655: lemma iSuc_mono [simp]: "(iSuc n < iSuc m) = (n < m)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: oheimb@11351: (* ----------------------------------------------------------------------- *) oheimb@11351: wenzelm@11655: lemma ile_def2: "(m \ n) = (m < n \ m = (n::inat))" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma ile_refl [simp]: "n \ (n::inat)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma ile_trans: "i \ j ==> j \ k ==> i \ (k::inat)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma ile_iless_trans: "i \ j ==> j < k ==> i < (k::inat)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma iless_ile_trans: "i < j ==> j \ k ==> i < (k::inat)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma Infty_ub [simp]: "n \ \" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma i0_lb [simp]: "(0::inat) \ n" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma Infty_ileE [elim!]: "\ \ Fin m ==> R" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma Fin_ile_mono [simp]: "(Fin n \ Fin m) = (n \ m)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma ilessI1: "n \ m ==> n \ m ==> n < (m::inat)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma ileI1: "m < n ==> iSuc m \ n" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11655: lemma Suc_ile_eq: "(Fin (Suc m) \ n) = (Fin m < n)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11655: lemma iSuc_ile_mono [simp]: "(iSuc n \ iSuc m) = (n \ m)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11655: lemma iless_Suc_eq [simp]: "(Fin m < iSuc n) = (Fin m \ n)" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma not_iSuc_ilei0 [simp]: "\ iSuc n \ 0" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma ile_iSuc [simp]: "n \ iSuc n" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: wenzelm@11355: lemma Fin_ile: "n \ Fin m ==> \k. n = Fin k" wenzelm@11357: by (simp add: inat_defs split:inat_splits, arith?) oheimb@11351: oheimb@11351: lemma chain_incr: "\i. \j. Y i < Y j ==> \j. Fin k < Y j" wenzelm@11355: apply (induct_tac k) wenzelm@11355: apply (simp (no_asm) only: Fin_0) wenzelm@11355: apply (fast intro: ile_iless_trans i0_lb) wenzelm@11355: apply (erule exE) wenzelm@11355: apply (drule spec) wenzelm@11355: apply (erule exE) wenzelm@11355: apply (drule ileI1) wenzelm@11355: apply (rule iSuc_Fin [THEN subst]) wenzelm@11355: apply (rule exI) wenzelm@11355: apply (erule (1) ile_iless_trans) wenzelm@11355: done oheimb@11351: oheimb@11351: end