src/HOL/Library/Nat_Infinity.thy
author nipkow
Sun, 10 May 2009 14:21:41 +0200
changeset 31084 f4db921165ce
parent 31077 28dd6fd3d184
child 31094 7d6edb28bdbc
permissions -rw-r--r--
fixed HOLCF proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11355
wenzelm
parents: 11351
diff changeset
     1
(*  Title:      HOL/Library/Nat_Infinity.thy
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
     2
    Author:     David von Oheimb, TU Muenchen;  Florian Haftmann, TU Muenchen
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
     3
*)
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
     4
14706
71590b7733b7 tuned document;
wenzelm
parents: 14691
diff changeset
     5
header {* Natural numbers with infinity *}
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
     6
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14981
diff changeset
     7
theory Nat_Infinity
30663
0b6aff7451b2 Main is (Complex_Main) base entry point in library theories
haftmann
parents: 29668
diff changeset
     8
imports Main
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14981
diff changeset
     9
begin
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    10
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    11
subsection {* Type definition *}
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    12
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    13
text {*
11355
wenzelm
parents: 11351
diff changeset
    14
  We extend the standard natural numbers by a special value indicating
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    15
  infinity.
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    16
*}
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    17
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    18
datatype inat = Fin nat | Infty
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    19
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 19736
diff changeset
    20
notation (xsymbols)
19736
wenzelm
parents: 15140
diff changeset
    21
  Infty  ("\<infinity>")
wenzelm
parents: 15140
diff changeset
    22
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 19736
diff changeset
    23
notation (HTML output)
19736
wenzelm
parents: 15140
diff changeset
    24
  Infty  ("\<infinity>")
wenzelm
parents: 15140
diff changeset
    25
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    26
31084
f4db921165ce fixed HOLCF proofs
nipkow
parents: 31077
diff changeset
    27
lemma not_Infty_eq[iff]: "(x ~= Infty) = (EX i. x = Fin i)"
f4db921165ce fixed HOLCF proofs
nipkow
parents: 31077
diff changeset
    28
by (cases x) auto
f4db921165ce fixed HOLCF proofs
nipkow
parents: 31077
diff changeset
    29
f4db921165ce fixed HOLCF proofs
nipkow
parents: 31077
diff changeset
    30
lemma not_Fin_eq [iff]: "(ALL y. x ~= Fin y) = (x = Infty)"
31077
28dd6fd3d184 more lemmas
nipkow
parents: 30663
diff changeset
    31
by (cases x) auto
28dd6fd3d184 more lemmas
nipkow
parents: 30663
diff changeset
    32
28dd6fd3d184 more lemmas
nipkow
parents: 30663
diff changeset
    33
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    34
subsection {* Constructors and numbers *}
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    35
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    36
instantiation inat :: "{zero, one, number}"
25594
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    37
begin
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    38
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    39
definition
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    40
  "0 = Fin 0"
25594
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    41
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    42
definition
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    43
  [code inline]: "1 = Fin 1"
25594
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    44
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    45
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27823
diff changeset
    46
  [code inline, code del]: "number_of k = Fin (number_of k)"
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    47
25594
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    48
instance ..
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    49
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    50
end
43c718438f9f switched import from Main to PreList
haftmann
parents: 25134
diff changeset
    51
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    52
definition iSuc :: "inat \<Rightarrow> inat" where
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    53
  "iSuc i = (case i of Fin n \<Rightarrow> Fin (Suc n) | \<infinity> \<Rightarrow> \<infinity>)"
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    54
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    55
lemma Fin_0: "Fin 0 = 0"
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    56
  by (simp add: zero_inat_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    57
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    58
lemma Fin_1: "Fin 1 = 1"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    59
  by (simp add: one_inat_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    60
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    61
lemma Fin_number: "Fin (number_of k) = number_of k"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    62
  by (simp add: number_of_inat_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    63
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    64
lemma one_iSuc: "1 = iSuc 0"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    65
  by (simp add: zero_inat_def one_inat_def iSuc_def)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    66
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    67
lemma Infty_ne_i0 [simp]: "\<infinity> \<noteq> 0"
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    68
  by (simp add: zero_inat_def)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    69
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    70
lemma i0_ne_Infty [simp]: "0 \<noteq> \<infinity>"
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    71
  by (simp add: zero_inat_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    72
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    73
lemma zero_inat_eq [simp]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    74
  "number_of k = (0\<Colon>inat) \<longleftrightarrow> number_of k = (0\<Colon>nat)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    75
  "(0\<Colon>inat) = number_of k \<longleftrightarrow> number_of k = (0\<Colon>nat)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    76
  unfolding zero_inat_def number_of_inat_def by simp_all
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    77
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    78
lemma one_inat_eq [simp]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    79
  "number_of k = (1\<Colon>inat) \<longleftrightarrow> number_of k = (1\<Colon>nat)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    80
  "(1\<Colon>inat) = number_of k \<longleftrightarrow> number_of k = (1\<Colon>nat)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    81
  unfolding one_inat_def number_of_inat_def by simp_all
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    82
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    83
lemma zero_one_inat_neq [simp]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    84
  "\<not> 0 = (1\<Colon>inat)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    85
  "\<not> 1 = (0\<Colon>inat)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    86
  unfolding zero_inat_def one_inat_def by simp_all
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
    87
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    88
lemma Infty_ne_i1 [simp]: "\<infinity> \<noteq> 1"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    89
  by (simp add: one_inat_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    90
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    91
lemma i1_ne_Infty [simp]: "1 \<noteq> \<infinity>"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    92
  by (simp add: one_inat_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    93
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    94
lemma Infty_ne_number [simp]: "\<infinity> \<noteq> number_of k"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    95
  by (simp add: number_of_inat_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    96
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    97
lemma number_ne_Infty [simp]: "number_of k \<noteq> \<infinity>"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    98
  by (simp add: number_of_inat_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
    99
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   100
lemma iSuc_Fin: "iSuc (Fin n) = Fin (Suc n)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   101
  by (simp add: iSuc_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   102
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   103
lemma iSuc_number_of: "iSuc (number_of k) = Fin (Suc (number_of k))"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   104
  by (simp add: iSuc_Fin number_of_inat_def)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   105
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   106
lemma iSuc_Infty [simp]: "iSuc \<infinity> = \<infinity>"
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   107
  by (simp add: iSuc_def)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   108
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   109
lemma iSuc_ne_0 [simp]: "iSuc n \<noteq> 0"
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   110
  by (simp add: iSuc_def zero_inat_def split: inat.splits)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   111
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   112
lemma zero_ne_iSuc [simp]: "0 \<noteq> iSuc n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   113
  by (rule iSuc_ne_0 [symmetric])
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   114
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   115
lemma iSuc_inject [simp]: "iSuc m = iSuc n \<longleftrightarrow> m = n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   116
  by (simp add: iSuc_def split: inat.splits)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   117
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   118
lemma number_of_inat_inject [simp]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   119
  "(number_of k \<Colon> inat) = number_of l \<longleftrightarrow> (number_of k \<Colon> nat) = number_of l"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   120
  by (simp add: number_of_inat_def)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   121
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   122
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   123
subsection {* Addition *}
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   124
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   125
instantiation inat :: comm_monoid_add
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   126
begin
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   127
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   128
definition
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   129
  [code del]: "m + n = (case m of \<infinity> \<Rightarrow> \<infinity> | Fin m \<Rightarrow> (case n of \<infinity> \<Rightarrow> \<infinity> | Fin n \<Rightarrow> Fin (m + n)))"
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   130
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   131
lemma plus_inat_simps [simp, code]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   132
  "Fin m + Fin n = Fin (m + n)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   133
  "\<infinity> + q = \<infinity>"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   134
  "q + \<infinity> = \<infinity>"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   135
  by (simp_all add: plus_inat_def split: inat.splits)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   136
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   137
instance proof
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   138
  fix n m q :: inat
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   139
  show "n + m + q = n + (m + q)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   140
    by (cases n, auto, cases m, auto, cases q, auto)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   141
  show "n + m = m + n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   142
    by (cases n, auto, cases m, auto)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   143
  show "0 + n = n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   144
    by (cases n) (simp_all add: zero_inat_def)
26089
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   145
qed
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   146
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   147
end
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   148
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   149
lemma plus_inat_0 [simp]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   150
  "0 + (q\<Colon>inat) = q"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   151
  "(q\<Colon>inat) + 0 = q"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   152
  by (simp_all add: plus_inat_def zero_inat_def split: inat.splits)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   153
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   154
lemma plus_inat_number [simp]:
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 28562
diff changeset
   155
  "(number_of k \<Colon> inat) + number_of l = (if k < Int.Pls then number_of l
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 28562
diff changeset
   156
    else if l < Int.Pls then number_of k else number_of (k + l))"
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   157
  unfolding number_of_inat_def plus_inat_simps nat_arith(1) if_distrib [symmetric, of _ Fin] ..
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   158
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   159
lemma iSuc_number [simp]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   160
  "iSuc (number_of k) = (if neg (number_of k \<Colon> int) then 1 else number_of (Int.succ k))"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   161
  unfolding iSuc_number_of
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   162
  unfolding one_inat_def number_of_inat_def Suc_nat_number_of if_distrib [symmetric] ..
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   163
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   164
lemma iSuc_plus_1:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   165
  "iSuc n = n + 1"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   166
  by (cases n) (simp_all add: iSuc_Fin one_inat_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   167
  
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   168
lemma plus_1_iSuc:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   169
  "1 + q = iSuc q"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   170
  "q + 1 = iSuc q"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   171
  unfolding iSuc_plus_1 by (simp_all add: add_ac)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   172
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   173
29014
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   174
subsection {* Multiplication *}
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   175
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   176
instantiation inat :: comm_semiring_1
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   177
begin
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   178
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   179
definition
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   180
  times_inat_def [code del]:
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   181
  "m * n = (case m of \<infinity> \<Rightarrow> if n = 0 then 0 else \<infinity> | Fin m \<Rightarrow>
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   182
    (case n of \<infinity> \<Rightarrow> if m = 0 then 0 else \<infinity> | Fin n \<Rightarrow> Fin (m * n)))"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   183
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   184
lemma times_inat_simps [simp, code]:
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   185
  "Fin m * Fin n = Fin (m * n)"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   186
  "\<infinity> * \<infinity> = \<infinity>"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   187
  "\<infinity> * Fin n = (if n = 0 then 0 else \<infinity>)"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   188
  "Fin m * \<infinity> = (if m = 0 then 0 else \<infinity>)"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   189
  unfolding times_inat_def zero_inat_def
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   190
  by (simp_all split: inat.split)
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   191
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   192
instance proof
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   193
  fix a b c :: inat
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   194
  show "(a * b) * c = a * (b * c)"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   195
    unfolding times_inat_def zero_inat_def
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   196
    by (simp split: inat.split)
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   197
  show "a * b = b * a"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   198
    unfolding times_inat_def zero_inat_def
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   199
    by (simp split: inat.split)
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   200
  show "1 * a = a"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   201
    unfolding times_inat_def zero_inat_def one_inat_def
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   202
    by (simp split: inat.split)
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   203
  show "(a + b) * c = a * c + b * c"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   204
    unfolding times_inat_def zero_inat_def
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   205
    by (simp split: inat.split add: left_distrib)
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   206
  show "0 * a = 0"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   207
    unfolding times_inat_def zero_inat_def
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   208
    by (simp split: inat.split)
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   209
  show "a * 0 = 0"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   210
    unfolding times_inat_def zero_inat_def
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   211
    by (simp split: inat.split)
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   212
  show "(0::inat) \<noteq> 1"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   213
    unfolding zero_inat_def one_inat_def
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   214
    by simp
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   215
qed
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   216
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   217
end
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   218
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   219
lemma mult_iSuc: "iSuc m * n = n + m * n"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29337
diff changeset
   220
  unfolding iSuc_plus_1 by (simp add: algebra_simps)
29014
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   221
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   222
lemma mult_iSuc_right: "m * iSuc n = m + m * n"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29337
diff changeset
   223
  unfolding iSuc_plus_1 by (simp add: algebra_simps)
29014
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   224
29023
ef3adebc6d98 instance inat :: semiring_char_0
huffman
parents: 29014
diff changeset
   225
lemma of_nat_eq_Fin: "of_nat n = Fin n"
ef3adebc6d98 instance inat :: semiring_char_0
huffman
parents: 29014
diff changeset
   226
  apply (induct n)
ef3adebc6d98 instance inat :: semiring_char_0
huffman
parents: 29014
diff changeset
   227
  apply (simp add: Fin_0)
ef3adebc6d98 instance inat :: semiring_char_0
huffman
parents: 29014
diff changeset
   228
  apply (simp add: plus_1_iSuc iSuc_Fin)
ef3adebc6d98 instance inat :: semiring_char_0
huffman
parents: 29014
diff changeset
   229
  done
ef3adebc6d98 instance inat :: semiring_char_0
huffman
parents: 29014
diff changeset
   230
ef3adebc6d98 instance inat :: semiring_char_0
huffman
parents: 29014
diff changeset
   231
instance inat :: semiring_char_0
ef3adebc6d98 instance inat :: semiring_char_0
huffman
parents: 29014
diff changeset
   232
  by default (simp add: of_nat_eq_Fin)
ef3adebc6d98 instance inat :: semiring_char_0
huffman
parents: 29014
diff changeset
   233
29014
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   234
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   235
subsection {* Ordering *}
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   236
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   237
instantiation inat :: ordered_ab_semigroup_add
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   238
begin
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   239
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   240
definition
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   241
  [code del]: "m \<le> n = (case n of Fin n1 \<Rightarrow> (case m of Fin m1 \<Rightarrow> m1 \<le> n1 | \<infinity> \<Rightarrow> False)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   242
    | \<infinity> \<Rightarrow> True)"
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   243
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   244
definition
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   245
  [code del]: "m < n = (case m of Fin m1 \<Rightarrow> (case n of Fin n1 \<Rightarrow> m1 < n1 | \<infinity> \<Rightarrow> True)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   246
    | \<infinity> \<Rightarrow> False)"
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   247
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   248
lemma inat_ord_simps [simp]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   249
  "Fin m \<le> Fin n \<longleftrightarrow> m \<le> n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   250
  "Fin m < Fin n \<longleftrightarrow> m < n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   251
  "q \<le> \<infinity>"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   252
  "q < \<infinity> \<longleftrightarrow> q \<noteq> \<infinity>"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   253
  "\<infinity> \<le> q \<longleftrightarrow> q = \<infinity>"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   254
  "\<infinity> < q \<longleftrightarrow> False"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   255
  by (simp_all add: less_eq_inat_def less_inat_def split: inat.splits)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   256
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   257
lemma inat_ord_code [code]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   258
  "Fin m \<le> Fin n \<longleftrightarrow> m \<le> n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   259
  "Fin m < Fin n \<longleftrightarrow> m < n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   260
  "q \<le> \<infinity> \<longleftrightarrow> True"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   261
  "Fin m < \<infinity> \<longleftrightarrow> True"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   262
  "\<infinity> \<le> Fin n \<longleftrightarrow> False"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   263
  "\<infinity> < q \<longleftrightarrow> False"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   264
  by simp_all
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   265
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   266
instance by default
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   267
  (auto simp add: less_eq_inat_def less_inat_def plus_inat_def split: inat.splits)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   268
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   269
end
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   270
31077
28dd6fd3d184 more lemmas
nipkow
parents: 30663
diff changeset
   271
instance inat :: linorder
28dd6fd3d184 more lemmas
nipkow
parents: 30663
diff changeset
   272
by intro_classes (auto simp add: less_eq_inat_def split: inat.splits)
28dd6fd3d184 more lemmas
nipkow
parents: 30663
diff changeset
   273
29014
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   274
instance inat :: pordered_comm_semiring
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   275
proof
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   276
  fix a b c :: inat
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   277
  assume "a \<le> b" and "0 \<le> c"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   278
  thus "c * a \<le> c * b"
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   279
    unfolding times_inat_def less_eq_inat_def zero_inat_def
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   280
    by (simp split: inat.splits)
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   281
qed
e515f42d1db7 multiplication for type inat
huffman
parents: 29012
diff changeset
   282
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   283
lemma inat_ord_number [simp]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   284
  "(number_of m \<Colon> inat) \<le> number_of n \<longleftrightarrow> (number_of m \<Colon> nat) \<le> number_of n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   285
  "(number_of m \<Colon> inat) < number_of n \<longleftrightarrow> (number_of m \<Colon> nat) < number_of n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   286
  by (simp_all add: number_of_inat_def)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   287
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   288
lemma i0_lb [simp]: "(0\<Colon>inat) \<le> n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   289
  by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   290
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   291
lemma i0_neq [simp]: "n \<le> (0\<Colon>inat) \<longleftrightarrow> n = 0"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   292
  by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   293
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   294
lemma Infty_ileE [elim!]: "\<infinity> \<le> Fin m \<Longrightarrow> R"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   295
  by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   296
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   297
lemma Infty_ilessE [elim!]: "\<infinity> < Fin m \<Longrightarrow> R"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   298
  by simp
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   299
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   300
lemma not_ilessi0 [simp]: "\<not> n < (0\<Colon>inat)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   301
  by (simp add: zero_inat_def less_inat_def split: inat.splits)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   302
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   303
lemma i0_eq [simp]: "(0\<Colon>inat) < n \<longleftrightarrow> n \<noteq> 0"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   304
  by (simp add: zero_inat_def less_inat_def split: inat.splits)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   305
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   306
lemma iSuc_ile_mono [simp]: "iSuc n \<le> iSuc m \<longleftrightarrow> n \<le> m"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   307
  by (simp add: iSuc_def less_eq_inat_def split: inat.splits)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   308
 
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   309
lemma iSuc_mono [simp]: "iSuc n < iSuc m \<longleftrightarrow> n < m"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   310
  by (simp add: iSuc_def less_inat_def split: inat.splits)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   311
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   312
lemma ile_iSuc [simp]: "n \<le> iSuc n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   313
  by (simp add: iSuc_def less_eq_inat_def split: inat.splits)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   314
11355
wenzelm
parents: 11351
diff changeset
   315
lemma not_iSuc_ilei0 [simp]: "\<not> iSuc n \<le> 0"
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   316
  by (simp add: zero_inat_def iSuc_def less_eq_inat_def split: inat.splits)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   317
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   318
lemma i0_iless_iSuc [simp]: "0 < iSuc n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   319
  by (simp add: zero_inat_def iSuc_def less_inat_def split: inat.splits)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   320
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   321
lemma ileI1: "m < n \<Longrightarrow> iSuc m \<le> n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   322
  by (simp add: iSuc_def less_eq_inat_def less_inat_def split: inat.splits)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   323
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   324
lemma Suc_ile_eq: "Fin (Suc m) \<le> n \<longleftrightarrow> Fin m < n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   325
  by (cases n) auto
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   326
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   327
lemma iless_Suc_eq [simp]: "Fin m < iSuc n \<longleftrightarrow> Fin m \<le> n"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   328
  by (auto simp add: iSuc_def less_inat_def split: inat.splits)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   329
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   330
lemma min_inat_simps [simp]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   331
  "min (Fin m) (Fin n) = Fin (min m n)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   332
  "min q 0 = 0"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   333
  "min 0 q = 0"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   334
  "min q \<infinity> = q"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   335
  "min \<infinity> q = q"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   336
  by (auto simp add: min_def)
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   337
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   338
lemma max_inat_simps [simp]:
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   339
  "max (Fin m) (Fin n) = Fin (max m n)"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   340
  "max q 0 = q"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   341
  "max 0 q = q"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   342
  "max q \<infinity> = \<infinity>"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   343
  "max \<infinity> q = \<infinity>"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   344
  by (simp_all add: max_def)
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   345
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   346
lemma Fin_ile: "n \<le> Fin m \<Longrightarrow> \<exists>k. n = Fin k"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   347
  by (cases n) simp_all
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   348
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   349
lemma Fin_iless: "n < Fin m \<Longrightarrow> \<exists>k. n = Fin k"
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   350
  by (cases n) simp_all
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   352
lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. Fin k < Y j"
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   353
apply (induct_tac k)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   354
 apply (simp (no_asm) only: Fin_0)
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   355
 apply (fast intro: le_less_trans [OF i0_lb])
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   356
apply (erule exE)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   357
apply (drule spec)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   358
apply (erule exE)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   359
apply (drule ileI1)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   360
apply (rule iSuc_Fin [THEN subst])
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   361
apply (rule exI)
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   362
apply (erule (1) le_less_trans)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   363
done
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   364
29337
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   365
instantiation inat :: "{bot, top}"
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   366
begin
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   367
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   368
definition bot_inat :: inat where
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   369
  "bot_inat = 0"
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   370
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   371
definition top_inat :: inat where
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   372
  "top_inat = \<infinity>"
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   373
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   374
instance proof
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   375
qed (simp_all add: bot_inat_def top_inat_def)
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   376
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   377
end
450805a4a91f added instance for bot, top
haftmann
parents: 29023
diff changeset
   378
26089
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   379
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   380
subsection {* Well-ordering *}
26089
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   381
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   382
lemma less_FinE:
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   383
  "[| n < Fin m; !!k. n = Fin k ==> k < m ==> P |] ==> P"
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   384
by (induct n) auto
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   385
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   386
lemma less_InftyE:
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   387
  "[| n < Infty; !!k. n = Fin k ==> P |] ==> P"
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   388
by (induct n) auto
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   389
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   390
lemma inat_less_induct:
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   391
  assumes prem: "!!n. \<forall>m::inat. m < n --> P m ==> P n" shows "P n"
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   392
proof -
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   393
  have P_Fin: "!!k. P (Fin k)"
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   394
    apply (rule nat_less_induct)
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   395
    apply (rule prem, clarify)
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   396
    apply (erule less_FinE, simp)
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   397
    done
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   398
  show ?thesis
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   399
  proof (induct n)
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   400
    fix nat
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   401
    show "P (Fin nat)" by (rule P_Fin)
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   402
  next
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   403
    show "P Infty"
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   404
      apply (rule prem, clarify)
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   405
      apply (erule less_InftyE)
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   406
      apply (simp add: P_Fin)
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   407
      done
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   408
  qed
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   409
qed
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   410
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   411
instance inat :: wellorder
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   412
proof
27823
52971512d1a2 moved class wellorder to theory Orderings
haftmann
parents: 27487
diff changeset
   413
  fix P and n
52971512d1a2 moved class wellorder to theory Orderings
haftmann
parents: 27487
diff changeset
   414
  assume hyp: "(\<And>n\<Colon>inat. (\<And>m\<Colon>inat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
52971512d1a2 moved class wellorder to theory Orderings
haftmann
parents: 27487
diff changeset
   415
  show "P n" by (blast intro: inat_less_induct hyp)
26089
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   416
qed
373221497340 proved linorder and wellorder class instances
huffman
parents: 25691
diff changeset
   417
27110
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   418
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   419
subsection {* Traditional theorem names *}
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   420
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   421
lemmas inat_defs = zero_inat_def one_inat_def number_of_inat_def iSuc_def
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   422
  plus_inat_def less_eq_inat_def less_inat_def
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   423
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   424
lemmas inat_splits = inat.splits
194aa674c2a1 refactoring; addition, numerals
haftmann
parents: 26089
diff changeset
   425
31077
28dd6fd3d184 more lemmas
nipkow
parents: 30663
diff changeset
   426
28dd6fd3d184 more lemmas
nipkow
parents: 30663
diff changeset
   427
instance inat :: linorder
28dd6fd3d184 more lemmas
nipkow
parents: 30663
diff changeset
   428
by intro_classes (auto simp add: inat_defs split: inat.splits)
28dd6fd3d184 more lemmas
nipkow
parents: 30663
diff changeset
   429
11351
c5c403d30c77 added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff changeset
   430
end