src/HOL/Library/Nat_Infinity.thy
author wenzelm
Fri, 05 Oct 2001 21:52:39 +0200
changeset 11701 3d51fbf81c17
parent 11655 923e4d0d36d5
child 14565 c6dc17aab88a
permissions -rw-r--r--
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
"num" syntax (still with "#"), Numeral0, Numeral1;
wenzelm@11355
     1
(*  Title:      HOL/Library/Nat_Infinity.thy
wenzelm@11355
     2
    ID:         $Id$
wenzelm@11355
     3
    Author:     David von Oheimb, TU Muenchen
oheimb@11351
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
oheimb@11351
     5
*)
oheimb@11351
     6
oheimb@11351
     7
header {*
oheimb@11351
     8
  \title{Natural numbers with infinity}
oheimb@11351
     9
  \author{David von Oheimb}
oheimb@11351
    10
*}
oheimb@11351
    11
wenzelm@11355
    12
theory Nat_Infinity = Main:
oheimb@11351
    13
oheimb@11351
    14
subsection "Definitions"
oheimb@11351
    15
oheimb@11351
    16
text {*
wenzelm@11355
    17
  We extend the standard natural numbers by a special value indicating
wenzelm@11355
    18
  infinity.  This includes extending the ordering relations @{term "op
wenzelm@11355
    19
  <"} and @{term "op \<le>"}.
oheimb@11351
    20
*}
oheimb@11351
    21
oheimb@11351
    22
datatype inat = Fin nat | Infty
oheimb@11351
    23
oheimb@11351
    24
instance inat :: ord ..
oheimb@11351
    25
instance inat :: zero ..
oheimb@11351
    26
oheimb@11351
    27
consts
wenzelm@11355
    28
  iSuc :: "inat => inat"
oheimb@11351
    29
oheimb@11351
    30
syntax (xsymbols)
wenzelm@11355
    31
  Infty :: inat    ("\<infinity>")
oheimb@11351
    32
oheimb@11351
    33
defs
wenzelm@11701
    34
  Zero_inat_def: "0 == Fin 0"
wenzelm@11355
    35
  iSuc_def: "iSuc i == case i of Fin n  => Fin (Suc n) | \<infinity> => \<infinity>"
wenzelm@11355
    36
  iless_def: "m < n ==
wenzelm@11355
    37
    case m of Fin m1 => (case n of Fin n1 => m1 < n1 | \<infinity> => True)
wenzelm@11355
    38
    | \<infinity>  => False"
wenzelm@11355
    39
  ile_def: "(m::inat) \<le> n == \<not> (n < m)"
oheimb@11351
    40
wenzelm@11701
    41
lemmas inat_defs = Zero_inat_def iSuc_def iless_def ile_def
oheimb@11351
    42
lemmas inat_splits = inat.split inat.split_asm
oheimb@11351
    43
wenzelm@11355
    44
text {*
wenzelm@11357
    45
  Below is a not quite complete set of theorems.  Use the method
wenzelm@11357
    46
  @{text "(simp add: inat_defs split:inat_splits, arith?)"} to prove
wenzelm@11357
    47
  new theorems or solve arithmetic subgoals involving @{typ inat} on
wenzelm@11357
    48
  the fly.
oheimb@11351
    49
*}
oheimb@11351
    50
oheimb@11351
    51
subsection "Constructors"
oheimb@11351
    52
oheimb@11351
    53
lemma Fin_0: "Fin 0 = 0"
wenzelm@11357
    54
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    55
oheimb@11351
    56
lemma Infty_ne_i0 [simp]: "\<infinity> \<noteq> 0"
wenzelm@11357
    57
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    58
oheimb@11351
    59
lemma i0_ne_Infty [simp]: "0 \<noteq> \<infinity>"
wenzelm@11357
    60
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    61
oheimb@11351
    62
lemma iSuc_Fin [simp]: "iSuc (Fin n) = Fin (Suc n)"
wenzelm@11357
    63
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    64
oheimb@11351
    65
lemma iSuc_Infty [simp]: "iSuc \<infinity> = \<infinity>"
wenzelm@11357
    66
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    67
oheimb@11351
    68
lemma iSuc_ne_0 [simp]: "iSuc n \<noteq> 0"
wenzelm@11357
    69
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    70
oheimb@11351
    71
lemma iSuc_inject [simp]: "(iSuc x = iSuc y) = (x = y)"
wenzelm@11357
    72
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    73
oheimb@11351
    74
oheimb@11351
    75
subsection "Ordering relations"
oheimb@11351
    76
oheimb@11351
    77
lemma Infty_ilessE [elim!]: "\<infinity> < Fin m ==> R"
wenzelm@11357
    78
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    79
wenzelm@11355
    80
lemma iless_linear: "m < n \<or> m = n \<or> n < (m::inat)"
wenzelm@11357
    81
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    82
oheimb@11351
    83
lemma iless_not_refl [simp]: "\<not> n < (n::inat)"
wenzelm@11357
    84
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    85
oheimb@11351
    86
lemma iless_trans: "i < j ==> j < k ==> i < (k::inat)"
wenzelm@11357
    87
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    88
oheimb@11351
    89
lemma iless_not_sym: "n < m ==> \<not> m < (n::inat)"
wenzelm@11357
    90
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    91
oheimb@11351
    92
lemma Fin_iless_mono [simp]: "(Fin n < Fin m) = (n < m)"
wenzelm@11357
    93
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    94
oheimb@11351
    95
lemma Fin_iless_Infty [simp]: "Fin n < \<infinity>"
wenzelm@11357
    96
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
    97
wenzelm@11655
    98
lemma Infty_eq [simp]: "(n < \<infinity>) = (n \<noteq> \<infinity>)"
wenzelm@11357
    99
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   100
oheimb@11351
   101
lemma i0_eq [simp]: "((0::inat) < n) = (n \<noteq> 0)"
wenzelm@11357
   102
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   103
oheimb@11351
   104
lemma i0_iless_iSuc [simp]: "0 < iSuc n"
wenzelm@11357
   105
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   106
oheimb@11351
   107
lemma not_ilessi0 [simp]: "\<not> n < (0::inat)"
wenzelm@11357
   108
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   109
oheimb@11351
   110
lemma Fin_iless: "n < Fin m ==> \<exists>k. n = Fin k"
wenzelm@11357
   111
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   112
wenzelm@11655
   113
lemma iSuc_mono [simp]: "(iSuc n < iSuc m) = (n < m)"
wenzelm@11357
   114
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   115
oheimb@11351
   116
oheimb@11351
   117
(* ----------------------------------------------------------------------- *)
oheimb@11351
   118
wenzelm@11655
   119
lemma ile_def2: "(m \<le> n) = (m < n \<or> m = (n::inat))"
wenzelm@11357
   120
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   121
wenzelm@11355
   122
lemma ile_refl [simp]: "n \<le> (n::inat)"
wenzelm@11357
   123
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   124
wenzelm@11355
   125
lemma ile_trans: "i \<le> j ==> j \<le> k ==> i \<le> (k::inat)"
wenzelm@11357
   126
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   127
wenzelm@11355
   128
lemma ile_iless_trans: "i \<le> j ==> j < k ==> i < (k::inat)"
wenzelm@11357
   129
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   130
wenzelm@11355
   131
lemma iless_ile_trans: "i < j ==> j \<le> k ==> i < (k::inat)"
wenzelm@11357
   132
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   133
wenzelm@11355
   134
lemma Infty_ub [simp]: "n \<le> \<infinity>"
wenzelm@11357
   135
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   136
wenzelm@11355
   137
lemma i0_lb [simp]: "(0::inat) \<le> n"
wenzelm@11357
   138
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   139
wenzelm@11355
   140
lemma Infty_ileE [elim!]: "\<infinity> \<le> Fin m ==> R"
wenzelm@11357
   141
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   142
wenzelm@11355
   143
lemma Fin_ile_mono [simp]: "(Fin n \<le> Fin m) = (n \<le> m)"
wenzelm@11357
   144
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   145
wenzelm@11355
   146
lemma ilessI1: "n \<le> m ==> n \<noteq> m ==> n < (m::inat)"
wenzelm@11357
   147
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   148
wenzelm@11355
   149
lemma ileI1: "m < n ==> iSuc m \<le> n"
wenzelm@11357
   150
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   151
wenzelm@11655
   152
lemma Suc_ile_eq: "(Fin (Suc m) \<le> n) = (Fin m < n)"
wenzelm@11357
   153
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   154
wenzelm@11655
   155
lemma iSuc_ile_mono [simp]: "(iSuc n \<le> iSuc m) = (n \<le> m)"
wenzelm@11357
   156
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   157
wenzelm@11655
   158
lemma iless_Suc_eq [simp]: "(Fin m < iSuc n) = (Fin m \<le> n)"
wenzelm@11357
   159
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   160
wenzelm@11355
   161
lemma not_iSuc_ilei0 [simp]: "\<not> iSuc n \<le> 0"
wenzelm@11357
   162
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   163
wenzelm@11355
   164
lemma ile_iSuc [simp]: "n \<le> iSuc n"
wenzelm@11357
   165
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   166
wenzelm@11355
   167
lemma Fin_ile: "n \<le> Fin m ==> \<exists>k. n = Fin k"
wenzelm@11357
   168
  by (simp add: inat_defs split:inat_splits, arith?)
oheimb@11351
   169
oheimb@11351
   170
lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. Fin k < Y j"
wenzelm@11355
   171
  apply (induct_tac k)
wenzelm@11355
   172
   apply (simp (no_asm) only: Fin_0)
wenzelm@11355
   173
   apply (fast intro: ile_iless_trans i0_lb)
wenzelm@11355
   174
  apply (erule exE)
wenzelm@11355
   175
  apply (drule spec)
wenzelm@11355
   176
  apply (erule exE)
wenzelm@11355
   177
  apply (drule ileI1)
wenzelm@11355
   178
  apply (rule iSuc_Fin [THEN subst])
wenzelm@11355
   179
  apply (rule exI)
wenzelm@11355
   180
  apply (erule (1) ile_iless_trans)
wenzelm@11355
   181
  done
oheimb@11351
   182
oheimb@11351
   183
end