src/HOL/Library/Extended.thy
author nipkow
Tue, 05 Mar 2013 15:26:57 +0100
changeset 51338 054d1653950f
child 51357 ac4c1cf1b001
permissions -rw-r--r--
New theory of infinity-extended types; should replace Extended_xyz eventually
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51338
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
     1
(*  Author:     Tobias Nipkow, TU München
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
     2
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
     3
A theory of types extended with a greatest and a least element.
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
     4
Oriented towards numeric types, hence "\<infinity>" and "-\<infinity>".
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
     5
*)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
     6
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
     7
theory Extended
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
     8
imports Main
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
     9
begin
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    10
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    11
datatype 'a extended = Fin 'a | Pinf ("\<infinity>") | Minf ("-\<infinity>")
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    12
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    13
lemmas extended_cases2 = extended.exhaust[case_product extended.exhaust]
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    14
lemmas extended_cases3 = extended.exhaust[case_product extended_cases2]
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    15
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    16
instantiation extended :: (order)order
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    17
begin
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    18
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    19
fun less_eq_extended :: "'a extended \<Rightarrow> 'a extended \<Rightarrow> bool" where
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    20
"Fin x \<le> Fin y = (x \<le> y)" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    21
"_     \<le> Pinf  = True" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    22
"Minf  \<le> _     = True" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    23
"(_::'a extended) \<le> _     = False"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    24
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    25
lemma less_eq_extended_cases:
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    26
  "x \<le> y = (case x of Fin x \<Rightarrow> (case y of Fin y \<Rightarrow> x \<le> y | Pinf \<Rightarrow> True | Minf \<Rightarrow> False)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    27
            | Pinf \<Rightarrow> y=Pinf | Minf \<Rightarrow> True)"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    28
by(induct x y rule: less_eq_extended.induct)(auto split: extended.split)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    29
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    30
definition less_extended :: "'a extended \<Rightarrow> 'a extended \<Rightarrow> bool" where
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    31
"((x::'a extended) < y) = (x \<le> y & \<not> x \<ge> y)"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    32
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    33
instance
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    34
proof
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    35
  case goal1 show ?case by(rule less_extended_def)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    36
next
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    37
  case goal2 show ?case by(cases x) auto
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    38
next
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    39
  case goal3 thus ?case by(auto simp: less_eq_extended_cases split:extended.splits)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    40
next
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    41
  case goal4 thus ?case by(auto simp: less_eq_extended_cases split:extended.splits)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    42
qed
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    43
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    44
end
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    45
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    46
instance extended :: (linorder)linorder
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    47
proof
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    48
  case goal1 thus ?case by(auto simp: less_eq_extended_cases split:extended.splits)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    49
qed
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    50
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    51
lemma Minf_le[simp]: "Minf \<le> y"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    52
by(cases y) auto
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    53
lemma le_Pinf[simp]: "x \<le> Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    54
by(cases x) auto
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    55
lemma le_Minf[simp]: "x \<le> Minf \<longleftrightarrow> x = Minf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    56
by(cases x) auto
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    57
lemma Pinf_le[simp]: "Pinf \<le> x \<longleftrightarrow> x = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    58
by(cases x) auto
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    59
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    60
lemma less_extended_simps[simp]:
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    61
  "Fin x < Fin y = (x < y)"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    62
  "Fin x < Pinf  = True"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    63
  "Fin x < Minf  = False"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    64
  "Pinf < h      = False"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    65
  "Minf < Fin x  = True"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    66
  "Minf < Pinf   = True"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    67
  "l    < Minf   = False"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    68
by (auto simp add: less_extended_def)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    69
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    70
lemma min_extended_simps[simp]:
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    71
  "min (Fin x) (Fin y) = Fin(min x y)"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    72
  "min xx      Pinf    = xx"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    73
  "min xx      Minf    = Minf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    74
  "min Pinf    yy      = yy"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    75
  "min Minf    yy      = Minf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    76
by (auto simp add: min_def)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    77
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    78
lemma max_extended_simps[simp]:
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    79
  "max (Fin x) (Fin y) = Fin(max x y)"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    80
  "max xx      Pinf    = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    81
  "max xx      Minf    = xx"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    82
  "max Pinf    yy      = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    83
  "max Minf    yy      = yy"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    84
by (auto simp add: max_def)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    85
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    86
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    87
instantiation extended :: (plus)plus
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    88
begin
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    89
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    90
text {* The following definition of of addition is totalized
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    91
to make it asociative and commutative. Normally the sum of plus and minus infinity is undefined. *}
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    92
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    93
fun plus_extended where
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    94
"Fin x + Fin y = Fin(x+y)" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    95
"Fin x + Pinf  = Pinf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    96
"Pinf  + Fin x = Pinf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    97
"Pinf  + Pinf  = Pinf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    98
"Minf  + Fin y = Minf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
    99
"Fin x + Minf  = Minf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   100
"Minf  + Minf  = Minf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   101
"Minf  + Pinf  = Pinf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   102
"Pinf  + Minf  = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   103
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   104
instance ..
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   105
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   106
end
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   107
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   108
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   109
instance extended :: (ab_semigroup_add)ab_semigroup_add
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   110
proof
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   111
  fix a b c :: "'a extended"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   112
  show "a + b = b + a"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   113
    by (induct a b rule: plus_extended.induct) (simp_all add: ac_simps)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   114
  show "a + b + c = a + (b + c)"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   115
    by (cases rule: extended_cases3[of a b c]) (simp_all add: ac_simps)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   116
qed
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   117
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   118
instance extended :: (ordered_ab_semigroup_add)ordered_ab_semigroup_add
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   119
proof
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   120
  fix a b c :: "'a extended"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   121
  assume "a \<le> b" then show "c + a \<le> c + b"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   122
    by (cases rule: extended_cases3[of a b c]) (auto simp: add_left_mono)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   123
qed
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   124
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   125
instantiation extended :: (comm_monoid_add)comm_monoid_add
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   126
begin
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   127
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   128
definition "0 = Fin 0"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   129
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   130
instance
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   131
proof
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   132
  fix x :: "'a extended" show "0 + x = x" unfolding zero_extended_def by(cases x)auto
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   133
qed
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   134
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   135
end
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   136
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   137
instantiation extended :: (uminus)uminus
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   138
begin
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   139
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   140
fun uminus_extended where
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   141
"- (Fin x) = Fin (- x)" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   142
"- Pinf    = Minf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   143
"- Minf    = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   144
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   145
instance ..
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   146
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   147
end
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   148
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   149
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   150
instantiation extended :: (ab_group_add)minus
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   151
begin
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   152
definition "x - y = x + -(y::'a extended)"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   153
instance ..
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   154
end
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   155
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   156
lemma minus_extended_simps[simp]:
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   157
  "Fin x - Fin y = Fin(x - y)"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   158
  "Fin x - Pinf  = Minf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   159
  "Fin x - Minf  = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   160
  "Pinf  - Fin y = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   161
  "Pinf  - Minf  = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   162
  "Minf  - Fin y = Minf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   163
  "Minf  - Pinf  = Minf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   164
  "Minf  - Minf  = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   165
  "Pinf  - Pinf  = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   166
by (simp_all add: minus_extended_def)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   167
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   168
instantiation extended :: (lattice)bounded_lattice
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   169
begin
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   170
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   171
definition "bot = Minf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   172
definition "top = Pinf"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   173
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   174
fun inf_extended :: "'a extended \<Rightarrow> 'a extended \<Rightarrow> 'a extended" where
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   175
"inf_extended (Fin i) (Fin j) = Fin (inf i j)" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   176
"inf_extended a Minf = Minf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   177
"inf_extended Minf a = Minf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   178
"inf_extended Pinf a = a" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   179
"inf_extended a Pinf = a"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   180
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   181
fun sup_extended :: "'a extended \<Rightarrow> 'a extended \<Rightarrow> 'a extended" where
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   182
"sup_extended (Fin i) (Fin j) = Fin (sup i j)" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   183
"sup_extended a Pinf = Pinf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   184
"sup_extended Pinf a = Pinf" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   185
"sup_extended Minf a = a" |
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   186
"sup_extended a Minf = a"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   187
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   188
instance
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   189
proof
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   190
  fix x y z ::"'a extended"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   191
  show "inf x y \<le> x" "inf x y \<le> y" "\<lbrakk>x \<le> y; x \<le> z\<rbrakk> \<Longrightarrow> x \<le> inf y z"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   192
    "x \<le> sup x y" "y \<le> sup x y" "\<lbrakk>y \<le> x; z \<le> x\<rbrakk> \<Longrightarrow> sup y z \<le> x" "bot \<le> x" "x \<le> top"
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   193
    apply (atomize (full))
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   194
    apply (cases rule: extended_cases3[of x y z])
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   195
    apply (auto simp: bot_extended_def top_extended_def)
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   196
    done
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   197
qed
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   198
end
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   199
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   200
end
054d1653950f New theory of infinity-extended types; should replace Extended_xyz eventually
nipkow
parents:
diff changeset
   201