hoelzl@43919: (* Title: HOL/Library/Extended_Nat.thy haftmann@27110: Author: David von Oheimb, TU Muenchen; Florian Haftmann, TU Muenchen nipkow@41853: Contributions: David Trachtenherz, TU Muenchen oheimb@11351: *) oheimb@11351: hoelzl@43919: header {* Extended natural numbers (i.e. with infinity) *} oheimb@11351: hoelzl@43919: theory Extended_Nat haftmann@30663: imports Main nipkow@15131: begin oheimb@11351: hoelzl@43921: class infinity = hoelzl@43921: fixes infinity :: "'a" hoelzl@43921: hoelzl@43921: notation (xsymbols) hoelzl@43921: infinity ("\") hoelzl@43921: hoelzl@43921: notation (HTML output) hoelzl@43921: infinity ("\") hoelzl@43921: haftmann@27110: subsection {* Type definition *} oheimb@11351: oheimb@11351: text {* wenzelm@11355: We extend the standard natural numbers by a special value indicating haftmann@27110: infinity. oheimb@11351: *} oheimb@11351: hoelzl@43921: typedef (open) enat = "UNIV :: nat option set" .. hoelzl@43921: hoelzl@43924: definition enat :: "nat \ enat" where hoelzl@43924: "enat n = Abs_enat (Some n)" hoelzl@43921: hoelzl@43921: instantiation enat :: infinity hoelzl@43921: begin hoelzl@43921: definition "\ = Abs_enat None" hoelzl@43921: instance proof qed hoelzl@43921: end hoelzl@43921: hoelzl@43924: rep_datatype enat "\ :: enat" hoelzl@43921: proof - hoelzl@43924: fix P i assume "\j. P (enat j)" "P \" hoelzl@43921: then show "P i" hoelzl@43921: proof induct hoelzl@43921: case (Abs_enat y) then show ?case hoelzl@43921: by (cases y rule: option.exhaust) hoelzl@43924: (auto simp: enat_def infinity_enat_def) hoelzl@43921: qed hoelzl@43924: qed (auto simp add: enat_def infinity_enat_def Abs_enat_inject) wenzelm@19736: hoelzl@43924: declare [[coercion "enat::nat\enat"]] wenzelm@19736: noschinl@45934: lemmas enat2_cases = enat.exhaust[case_product enat.exhaust] noschinl@45934: lemmas enat3_cases = enat.exhaust[case_product enat.exhaust enat.exhaust] noschinl@45934: huffman@44019: lemma not_infinity_eq [iff]: "(x \ \) = (EX i. x = enat i)" huffman@44019: by (cases x) auto nipkow@31084: hoelzl@43924: lemma not_enat_eq [iff]: "(ALL y. x ~= enat y) = (x = \)" huffman@44019: by (cases x) auto nipkow@31077: hoelzl@43924: primrec the_enat :: "enat \ nat" huffman@44019: where "the_enat (enat n) = n" nipkow@41855: haftmann@27110: subsection {* Constructors and numbers *} haftmann@27110: hoelzl@43919: instantiation enat :: "{zero, one, number}" haftmann@25594: begin haftmann@25594: haftmann@25594: definition hoelzl@43924: "0 = enat 0" haftmann@25594: haftmann@25594: definition hoelzl@43924: [code_unfold]: "1 = enat 1" haftmann@25594: haftmann@25594: definition hoelzl@43924: [code_unfold, code del]: "number_of k = enat (number_of k)" oheimb@11351: haftmann@25594: instance .. haftmann@25594: haftmann@25594: end haftmann@25594: huffman@44019: definition eSuc :: "enat \ enat" where huffman@44019: "eSuc i = (case i of enat n \ enat (Suc n) | \ \ \)" oheimb@11351: hoelzl@43924: lemma enat_0: "enat 0 = 0" hoelzl@43919: by (simp add: zero_enat_def) haftmann@27110: hoelzl@43924: lemma enat_1: "enat 1 = 1" hoelzl@43919: by (simp add: one_enat_def) haftmann@27110: hoelzl@43924: lemma enat_number: "enat (number_of k) = number_of k" hoelzl@43919: by (simp add: number_of_enat_def) haftmann@27110: huffman@44019: lemma one_eSuc: "1 = eSuc 0" huffman@44019: by (simp add: zero_enat_def one_enat_def eSuc_def) oheimb@11351: huffman@44019: lemma infinity_ne_i0 [simp]: "(\::enat) \ 0" hoelzl@43919: by (simp add: zero_enat_def) oheimb@11351: huffman@44019: lemma i0_ne_infinity [simp]: "0 \ (\::enat)" hoelzl@43919: by (simp add: zero_enat_def) haftmann@27110: hoelzl@43919: lemma zero_enat_eq [simp]: hoelzl@43919: "number_of k = (0\enat) \ number_of k = (0\nat)" hoelzl@43919: "(0\enat) = number_of k \ number_of k = (0\nat)" hoelzl@43919: unfolding zero_enat_def number_of_enat_def by simp_all haftmann@27110: hoelzl@43919: lemma one_enat_eq [simp]: hoelzl@43919: "number_of k = (1\enat) \ number_of k = (1\nat)" hoelzl@43919: "(1\enat) = number_of k \ number_of k = (1\nat)" hoelzl@43919: unfolding one_enat_def number_of_enat_def by simp_all haftmann@27110: hoelzl@43919: lemma zero_one_enat_neq [simp]: hoelzl@43919: "\ 0 = (1\enat)" hoelzl@43919: "\ 1 = (0\enat)" hoelzl@43919: unfolding zero_enat_def one_enat_def by simp_all oheimb@11351: huffman@44019: lemma infinity_ne_i1 [simp]: "(\::enat) \ 1" hoelzl@43919: by (simp add: one_enat_def) haftmann@27110: huffman@44019: lemma i1_ne_infinity [simp]: "1 \ (\::enat)" hoelzl@43919: by (simp add: one_enat_def) haftmann@27110: huffman@44019: lemma infinity_ne_number [simp]: "(\::enat) \ number_of k" hoelzl@43919: by (simp add: number_of_enat_def) haftmann@27110: huffman@44019: lemma number_ne_infinity [simp]: "number_of k \ (\::enat)" hoelzl@43919: by (simp add: number_of_enat_def) haftmann@27110: huffman@44019: lemma eSuc_enat: "eSuc (enat n) = enat (Suc n)" huffman@44019: by (simp add: eSuc_def) haftmann@27110: huffman@44019: lemma eSuc_number_of: "eSuc (number_of k) = enat (Suc (number_of k))" huffman@44019: by (simp add: eSuc_enat number_of_enat_def) oheimb@11351: huffman@44019: lemma eSuc_infinity [simp]: "eSuc \ = \" huffman@44019: by (simp add: eSuc_def) oheimb@11351: huffman@44019: lemma eSuc_ne_0 [simp]: "eSuc n \ 0" huffman@44019: by (simp add: eSuc_def zero_enat_def split: enat.splits) haftmann@27110: huffman@44019: lemma zero_ne_eSuc [simp]: "0 \ eSuc n" huffman@44019: by (rule eSuc_ne_0 [symmetric]) oheimb@11351: huffman@44019: lemma eSuc_inject [simp]: "eSuc m = eSuc n \ m = n" huffman@44019: by (simp add: eSuc_def split: enat.splits) haftmann@27110: hoelzl@43919: lemma number_of_enat_inject [simp]: hoelzl@43919: "(number_of k \ enat) = number_of l \ (number_of k \ nat) = number_of l" hoelzl@43919: by (simp add: number_of_enat_def) oheimb@11351: oheimb@11351: haftmann@27110: subsection {* Addition *} haftmann@27110: hoelzl@43919: instantiation enat :: comm_monoid_add haftmann@27110: begin haftmann@27110: blanchet@38167: definition [nitpick_simp]: hoelzl@43924: "m + n = (case m of \ \ \ | enat m \ (case n of \ \ \ | enat n \ enat (m + n)))" oheimb@11351: hoelzl@43919: lemma plus_enat_simps [simp, code]: hoelzl@43921: fixes q :: enat hoelzl@43924: shows "enat m + enat n = enat (m + n)" hoelzl@43921: and "\ + q = \" hoelzl@43921: and "q + \ = \" hoelzl@43919: by (simp_all add: plus_enat_def split: enat.splits) haftmann@27110: haftmann@27110: instance proof hoelzl@43919: fix n m q :: enat haftmann@27110: show "n + m + q = n + (m + q)" noschinl@45934: by (cases n m q rule: enat3_cases) auto haftmann@27110: show "n + m = m + n" noschinl@45934: by (cases n m rule: enat2_cases) auto haftmann@27110: show "0 + n = n" hoelzl@43919: by (cases n) (simp_all add: zero_enat_def) huffman@26089: qed huffman@26089: haftmann@27110: end oheimb@11351: hoelzl@43919: lemma plus_enat_number [simp]: hoelzl@43919: "(number_of k \ enat) + number_of l = (if k < Int.Pls then number_of l huffman@29012: else if l < Int.Pls then number_of k else number_of (k + l))" hoelzl@43924: unfolding number_of_enat_def plus_enat_simps nat_arith(1) if_distrib [symmetric, of _ enat] .. oheimb@11351: huffman@44019: lemma eSuc_number [simp]: huffman@44019: "eSuc (number_of k) = (if neg (number_of k \ int) then 1 else number_of (Int.succ k))" huffman@44019: unfolding eSuc_number_of hoelzl@43919: unfolding one_enat_def number_of_enat_def Suc_nat_number_of if_distrib [symmetric] .. oheimb@11351: huffman@44019: lemma eSuc_plus_1: huffman@44019: "eSuc n = n + 1" huffman@44019: by (cases n) (simp_all add: eSuc_enat one_enat_def) haftmann@27110: huffman@44019: lemma plus_1_eSuc: huffman@44019: "1 + q = eSuc q" huffman@44019: "q + 1 = eSuc q" huffman@44019: by (simp_all add: eSuc_plus_1 add_ac) nipkow@41853: huffman@44019: lemma iadd_Suc: "eSuc m + n = eSuc (m + n)" huffman@44019: by (simp_all add: eSuc_plus_1 add_ac) oheimb@11351: huffman@44019: lemma iadd_Suc_right: "m + eSuc n = eSuc (m + n)" huffman@44019: by (simp only: add_commute[of m] iadd_Suc) nipkow@41853: hoelzl@43919: lemma iadd_is_0: "(m + n = (0::enat)) = (m = 0 \ n = 0)" huffman@44019: by (cases m, cases n, simp_all add: zero_enat_def) oheimb@11351: huffman@29014: subsection {* Multiplication *} huffman@29014: hoelzl@43919: instantiation enat :: comm_semiring_1 huffman@29014: begin huffman@29014: hoelzl@43919: definition times_enat_def [nitpick_simp]: hoelzl@43924: "m * n = (case m of \ \ if n = 0 then 0 else \ | enat m \ hoelzl@43924: (case n of \ \ if m = 0 then 0 else \ | enat n \ enat (m * n)))" huffman@29014: hoelzl@43919: lemma times_enat_simps [simp, code]: hoelzl@43924: "enat m * enat n = enat (m * n)" hoelzl@43921: "\ * \ = (\::enat)" hoelzl@43924: "\ * enat n = (if n = 0 then 0 else \)" hoelzl@43924: "enat m * \ = (if m = 0 then 0 else \)" hoelzl@43919: unfolding times_enat_def zero_enat_def hoelzl@43919: by (simp_all split: enat.split) huffman@29014: huffman@29014: instance proof hoelzl@43919: fix a b c :: enat huffman@29014: show "(a * b) * c = a * (b * c)" hoelzl@43919: unfolding times_enat_def zero_enat_def hoelzl@43919: by (simp split: enat.split) huffman@29014: show "a * b = b * a" hoelzl@43919: unfolding times_enat_def zero_enat_def hoelzl@43919: by (simp split: enat.split) huffman@29014: show "1 * a = a" hoelzl@43919: unfolding times_enat_def zero_enat_def one_enat_def hoelzl@43919: by (simp split: enat.split) huffman@29014: show "(a + b) * c = a * c + b * c" hoelzl@43919: unfolding times_enat_def zero_enat_def hoelzl@43919: by (simp split: enat.split add: left_distrib) huffman@29014: show "0 * a = 0" hoelzl@43919: unfolding times_enat_def zero_enat_def hoelzl@43919: by (simp split: enat.split) huffman@29014: show "a * 0 = 0" hoelzl@43919: unfolding times_enat_def zero_enat_def hoelzl@43919: by (simp split: enat.split) hoelzl@43919: show "(0::enat) \ 1" hoelzl@43919: unfolding zero_enat_def one_enat_def huffman@29014: by simp huffman@29014: qed huffman@29014: huffman@29014: end huffman@29014: huffman@44019: lemma mult_eSuc: "eSuc m * n = n + m * n" huffman@44019: unfolding eSuc_plus_1 by (simp add: algebra_simps) huffman@29014: huffman@44019: lemma mult_eSuc_right: "m * eSuc n = m + m * n" huffman@44019: unfolding eSuc_plus_1 by (simp add: algebra_simps) huffman@29014: hoelzl@43924: lemma of_nat_eq_enat: "of_nat n = enat n" huffman@29023: apply (induct n) hoelzl@43924: apply (simp add: enat_0) huffman@44019: apply (simp add: plus_1_eSuc eSuc_enat) huffman@29023: done huffman@29023: hoelzl@43919: instance enat :: number_semiring huffman@43532: proof hoelzl@43919: fix n show "number_of (int n) = (of_nat n :: enat)" hoelzl@43924: unfolding number_of_enat_def number_of_int of_nat_id of_nat_eq_enat .. huffman@43532: qed huffman@43532: hoelzl@43919: instance enat :: semiring_char_0 proof hoelzl@43924: have "inj enat" by (rule injI) simp hoelzl@43924: then show "inj (\n. of_nat n :: enat)" by (simp add: of_nat_eq_enat) haftmann@38621: qed huffman@29023: huffman@44019: lemma imult_is_0 [simp]: "((m::enat) * n = 0) = (m = 0 \ n = 0)" huffman@44019: by (auto simp add: times_enat_def zero_enat_def split: enat.split) nipkow@41853: huffman@44019: lemma imult_is_infinity: "((a::enat) * b = \) = (a = \ \ b \ 0 \ b = \ \ a \ 0)" huffman@44019: by (auto simp add: times_enat_def zero_enat_def split: enat.split) nipkow@41853: nipkow@41853: nipkow@41853: subsection {* Subtraction *} nipkow@41853: hoelzl@43919: instantiation enat :: minus nipkow@41853: begin nipkow@41853: hoelzl@43919: definition diff_enat_def: hoelzl@43924: "a - b = (case a of (enat x) \ (case b of (enat y) \ enat (x - y) | \ \ 0) nipkow@41853: | \ \ \)" nipkow@41853: nipkow@41853: instance .. nipkow@41853: nipkow@41853: end nipkow@41853: huffman@44019: lemma idiff_enat_enat [simp,code]: "enat a - enat b = enat (a - b)" huffman@44019: by (simp add: diff_enat_def) nipkow@41853: huffman@44019: lemma idiff_infinity [simp,code]: "\ - n = (\::enat)" huffman@44019: by (simp add: diff_enat_def) nipkow@41853: huffman@44019: lemma idiff_infinity_right [simp,code]: "enat a - \ = 0" huffman@44019: by (simp add: diff_enat_def) nipkow@41853: huffman@44019: lemma idiff_0 [simp]: "(0::enat) - n = 0" huffman@44019: by (cases n, simp_all add: zero_enat_def) nipkow@41853: huffman@44019: lemmas idiff_enat_0 [simp] = idiff_0 [unfolded zero_enat_def] nipkow@41853: huffman@44019: lemma idiff_0_right [simp]: "(n::enat) - 0 = n" huffman@44019: by (cases n) (simp_all add: zero_enat_def) nipkow@41853: huffman@44019: lemmas idiff_enat_0_right [simp] = idiff_0_right [unfolded zero_enat_def] nipkow@41853: huffman@44019: lemma idiff_self [simp]: "n \ \ \ (n::enat) - n = 0" huffman@44019: by (auto simp: zero_enat_def) nipkow@41853: huffman@44019: lemma eSuc_minus_eSuc [simp]: "eSuc n - eSuc m = n - m" huffman@44019: by (simp add: eSuc_def split: enat.split) nipkow@41855: huffman@44019: lemma eSuc_minus_1 [simp]: "eSuc n - 1 = n" huffman@44019: by (simp add: one_enat_def eSuc_enat[symmetric] zero_enat_def[symmetric]) nipkow@41855: hoelzl@43924: (*lemmas idiff_self_eq_0_enat = idiff_self_eq_0[unfolded zero_enat_def]*) nipkow@41853: haftmann@27110: subsection {* Ordering *} haftmann@27110: hoelzl@43919: instantiation enat :: linordered_ab_semigroup_add haftmann@27110: begin oheimb@11351: blanchet@38167: definition [nitpick_simp]: hoelzl@43924: "m \ n = (case n of enat n1 \ (case m of enat m1 \ m1 \ n1 | \ \ False) haftmann@27110: | \ \ True)" oheimb@11351: blanchet@38167: definition [nitpick_simp]: hoelzl@43924: "m < n = (case m of enat m1 \ (case n of enat n1 \ m1 < n1 | \ \ True) haftmann@27110: | \ \ False)" oheimb@11351: hoelzl@43919: lemma enat_ord_simps [simp]: hoelzl@43924: "enat m \ enat n \ m \ n" hoelzl@43924: "enat m < enat n \ m < n" hoelzl@43921: "q \ (\::enat)" hoelzl@43921: "q < (\::enat) \ q \ \" hoelzl@43921: "(\::enat) \ q \ q = \" hoelzl@43921: "(\::enat) < q \ False" hoelzl@43919: by (simp_all add: less_eq_enat_def less_enat_def split: enat.splits) oheimb@11351: noschinl@45934: lemma number_of_le_enat_iff[simp]: noschinl@45934: shows "number_of m \ enat n \ number_of m \ n" noschinl@45934: by (auto simp: number_of_enat_def) noschinl@45934: noschinl@45934: lemma number_of_less_enat_iff[simp]: noschinl@45934: shows "number_of m < enat n \ number_of m < n" noschinl@45934: by (auto simp: number_of_enat_def) noschinl@45934: hoelzl@43919: lemma enat_ord_code [code]: hoelzl@43924: "enat m \ enat n \ m \ n" hoelzl@43924: "enat m < enat n \ m < n" hoelzl@43921: "q \ (\::enat) \ True" hoelzl@43924: "enat m < \ \ True" hoelzl@43924: "\ \ enat n \ False" hoelzl@43921: "(\::enat) < q \ False" haftmann@27110: by simp_all oheimb@11351: haftmann@27110: instance by default hoelzl@43919: (auto simp add: less_eq_enat_def less_enat_def plus_enat_def split: enat.splits) oheimb@11351: haftmann@27110: end haftmann@27110: hoelzl@43919: instance enat :: ordered_comm_semiring huffman@29014: proof hoelzl@43919: fix a b c :: enat huffman@29014: assume "a \ b" and "0 \ c" huffman@29014: thus "c * a \ c * b" hoelzl@43919: unfolding times_enat_def less_eq_enat_def zero_enat_def hoelzl@43919: by (simp split: enat.splits) huffman@29014: qed huffman@29014: hoelzl@43919: lemma enat_ord_number [simp]: hoelzl@43919: "(number_of m \ enat) \ number_of n \ (number_of m \ nat) \ number_of n" hoelzl@43919: "(number_of m \ enat) < number_of n \ (number_of m \ nat) < number_of n" hoelzl@43919: by (simp_all add: number_of_enat_def) oheimb@11351: hoelzl@43919: lemma i0_lb [simp]: "(0\enat) \ n" hoelzl@43919: by (simp add: zero_enat_def less_eq_enat_def split: enat.splits) oheimb@11351: hoelzl@43919: lemma ile0_eq [simp]: "n \ (0\enat) \ n = 0" hoelzl@43919: by (simp add: zero_enat_def less_eq_enat_def split: enat.splits) oheimb@11351: huffman@44019: lemma infinity_ileE [elim!]: "\ \ enat m \ R" huffman@44019: by (simp add: zero_enat_def less_eq_enat_def split: enat.splits) huffman@44019: huffman@44019: lemma infinity_ilessE [elim!]: "\ < enat m \ R" haftmann@27110: by simp oheimb@11351: hoelzl@43919: lemma not_iless0 [simp]: "\ n < (0\enat)" hoelzl@43919: by (simp add: zero_enat_def less_enat_def split: enat.splits) haftmann@27110: hoelzl@43919: lemma i0_less [simp]: "(0\enat) < n \ n \ 0" huffman@44019: by (simp add: zero_enat_def less_enat_def split: enat.splits) oheimb@11351: huffman@44019: lemma eSuc_ile_mono [simp]: "eSuc n \ eSuc m \ n \ m" huffman@44019: by (simp add: eSuc_def less_eq_enat_def split: enat.splits) haftmann@27110: huffman@44019: lemma eSuc_mono [simp]: "eSuc n < eSuc m \ n < m" huffman@44019: by (simp add: eSuc_def less_enat_def split: enat.splits) oheimb@11351: huffman@44019: lemma ile_eSuc [simp]: "n \ eSuc n" huffman@44019: by (simp add: eSuc_def less_eq_enat_def split: enat.splits) oheimb@11351: huffman@44019: lemma not_eSuc_ilei0 [simp]: "\ eSuc n \ 0" huffman@44019: by (simp add: zero_enat_def eSuc_def less_eq_enat_def split: enat.splits) haftmann@27110: huffman@44019: lemma i0_iless_eSuc [simp]: "0 < eSuc n" huffman@44019: by (simp add: zero_enat_def eSuc_def less_enat_def split: enat.splits) haftmann@27110: huffman@44019: lemma iless_eSuc0[simp]: "(n < eSuc 0) = (n = 0)" huffman@44019: by (simp add: zero_enat_def eSuc_def less_enat_def split: enat.split) nipkow@41853: huffman@44019: lemma ileI1: "m < n \ eSuc m \ n" huffman@44019: by (simp add: eSuc_def less_eq_enat_def less_enat_def split: enat.splits) haftmann@27110: hoelzl@43924: lemma Suc_ile_eq: "enat (Suc m) \ n \ enat m < n" haftmann@27110: by (cases n) auto haftmann@27110: huffman@44019: lemma iless_Suc_eq [simp]: "enat m < eSuc n \ enat m \ n" huffman@44019: by (auto simp add: eSuc_def less_enat_def split: enat.splits) oheimb@11351: huffman@44019: lemma imult_infinity: "(0::enat) < n \ \ * n = \" huffman@44019: by (simp add: zero_enat_def less_enat_def split: enat.splits) nipkow@41853: huffman@44019: lemma imult_infinity_right: "(0::enat) < n \ n * \ = \" huffman@44019: by (simp add: zero_enat_def less_enat_def split: enat.splits) nipkow@41853: hoelzl@43919: lemma enat_0_less_mult_iff: "(0 < (m::enat) * n) = (0 < m \ 0 < n)" huffman@44019: by (simp only: i0_less imult_is_0, simp) nipkow@41853: huffman@44019: lemma mono_eSuc: "mono eSuc" huffman@44019: by (simp add: mono_def) nipkow@41853: nipkow@41853: hoelzl@43919: lemma min_enat_simps [simp]: hoelzl@43924: "min (enat m) (enat n) = enat (min m n)" haftmann@27110: "min q 0 = 0" haftmann@27110: "min 0 q = 0" hoelzl@43921: "min q (\::enat) = q" hoelzl@43921: "min (\::enat) q = q" haftmann@27110: by (auto simp add: min_def) oheimb@11351: hoelzl@43919: lemma max_enat_simps [simp]: hoelzl@43924: "max (enat m) (enat n) = enat (max m n)" haftmann@27110: "max q 0 = q" haftmann@27110: "max 0 q = q" hoelzl@43921: "max q \ = (\::enat)" hoelzl@43921: "max \ q = (\::enat)" haftmann@27110: by (simp_all add: max_def) haftmann@27110: hoelzl@43924: lemma enat_ile: "n \ enat m \ \k. n = enat k" haftmann@27110: by (cases n) simp_all haftmann@27110: hoelzl@43924: lemma enat_iless: "n < enat m \ \k. n = enat k" haftmann@27110: by (cases n) simp_all oheimb@11351: hoelzl@43924: lemma chain_incr: "\i. \j. Y i < Y j ==> \j. enat k < Y j" nipkow@25134: apply (induct_tac k) hoelzl@43924: apply (simp (no_asm) only: enat_0) haftmann@27110: apply (fast intro: le_less_trans [OF i0_lb]) nipkow@25134: apply (erule exE) nipkow@25134: apply (drule spec) nipkow@25134: apply (erule exE) nipkow@25134: apply (drule ileI1) huffman@44019: apply (rule eSuc_enat [THEN subst]) nipkow@25134: apply (rule exI) haftmann@27110: apply (erule (1) le_less_trans) nipkow@25134: done oheimb@11351: hoelzl@43919: instantiation enat :: "{bot, top}" haftmann@29337: begin haftmann@29337: hoelzl@43919: definition bot_enat :: enat where hoelzl@43919: "bot_enat = 0" haftmann@29337: hoelzl@43919: definition top_enat :: enat where hoelzl@43919: "top_enat = \" haftmann@29337: haftmann@29337: instance proof hoelzl@43919: qed (simp_all add: bot_enat_def top_enat_def) haftmann@29337: haftmann@29337: end haftmann@29337: hoelzl@43924: lemma finite_enat_bounded: hoelzl@43924: assumes le_fin: "\y. y \ A \ y \ enat n" noschinl@42993: shows "finite A" noschinl@42993: proof (rule finite_subset) hoelzl@43924: show "finite (enat ` {..n})" by blast noschinl@42993: nipkow@44890: have "A \ {..enat n}" using le_fin by fastforce hoelzl@43924: also have "\ \ enat ` {..n}" noschinl@42993: by (rule subsetI) (case_tac x, auto) hoelzl@43924: finally show "A \ enat ` {..n}" . noschinl@42993: qed noschinl@42993: huffman@26089: huffman@45775: subsection {* Cancellation simprocs *} huffman@45775: huffman@45775: lemma enat_add_left_cancel: "a + b = a + c \ a = (\::enat) \ b = c" huffman@45775: unfolding plus_enat_def by (simp split: enat.split) huffman@45775: huffman@45775: lemma enat_add_left_cancel_le: "a + b \ a + c \ a = (\::enat) \ b \ c" huffman@45775: unfolding plus_enat_def by (simp split: enat.split) huffman@45775: huffman@45775: lemma enat_add_left_cancel_less: "a + b < a + c \ a \ (\::enat) \ b < c" huffman@45775: unfolding plus_enat_def by (simp split: enat.split) huffman@45775: huffman@45775: ML {* huffman@45775: structure Cancel_Enat_Common = huffman@45775: struct huffman@45775: (* copied from src/HOL/Tools/nat_numeral_simprocs.ML *) huffman@45775: fun find_first_t _ _ [] = raise TERM("find_first_t", []) huffman@45775: | find_first_t past u (t::terms) = huffman@45775: if u aconv t then (rev past @ terms) huffman@45775: else find_first_t (t::past) u terms huffman@45775: huffman@45775: val mk_sum = Arith_Data.long_mk_sum huffman@45775: val dest_sum = Arith_Data.dest_sum huffman@45775: val find_first = find_first_t [] huffman@45775: val trans_tac = Numeral_Simprocs.trans_tac huffman@45775: val norm_ss = HOL_basic_ss addsimps huffman@45775: @{thms add_ac semiring_numeral_0_eq_0 add_0_left add_0_right} huffman@45775: fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) huffman@45775: fun simplify_meta_eq ss cancel_th th = huffman@45775: Arith_Data.simplify_meta_eq @{thms semiring_numeral_0_eq_0} ss huffman@45775: ([th, cancel_th] MRS trans) huffman@45775: fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) huffman@45775: end huffman@45775: huffman@45775: structure Eq_Enat_Cancel = ExtractCommonTermFun huffman@45775: (open Cancel_Enat_Common huffman@45775: val mk_bal = HOLogic.mk_eq huffman@45775: val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} @{typ enat} huffman@45775: fun simp_conv _ _ = SOME @{thm enat_add_left_cancel} huffman@45775: ) huffman@45775: huffman@45775: structure Le_Enat_Cancel = ExtractCommonTermFun huffman@45775: (open Cancel_Enat_Common huffman@45775: val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less_eq} huffman@45775: val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} @{typ enat} huffman@45775: fun simp_conv _ _ = SOME @{thm enat_add_left_cancel_le} huffman@45775: ) huffman@45775: huffman@45775: structure Less_Enat_Cancel = ExtractCommonTermFun huffman@45775: (open Cancel_Enat_Common huffman@45775: val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less} huffman@45775: val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} @{typ enat} huffman@45775: fun simp_conv _ _ = SOME @{thm enat_add_left_cancel_less} huffman@45775: ) huffman@45775: *} huffman@45775: huffman@45775: simproc_setup enat_eq_cancel huffman@45775: ("(l::enat) + m = n" | "(l::enat) = m + n") = huffman@45775: {* fn phi => fn ss => fn ct => Eq_Enat_Cancel.proc ss (term_of ct) *} huffman@45775: huffman@45775: simproc_setup enat_le_cancel huffman@45775: ("(l::enat) + m \ n" | "(l::enat) \ m + n") = huffman@45775: {* fn phi => fn ss => fn ct => Le_Enat_Cancel.proc ss (term_of ct) *} huffman@45775: huffman@45775: simproc_setup enat_less_cancel huffman@45775: ("(l::enat) + m < n" | "(l::enat) < m + n") = huffman@45775: {* fn phi => fn ss => fn ct => Less_Enat_Cancel.proc ss (term_of ct) *} huffman@45775: huffman@45775: text {* TODO: add regression tests for these simprocs *} huffman@45775: huffman@45775: text {* TODO: add simprocs for combining and cancelling numerals *} huffman@45775: huffman@45775: haftmann@27110: subsection {* Well-ordering *} huffman@26089: hoelzl@43924: lemma less_enatE: hoelzl@43924: "[| n < enat m; !!k. n = enat k ==> k < m ==> P |] ==> P" huffman@26089: by (induct n) auto huffman@26089: huffman@44019: lemma less_infinityE: hoelzl@43924: "[| n < \; !!k. n = enat k ==> P |] ==> P" huffman@26089: by (induct n) auto huffman@26089: hoelzl@43919: lemma enat_less_induct: hoelzl@43919: assumes prem: "!!n. \m::enat. m < n --> P m ==> P n" shows "P n" huffman@26089: proof - hoelzl@43924: have P_enat: "!!k. P (enat k)" huffman@26089: apply (rule nat_less_induct) huffman@26089: apply (rule prem, clarify) hoelzl@43924: apply (erule less_enatE, simp) huffman@26089: done huffman@26089: show ?thesis huffman@26089: proof (induct n) huffman@26089: fix nat hoelzl@43924: show "P (enat nat)" by (rule P_enat) huffman@26089: next hoelzl@43921: show "P \" huffman@26089: apply (rule prem, clarify) huffman@44019: apply (erule less_infinityE) hoelzl@43924: apply (simp add: P_enat) huffman@26089: done huffman@26089: qed huffman@26089: qed huffman@26089: hoelzl@43919: instance enat :: wellorder huffman@26089: proof haftmann@27823: fix P and n hoelzl@43919: assume hyp: "(\n\enat. (\m\enat. m < n \ P m) \ P n)" hoelzl@43919: show "P n" by (blast intro: enat_less_induct hyp) huffman@26089: qed huffman@26089: noschinl@42993: subsection {* Complete Lattice *} noschinl@42993: hoelzl@43919: instantiation enat :: complete_lattice noschinl@42993: begin noschinl@42993: hoelzl@43919: definition inf_enat :: "enat \ enat \ enat" where hoelzl@43919: "inf_enat \ min" noschinl@42993: hoelzl@43919: definition sup_enat :: "enat \ enat \ enat" where hoelzl@43919: "sup_enat \ max" noschinl@42993: hoelzl@43919: definition Inf_enat :: "enat set \ enat" where hoelzl@43919: "Inf_enat A \ if A = {} then \ else (LEAST x. x \ A)" noschinl@42993: hoelzl@43919: definition Sup_enat :: "enat set \ enat" where hoelzl@43919: "Sup_enat A \ if A = {} then 0 noschinl@42993: else if finite A then Max A noschinl@42993: else \" noschinl@42993: instance proof hoelzl@43919: fix x :: "enat" and A :: "enat set" noschinl@42993: { assume "x \ A" then show "Inf A \ x" hoelzl@43919: unfolding Inf_enat_def by (auto intro: Least_le) } noschinl@42993: { assume "\y. y \ A \ x \ y" then show "x \ Inf A" hoelzl@43919: unfolding Inf_enat_def noschinl@42993: by (cases "A = {}") (auto intro: LeastI2_ex) } noschinl@42993: { assume "x \ A" then show "x \ Sup A" hoelzl@43919: unfolding Sup_enat_def by (cases "finite A") auto } noschinl@42993: { assume "\y. y \ A \ y \ x" then show "Sup A \ x" hoelzl@43924: unfolding Sup_enat_def using finite_enat_bounded by auto } hoelzl@43919: qed (simp_all add: inf_enat_def sup_enat_def) noschinl@42993: end noschinl@42993: hoelzl@43978: instance enat :: complete_linorder .. haftmann@27110: haftmann@27110: subsection {* Traditional theorem names *} haftmann@27110: huffman@44019: lemmas enat_defs = zero_enat_def one_enat_def number_of_enat_def eSuc_def hoelzl@43919: plus_enat_def less_eq_enat_def less_enat_def haftmann@27110: oheimb@11351: end