src/HOL/Library/Option_ord.thy
author blanchet
Thu, 05 Mar 2015 12:32:11 +0100
changeset 59607 a93592aedce4
parent 58881 b9556a055632
child 60500 903bb1495239
permissions -rw-r--r--
message tuning
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/Option_ord.thy
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
     3
*)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
     4
58881
b9556a055632 modernized header;
wenzelm
parents: 56166
diff changeset
     5
section {* Canonical order on option type *}
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
     6
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
     7
theory Option_ord
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
     8
imports Option Main
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
     9
begin
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    10
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    11
notation
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    12
  bot ("\<bottom>") and
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    13
  top ("\<top>") and
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    14
  inf  (infixl "\<sqinter>" 70) and
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    15
  sup  (infixl "\<squnion>" 65) and
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    16
  Inf  ("\<Sqinter>_" [900] 900) and
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    17
  Sup  ("\<Squnion>_" [900] 900)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    18
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    19
syntax (xsymbols)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    20
  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    21
  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    22
  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    23
  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    24
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    25
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    26
instantiation option :: (preorder) preorder
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    27
begin
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    28
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    29
definition less_eq_option where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 30662
diff changeset
    30
  "x \<le> y \<longleftrightarrow> (case x of None \<Rightarrow> True | Some x \<Rightarrow> (case y of None \<Rightarrow> False | Some y \<Rightarrow> x \<le> y))"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    31
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    32
definition less_option where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 30662
diff changeset
    33
  "x < y \<longleftrightarrow> (case y of None \<Rightarrow> False | Some y \<Rightarrow> (case x of None \<Rightarrow> True | Some x \<Rightarrow> x < y))"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    34
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    35
lemma less_eq_option_None [simp]: "None \<le> x"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    36
  by (simp add: less_eq_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    37
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    38
lemma less_eq_option_None_code [code]: "None \<le> x \<longleftrightarrow> True"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    39
  by simp
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    40
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    41
lemma less_eq_option_None_is_None: "x \<le> None \<Longrightarrow> x = None"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    42
  by (cases x) (simp_all add: less_eq_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    43
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    44
lemma less_eq_option_Some_None [simp, code]: "Some x \<le> None \<longleftrightarrow> False"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    45
  by (simp add: less_eq_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    46
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    47
lemma less_eq_option_Some [simp, code]: "Some x \<le> Some y \<longleftrightarrow> x \<le> y"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    48
  by (simp add: less_eq_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    49
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    50
lemma less_option_None [simp, code]: "x < None \<longleftrightarrow> False"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    51
  by (simp add: less_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    52
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    53
lemma less_option_None_is_Some: "None < x \<Longrightarrow> \<exists>z. x = Some z"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    54
  by (cases x) (simp_all add: less_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    55
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    56
lemma less_option_None_Some [simp]: "None < Some x"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    57
  by (simp add: less_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    58
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    59
lemma less_option_None_Some_code [code]: "None < Some x \<longleftrightarrow> True"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    60
  by simp
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    61
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    62
lemma less_option_Some [simp, code]: "Some x < Some y \<longleftrightarrow> x < y"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    63
  by (simp add: less_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    64
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    65
instance proof
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    66
qed (auto simp add: less_eq_option_def less_option_def less_le_not_le elim: order_trans split: option.splits)
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    67
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    68
end 
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    69
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    70
instance option :: (order) order proof
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    71
qed (auto simp add: less_eq_option_def less_option_def split: option.splits)
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    72
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    73
instance option :: (linorder) linorder proof
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    74
qed (auto simp add: less_eq_option_def less_option_def split: option.splits)
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    75
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
    76
instantiation option :: (order) order_bot
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    77
begin
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    78
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    79
definition bot_option where
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    80
  "\<bottom> = None"
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    81
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    82
instance proof
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    83
qed (simp add: bot_option_def)
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    84
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    85
end
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    86
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
    87
instantiation option :: (order_top) order_top
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    88
begin
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    89
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    90
definition top_option where
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    91
  "\<top> = Some \<top>"
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    92
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    93
instance proof
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    94
qed (simp add: top_option_def less_eq_option_def split: option.split)
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    95
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    96
end
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    97
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    98
instance option :: (wellorder) wellorder proof
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    99
  fix P :: "'a option \<Rightarrow> bool" and z :: "'a option"
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   100
  assume H: "\<And>x. (\<And>y. y < x \<Longrightarrow> P y) \<Longrightarrow> P x"
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   101
  have "P None" by (rule H) simp
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   102
  then have P_Some [case_names Some]:
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   103
    "\<And>z. (\<And>x. z = Some x \<Longrightarrow> (P o Some) x) \<Longrightarrow> P z"
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   104
  proof -
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   105
    fix z
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   106
    assume "\<And>x. z = Some x \<Longrightarrow> (P o Some) x"
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   107
    with `P None` show "P z" by (cases z) simp_all
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   108
  qed
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   109
  show "P z" proof (cases z rule: P_Some)
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   110
    case (Some w)
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   111
    show "(P o Some) w" proof (induct rule: less_induct)
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   112
      case (less x)
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   113
      have "P (Some x)" proof (rule H)
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   114
        fix y :: "'a option"
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   115
        assume "y < Some x"
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   116
        show "P y" proof (cases y rule: P_Some)
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   117
          case (Some v) with `y < Some x` have "v < x" by simp
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   118
          with less show "(P o Some) v" .
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   119
        qed
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   120
      qed
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   121
      then show ?case by simp
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   122
    qed
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   123
  qed
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   124
qed
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   125
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   126
instantiation option :: (inf) inf
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   127
begin
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   128
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   129
definition inf_option where
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   130
  "x \<sqinter> y = (case x of None \<Rightarrow> None | Some x \<Rightarrow> (case y of None \<Rightarrow> None | Some y \<Rightarrow> Some (x \<sqinter> y)))"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   131
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   132
lemma inf_None_1 [simp, code]:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   133
  "None \<sqinter> y = None"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   134
  by (simp add: inf_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   135
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   136
lemma inf_None_2 [simp, code]:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   137
  "x \<sqinter> None = None"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   138
  by (cases x) (simp_all add: inf_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   139
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   140
lemma inf_Some [simp, code]:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   141
  "Some x \<sqinter> Some y = Some (x \<sqinter> y)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   142
  by (simp add: inf_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   143
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   144
instance ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   145
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   146
end
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   147
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   148
instantiation option :: (sup) sup
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   149
begin
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   150
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   151
definition sup_option where
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   152
  "x \<squnion> y = (case x of None \<Rightarrow> y | Some x' \<Rightarrow> (case y of None \<Rightarrow> x | Some y \<Rightarrow> Some (x' \<squnion> y)))"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   153
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   154
lemma sup_None_1 [simp, code]:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   155
  "None \<squnion> y = y"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   156
  by (simp add: sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   157
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   158
lemma sup_None_2 [simp, code]:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   159
  "x \<squnion> None = x"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   160
  by (cases x) (simp_all add: sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   161
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   162
lemma sup_Some [simp, code]:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   163
  "Some x \<squnion> Some y = Some (x \<squnion> y)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   164
  by (simp add: sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   165
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   166
instance ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   167
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   168
end
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   169
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   170
instantiation option :: (semilattice_inf) semilattice_inf
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   171
begin
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   172
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   173
instance proof
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   174
  fix x y z :: "'a option"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   175
  show "x \<sqinter> y \<le> x"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   176
    by - (cases x, simp_all, cases y, simp_all)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   177
  show "x \<sqinter> y \<le> y"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   178
    by - (cases x, simp_all, cases y, simp_all)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   179
  show "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> y \<sqinter> z"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   180
    by - (cases x, simp_all, cases y, simp_all, cases z, simp_all)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   181
qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   182
  
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   183
end
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   184
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   185
instantiation option :: (semilattice_sup) semilattice_sup
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   186
begin
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   187
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   188
instance proof
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   189
  fix x y z :: "'a option"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   190
  show "x \<le> x \<squnion> y"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   191
    by - (cases x, simp_all, cases y, simp_all)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   192
  show "y \<le> x \<squnion> y"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   193
    by - (cases x, simp_all, cases y, simp_all)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   194
  fix x y z :: "'a option"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   195
  show "y \<le> x \<Longrightarrow> z \<le> x \<Longrightarrow> y \<squnion> z \<le> x"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   196
    by - (cases y, simp_all, cases z, simp_all, cases x, simp_all)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   197
qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   198
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   199
end
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   200
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   201
instance option :: (lattice) lattice ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   202
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   203
instance option :: (lattice) bounded_lattice_bot ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   204
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   205
instance option :: (bounded_lattice_top) bounded_lattice_top ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   206
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   207
instance option :: (bounded_lattice_top) bounded_lattice ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   208
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   209
instance option :: (distrib_lattice) distrib_lattice
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   210
proof
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   211
  fix x y z :: "'a option"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   212
  show "x \<squnion> y \<sqinter> z = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   213
    by - (cases x, simp_all, cases y, simp_all, cases z, simp_all add: sup_inf_distrib1 inf_commute)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   214
qed 
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   215
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   216
instantiation option :: (complete_lattice) complete_lattice
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   217
begin
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   218
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   219
definition Inf_option :: "'a option set \<Rightarrow> 'a option" where
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   220
  "\<Sqinter>A = (if None \<in> A then None else Some (\<Sqinter>Option.these A))"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   221
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   222
lemma None_in_Inf [simp]:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   223
  "None \<in> A \<Longrightarrow> \<Sqinter>A = None"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   224
  by (simp add: Inf_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   225
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   226
definition Sup_option :: "'a option set \<Rightarrow> 'a option" where
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   227
  "\<Squnion>A = (if A = {} \<or> A = {None} then None else Some (\<Squnion>Option.these A))"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   228
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   229
lemma empty_Sup [simp]:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   230
  "\<Squnion>{} = None"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   231
  by (simp add: Sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   232
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   233
lemma singleton_None_Sup [simp]:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   234
  "\<Squnion>{None} = None"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   235
  by (simp add: Sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   236
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   237
instance proof
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   238
  fix x :: "'a option" and A
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   239
  assume "x \<in> A"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   240
  then show "\<Sqinter>A \<le> x"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   241
    by (cases x) (auto simp add: Inf_option_def in_these_eq intro: Inf_lower)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   242
next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   243
  fix z :: "'a option" and A
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   244
  assume *: "\<And>x. x \<in> A \<Longrightarrow> z \<le> x"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   245
  show "z \<le> \<Sqinter>A"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   246
  proof (cases z)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   247
    case None then show ?thesis by simp
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   248
  next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   249
    case (Some y)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   250
    show ?thesis
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   251
      by (auto simp add: Inf_option_def in_these_eq Some intro!: Inf_greatest dest!: *)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   252
  qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   253
next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   254
  fix x :: "'a option" and A
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   255
  assume "x \<in> A"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   256
  then show "x \<le> \<Squnion>A"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   257
    by (cases x) (auto simp add: Sup_option_def in_these_eq intro: Sup_upper)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   258
next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   259
  fix z :: "'a option" and A
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   260
  assume *: "\<And>x. x \<in> A \<Longrightarrow> x \<le> z"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   261
  show "\<Squnion>A \<le> z "
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   262
  proof (cases z)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   263
    case None
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   264
    with * have "\<And>x. x \<in> A \<Longrightarrow> x = None" by (auto dest: less_eq_option_None_is_None)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   265
    then have "A = {} \<or> A = {None}" by blast
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   266
    then show ?thesis by (simp add: Sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   267
  next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   268
    case (Some y)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   269
    from * have "\<And>w. Some w \<in> A \<Longrightarrow> Some w \<le> z" .
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   270
    with Some have "\<And>w. w \<in> Option.these A \<Longrightarrow> w \<le> y"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   271
      by (simp add: in_these_eq)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   272
    then have "\<Squnion>Option.these A \<le> y" by (rule Sup_least)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   273
    with Some show ?thesis by (simp add: Sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   274
  qed
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
   275
next
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
   276
  show "\<Squnion>{} = (\<bottom>::'a option)"
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
   277
  by (auto simp: bot_option_def)
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
   278
next
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
   279
  show "\<Sqinter>{} = (\<top>::'a option)"
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
   280
  by (auto simp: top_option_def Inf_option_def)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   281
qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   282
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   283
end
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   284
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   285
lemma Some_Inf:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   286
  "Some (\<Sqinter>A) = \<Sqinter>(Some ` A)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   287
  by (auto simp add: Inf_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   288
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   289
lemma Some_Sup:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   290
  "A \<noteq> {} \<Longrightarrow> Some (\<Squnion>A) = \<Squnion>(Some ` A)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   291
  by (auto simp add: Sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   292
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   293
lemma Some_INF:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   294
  "Some (\<Sqinter>x\<in>A. f x) = (\<Sqinter>x\<in>A. Some (f x))"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 52729
diff changeset
   295
  using Some_Inf [of "f ` A"] by (simp add: comp_def)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   296
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   297
lemma Some_SUP:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   298
  "A \<noteq> {} \<Longrightarrow> Some (\<Squnion>x\<in>A. f x) = (\<Squnion>x\<in>A. Some (f x))"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 52729
diff changeset
   299
  using Some_Sup [of "f ` A"] by (simp add: comp_def)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   300
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   301
instantiation option :: (complete_distrib_lattice) complete_distrib_lattice
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   302
begin
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   303
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   304
instance proof
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   305
  fix a :: "'a option" and B
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   306
  show "a \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> b)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   307
  proof (cases a)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   308
    case None
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   309
    then show ?thesis by (simp add: INF_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   310
  next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   311
    case (Some c)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   312
    show ?thesis
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   313
    proof (cases "None \<in> B")
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   314
      case True
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   315
      then have "Some c = (\<Sqinter>b\<in>B. Some c \<squnion> b)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   316
        by (auto intro!: antisym INF_lower2 INF_greatest)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   317
      with True Some show ?thesis by simp
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   318
    next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   319
      case False then have B: "{x \<in> B. \<exists>y. x = Some y} = B" by auto (metis not_Some_eq)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   320
      from sup_Inf have "Some c \<squnion> Some (\<Sqinter>Option.these B) = Some (\<Sqinter>b\<in>Option.these B. c \<squnion> b)" by simp
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   321
      then have "Some c \<squnion> \<Sqinter>(Some ` Option.these B) = (\<Sqinter>x\<in>Some ` Option.these B. Some c \<squnion> x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 52729
diff changeset
   322
        by (simp add: Some_INF Some_Inf comp_def)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   323
      with Some B show ?thesis by (simp add: Some_image_these_eq)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   324
    qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   325
  qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   326
  show "a \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   327
  proof (cases a)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   328
    case None
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   329
    then show ?thesis by (simp add: SUP_def image_constant_conv bot_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   330
  next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   331
    case (Some c)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   332
    show ?thesis
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   333
    proof (cases "B = {} \<or> B = {None}")
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   334
      case True
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 52729
diff changeset
   335
      then show ?thesis by auto
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   336
    next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   337
      have B: "B = {x \<in> B. \<exists>y. x = Some y} \<union> {x \<in> B. x = None}"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   338
        by auto
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   339
      then have Sup_B: "\<Squnion>B = \<Squnion>({x \<in> B. \<exists>y. x = Some y} \<union> {x \<in> B. x = None})"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   340
        and SUP_B: "\<And>f. (\<Squnion>x \<in> B. f x) = (\<Squnion>x \<in> {x \<in> B. \<exists>y. x = Some y} \<union> {x \<in> B. x = None}. f x)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   341
        by simp_all
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   342
      have Sup_None: "\<Squnion>{x. x = None \<and> x \<in> B} = None"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   343
        by (simp add: bot_option_def [symmetric])
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   344
      have SUP_None: "(\<Squnion>x\<in>{x. x = None \<and> x \<in> B}. Some c \<sqinter> x) = None"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   345
        by (simp add: bot_option_def [symmetric])
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   346
      case False then have "Option.these B \<noteq> {}" by (simp add: these_not_empty_eq)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   347
      moreover from inf_Sup have "Some c \<sqinter> Some (\<Squnion>Option.these B) = Some (\<Squnion>b\<in>Option.these B. c \<sqinter> b)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   348
        by simp
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   349
      ultimately have "Some c \<sqinter> \<Squnion>(Some ` Option.these B) = (\<Squnion>x\<in>Some ` Option.these B. Some c \<sqinter> x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 52729
diff changeset
   350
        by (simp add: Some_SUP Some_Sup comp_def)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   351
      with Some show ?thesis
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   352
        by (simp add: Some_image_these_eq Sup_B SUP_B Sup_None SUP_None SUP_union Sup_union_distrib)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   353
    qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   354
  qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   355
qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   356
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   357
end
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   358
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   359
instantiation option :: (complete_linorder) complete_linorder
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   360
begin
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   361
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   362
instance ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   363
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   364
end
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   365
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   366
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   367
no_notation
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   368
  bot ("\<bottom>") and
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   369
  top ("\<top>") and
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   370
  inf  (infixl "\<sqinter>" 70) and
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   371
  sup  (infixl "\<squnion>" 65) and
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   372
  Inf  ("\<Sqinter>_" [900] 900) and
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   373
  Sup  ("\<Squnion>_" [900] 900)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   374
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   375
no_syntax (xsymbols)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   376
  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   377
  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   378
  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   379
  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   380
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   381
end
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   382