src/HOL/Library/Option_ord.thy
author wenzelm
Mon, 13 Jun 2022 11:10:39 +0200
changeset 75556 1f6fc2416a48
parent 74334 ead56ad40e15
child 81116 0fb1e2dd4122
permissions -rw-r--r--
clarified document structure; minor 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
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
     5
section \<open>Canonical order on option type\<close>
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
     6
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
     7
theory Option_ord
67006
b1278ed3cd46 prefer main entry points of HOL;
wenzelm
parents: 66453
diff changeset
     8
imports Main
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
     9
begin
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    10
74334
ead56ad40e15 bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
wenzelm
parents: 69861
diff changeset
    11
unbundle lattice_syntax
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
    12
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    13
instantiation option :: (preorder) preorder
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    14
begin
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    15
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    16
definition less_eq_option where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 30662
diff changeset
    17
  "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
    18
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    19
definition less_option where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 30662
diff changeset
    20
  "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
    21
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    22
lemma less_eq_option_None [simp]: "None \<le> x"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    23
  by (simp add: less_eq_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    24
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    25
lemma less_eq_option_None_code [code]: "None \<le> x \<longleftrightarrow> True"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    26
  by simp
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    27
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    28
lemma less_eq_option_None_is_None: "x \<le> None \<Longrightarrow> x = None"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    29
  by (cases x) (simp_all add: less_eq_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    30
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    31
lemma less_eq_option_Some_None [simp, code]: "Some x \<le> None \<longleftrightarrow> False"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    32
  by (simp add: less_eq_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    33
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    34
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
    35
  by (simp add: less_eq_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    36
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    37
lemma less_option_None [simp, code]: "x < None \<longleftrightarrow> False"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    38
  by (simp add: less_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    39
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    40
lemma less_option_None_is_Some: "None < x \<Longrightarrow> \<exists>z. x = Some z"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    41
  by (cases x) (simp_all add: less_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    42
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    43
lemma less_option_None_Some [simp]: "None < Some x"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    44
  by (simp add: less_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    45
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    46
lemma less_option_None_Some_code [code]: "None < Some x \<longleftrightarrow> True"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    47
  by simp
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    48
26258
20dfaa28e5e5 *** empty log message ***
haftmann
parents: 26241
diff changeset
    49
lemma less_option_Some [simp, code]: "Some x < Some y \<longleftrightarrow> x < y"
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    50
  by (simp add: less_option_def)
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    51
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    52
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    53
  by standard
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    54
    (auto simp add: less_eq_option_def less_option_def less_le_not_le
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    55
      elim: order_trans split: option.splits)
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    56
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    57
end
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    58
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    59
instance option :: (order) order
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    60
  by standard (auto simp add: less_eq_option_def less_option_def split: option.splits)
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    61
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    62
instance option :: (linorder) linorder
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    63
  by standard (auto simp add: less_eq_option_def less_option_def split: option.splits)
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    64
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
    65
instantiation option :: (order) order_bot
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    66
begin
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    67
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    68
definition bot_option where "\<bottom> = None"
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    69
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    70
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    71
  by standard (simp add: bot_option_def)
30662
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
end
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    74
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
    75
instantiation option :: (order_top) order_top
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    76
begin
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    77
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    78
definition top_option where "\<top> = Some \<top>"
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    79
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    80
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    81
  by standard (simp add: top_option_def less_eq_option_def split: option.split)
26241
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    82
b7d8c2338585 added Option_ord.thy
haftmann
parents:
diff changeset
    83
end
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    84
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    85
instance option :: (wellorder) wellorder
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    86
proof
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    87
  fix P :: "'a option \<Rightarrow> bool"
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    88
  fix z :: "'a option"
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    89
  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
    90
  have "P None" by (rule H) simp
67091
1393c2340eec more symbols;
wenzelm
parents: 67006
diff changeset
    91
  then have P_Some [case_names Some]: "P z" if "\<And>x. z = Some x \<Longrightarrow> (P \<circ> Some) x" for z
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    92
    using \<open>P None\<close> that by (cases z) simp_all
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    93
  show "P z"
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    94
  proof (cases z rule: P_Some)
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    95
    case (Some w)
67091
1393c2340eec more symbols;
wenzelm
parents: 67006
diff changeset
    96
    show "(P \<circ> Some) w"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    97
    proof (induct rule: less_induct)
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
    98
      case (less x)
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    99
      have "P (Some x)"
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   100
      proof (rule H)
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   101
        fix y :: "'a option"
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   102
        assume "y < Some x"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   103
        show "P y"
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   104
        proof (cases y rule: P_Some)
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   105
          case (Some v)
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   106
          with \<open>y < Some x\<close> have "v < x" by simp
67091
1393c2340eec more symbols;
wenzelm
parents: 67006
diff changeset
   107
          with less show "(P \<circ> Some) v" .
30662
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
      qed
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   110
      then show ?case by simp
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   111
    qed
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   112
  qed
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   113
qed
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   114
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   115
instantiation option :: (inf) inf
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   116
begin
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   117
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   118
definition inf_option where
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   119
  "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
   120
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   121
lemma inf_None_1 [simp, code]: "None \<sqinter> y = None"
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   122
  by (simp add: inf_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   123
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   124
lemma inf_None_2 [simp, code]: "x \<sqinter> None = None"
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   125
  by (cases x) (simp_all add: inf_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   126
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   127
lemma inf_Some [simp, code]: "Some x \<sqinter> Some y = Some (x \<sqinter> y)"
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   128
  by (simp add: inf_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   129
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   130
instance ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   131
30662
396be15b95d5 added instances for bot, top, wellorder
haftmann
parents: 28562
diff changeset
   132
end
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   133
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   134
instantiation option :: (sup) sup
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   135
begin
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   136
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   137
definition sup_option where
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   138
  "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
   139
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   140
lemma sup_None_1 [simp, code]: "None \<squnion> y = y"
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   141
  by (simp add: sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   142
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   143
lemma sup_None_2 [simp, code]: "x \<squnion> None = x"
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   144
  by (cases x) (simp_all add: sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   145
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   146
lemma sup_Some [simp, code]: "Some x \<squnion> Some y = Some (x \<squnion> y)"
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   147
  by (simp add: sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   148
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   149
instance ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   150
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   151
end
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   152
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   153
instance option :: (semilattice_inf) semilattice_inf
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   154
proof
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   155
  fix x y z :: "'a option"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   156
  show "x \<sqinter> y \<le> x"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   157
    by (cases x, simp_all, cases y, simp_all)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   158
  show "x \<sqinter> y \<le> y"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   159
    by (cases x, simp_all, cases y, simp_all)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   160
  show "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> y \<sqinter> z"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   161
    by (cases x, simp_all, cases y, simp_all, cases z, simp_all)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   162
qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   163
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   164
instance option :: (semilattice_sup) semilattice_sup
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   165
proof
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   166
  fix x y z :: "'a option"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   167
  show "x \<le> x \<squnion> y"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   168
    by (cases x, simp_all, cases y, simp_all)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   169
  show "y \<le> x \<squnion> y"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   170
    by (cases x, simp_all, cases y, simp_all)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   171
  fix x y z :: "'a option"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   172
  show "y \<le> x \<Longrightarrow> z \<le> x \<Longrightarrow> y \<squnion> z \<le> x"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   173
    by (cases y, simp_all, cases z, simp_all, cases x, simp_all)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   174
qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   175
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   176
instance option :: (lattice) lattice ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   177
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   178
instance option :: (lattice) bounded_lattice_bot ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   179
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   180
instance option :: (bounded_lattice_top) bounded_lattice_top ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   181
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   182
instance option :: (bounded_lattice_top) bounded_lattice ..
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   183
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   184
instance option :: (distrib_lattice) distrib_lattice
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   185
proof
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   186
  fix x y z :: "'a option"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   187
  show "x \<squnion> y \<sqinter> z = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   188
    by (cases x, simp_all, cases y, simp_all, cases z, simp_all add: sup_inf_distrib1 inf_commute)
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   189
qed
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   191
instantiation option :: (complete_lattice) complete_lattice
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   192
begin
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   193
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   194
definition Inf_option :: "'a option set \<Rightarrow> 'a option" where
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   195
  "\<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
   196
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   197
lemma None_in_Inf [simp]: "None \<in> A \<Longrightarrow> \<Sqinter>A = None"
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   198
  by (simp add: Inf_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   199
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   200
definition Sup_option :: "'a option set \<Rightarrow> 'a option" where
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   201
  "\<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
   202
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   203
lemma empty_Sup [simp]: "\<Squnion>{} = None"
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   204
  by (simp add: Sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   205
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   206
lemma singleton_None_Sup [simp]: "\<Squnion>{None} = None"
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   207
  by (simp add: Sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   208
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   209
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   210
proof
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   211
  fix x :: "'a option" and A
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   212
  assume "x \<in> A"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   213
  then show "\<Sqinter>A \<le> x"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   214
    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
   215
next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   216
  fix z :: "'a option" and A
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   217
  assume *: "\<And>x. x \<in> A \<Longrightarrow> z \<le> x"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   218
  show "z \<le> \<Sqinter>A"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   219
  proof (cases z)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   220
    case None then show ?thesis by simp
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   221
  next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   222
    case (Some y)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   223
    show ?thesis
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   224
      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
   225
  qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   226
next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   227
  fix x :: "'a option" and A
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   228
  assume "x \<in> A"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   229
  then show "x \<le> \<Squnion>A"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   230
    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
   231
next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   232
  fix z :: "'a option" and A
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   233
  assume *: "\<And>x. x \<in> A \<Longrightarrow> x \<le> z"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   234
  show "\<Squnion>A \<le> z "
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   235
  proof (cases z)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   236
    case None
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   237
    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
   238
    then have "A = {} \<or> A = {None}" by blast
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   239
    then show ?thesis by (simp add: Sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   240
  next
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   241
    case (Some y)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   242
    from * have "\<And>w. Some w \<in> A \<Longrightarrow> Some w \<le> z" .
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   243
    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
   244
      by (simp add: in_these_eq)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   245
    then have "\<Squnion>Option.these A \<le> y" by (rule Sup_least)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   246
    with Some show ?thesis by (simp add: Sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   247
  qed
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
   248
next
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
   249
  show "\<Squnion>{} = (\<bottom>::'a option)"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   250
    by (auto simp: bot_option_def)
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 49190
diff changeset
   251
  show "\<Sqinter>{} = (\<top>::'a option)"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   252
    by (auto simp: top_option_def Inf_option_def)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   253
qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   254
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   255
end
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   256
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   257
lemma Some_Inf:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   258
  "Some (\<Sqinter>A) = \<Sqinter>(Some ` A)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   259
  by (auto simp add: Inf_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   260
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   261
lemma Some_Sup:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   262
  "A \<noteq> {} \<Longrightarrow> Some (\<Squnion>A) = \<Squnion>(Some ` A)"
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   263
  by (auto simp add: Sup_option_def)
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   264
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   265
lemma Some_INF:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   266
  "Some (\<Sqinter>x\<in>A. f x) = (\<Sqinter>x\<in>A. Some (f x))"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69661
diff changeset
   267
  by (simp add: Some_Inf image_comp)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   268
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   269
lemma Some_SUP:
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   270
  "A \<noteq> {} \<Longrightarrow> Some (\<Squnion>x\<in>A. f x) = (\<Squnion>x\<in>A. Some (f x))"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69661
diff changeset
   271
  by (simp add: Some_Sup image_comp)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   272
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 68980
diff changeset
   273
lemma option_Inf_Sup: "\<Sqinter>(Sup ` A) \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
b021008c5397 removed legacy input syntax
haftmann
parents: 68980
diff changeset
   274
  for A :: "('a::complete_distrib_lattice option) set set"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   275
proof (cases "{} \<in> A")
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   276
  case True
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   277
  then show ?thesis
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   278
    by (rule INF_lower2, simp_all)
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   279
next
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   280
  case False
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   281
  from this have X: "{} \<notin> A"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   282
    by simp
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   283
  then show ?thesis
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   284
  proof (cases "{None} \<in> A")
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   285
    case True
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   286
    then show ?thesis
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   287
      by (rule INF_lower2, simp_all)
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   288
  next
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   289
    case False
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   290
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   291
    {fix y
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   292
      assume A: "y \<in> A"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   293
      have "Sup (y - {None}) = Sup y"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   294
        by (metis (no_types, lifting) Sup_option_def insert_Diff_single these_insert_None these_not_empty_eq)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   295
      from A and this have "(\<exists>z. y - {None} = z - {None} \<and> z \<in> A) \<and> \<Squnion>y = \<Squnion>(y - {None})"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   296
        by auto
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   297
    }
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   298
    from this have A: "Sup ` A = (Sup ` {y - {None} | y. y\<in>A})"
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   299
      by (auto simp add: image_def)
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   300
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   301
    have [simp]: "\<And>y. y \<in> A \<Longrightarrow> \<exists>ya. {ya. \<exists>x. x \<in> y \<and> (\<exists>y. x = Some y) \<and> ya = the x} 
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   302
          = {y. \<exists>x\<in>ya - {None}. y = the x} \<and> ya \<in> A"
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   303
      by (rule exI, auto)
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   304
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   305
    have [simp]: "\<And>y. y \<in> A \<Longrightarrow>
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   306
         (\<exists>ya. y - {None} = ya - {None} \<and> ya \<in> A) \<and> \<Squnion>{ya. \<exists>x\<in>y - {None}. ya = the x} 
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   307
          = \<Squnion>{ya. \<exists>x. x \<in> y \<and> (\<exists>y. x = Some y) \<and> ya = the x}"
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   308
      apply (safe, blast)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   309
      by (rule arg_cong [of _ _ Sup], auto)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   310
    {fix y
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   311
      assume [simp]: "y \<in> A"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   312
      have "\<exists>x. (\<exists>y. x = {ya. \<exists>x\<in>y - {None}. ya = the x} \<and> y \<in> A) \<and> \<Squnion>{ya. \<exists>x. x \<in> y \<and> (\<exists>y. x = Some y) \<and> ya = the x} = \<Squnion>x"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   313
      and "\<exists>x. (\<exists>y. x = y - {None} \<and> y \<in> A) \<and> \<Squnion>{ya. \<exists>x\<in>y - {None}. ya = the x} = \<Squnion>{y. \<exists>xa. xa \<in> x \<and> (\<exists>y. xa = Some y) \<and> y = the xa}"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   314
         apply (rule exI [of _ "{ya. \<exists>x. x \<in> y \<and> (\<exists>y. x = Some y) \<and> ya = the x}"], simp)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   315
        by (rule exI [of _ "y - {None}"], simp)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   316
    }
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   317
    from this have C: "(\<lambda>x.  (\<Squnion>Option.these x)) ` {y - {None} |y. y \<in> A} =  (Sup ` {the ` (y - {None}) |y. y \<in> A})"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   318
      by (simp add: image_def Option.these_def, safe, simp_all)
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   319
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   320
    have D: "\<forall> f . \<exists>Y\<in>A. f Y \<notin> Y \<Longrightarrow> False"
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   321
      by (drule spec [of _ "\<lambda> Y . SOME x . x \<in> Y"], simp add: X some_in_eq)
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   322
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   323
    define F where "F = (\<lambda> Y . SOME x::'a option . x \<in> (Y - {None}))"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   324
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   325
    have G: "\<And> Y . Y \<in> A \<Longrightarrow> \<exists> x . x \<in> Y - {None}"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   326
      by (metis False X all_not_in_conv insert_Diff_single these_insert_None these_not_empty_eq)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   327
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   328
    have F: "\<And> Y . Y \<in> A \<Longrightarrow> F Y \<in> (Y - {None})"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   329
      by (metis F_def G empty_iff some_in_eq)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   330
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   331
    have "Some \<bottom> \<le> Inf (F ` A)"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   332
      by (metis (no_types, lifting) Diff_iff F Inf_option_def bot.extremum image_iff 
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   333
          less_eq_option_Some singletonI)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   334
      
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   335
    from this have "Inf (F ` A) \<noteq> None"
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   336
      by (cases "\<Sqinter>x\<in>A. F x", simp_all)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   337
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   338
    from this have "Inf (F ` A) \<noteq> None \<and> Inf (F ` A) \<in> Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   339
      using F by auto
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   340
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   341
    from this have "\<exists> x . x \<noteq> None \<and> x \<in> Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}"
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   342
      by blast
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   343
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   344
    from this have E:" Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y} = {None} \<Longrightarrow> False"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   345
      by blast
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   346
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   347
    have [simp]: "((\<Squnion>x\<in>{f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}. \<Sqinter>x) = None) = False"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   348
      by (metis (no_types, lifting) E Sup_option_def \<open>\<exists>x. x \<noteq> None \<and> x \<in> Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}\<close> 
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   349
          ex_in_conv option.simps(3))
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   350
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   351
    have B: "Option.these ((\<lambda>x. Some (\<Squnion>Option.these x)) ` {y - {None} |y. y \<in> A}) 
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   352
        = ((\<lambda>x. (\<Squnion> Option.these x)) ` {y - {None} |y. y \<in> A})"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   353
      by (metis image_image these_image_Some_eq)
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   354
    {
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   355
      fix f
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   356
      assume A: "\<And> Y . (\<exists>y. Y = the ` (y - {None}) \<and> y \<in> A) \<Longrightarrow> f Y \<in> Y"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   357
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   358
      have "\<And>xa. xa \<in> A \<Longrightarrow> f {y. \<exists>a\<in>xa - {None}. y = the a} = f (the ` (xa - {None}))"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   359
        by  (simp add: image_def)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   360
      from this have [simp]: "\<And>xa. xa \<in> A \<Longrightarrow> \<exists>x\<in>A. f {y. \<exists>a\<in>xa - {None}. y = the a} = f (the ` (x - {None}))"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   361
        by blast
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   362
      have "\<And>xa. xa \<in> A \<Longrightarrow> f (the ` (xa - {None})) = f {y. \<exists>a \<in> xa - {None}. y = the a} \<and> xa \<in> A"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   363
        by (simp add: image_def)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   364
      from this have [simp]: "\<And>xa. xa \<in> A \<Longrightarrow> \<exists>x. f (the ` (xa - {None})) = f {y. \<exists>a\<in>x - {None}. y = the a} \<and> x \<in> A"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   365
        by blast
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   366
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   367
      {
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   368
        fix Y
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   369
        have "Y \<in> A \<Longrightarrow> Some (f (the ` (Y - {None}))) \<in> Y"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   370
          using A [of "the ` (Y - {None})"] apply (simp add: image_def)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   371
          using option.collapse by fastforce
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   372
      }
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   373
      from this have [simp]: "\<And> Y . Y \<in> A \<Longrightarrow> Some (f (the ` (Y - {None}))) \<in> Y"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   374
        by blast
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   375
      have [simp]: "(\<Sqinter>x\<in>A. Some (f {y. \<exists>x\<in>x - {None}. y = the x})) = \<Sqinter>{Some (f {y. \<exists>a\<in>x - {None}. y = the a}) |x. x \<in> A}"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   376
        by (simp add: Setcompr_eq_image)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   377
      
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   378
      have [simp]: "\<exists>x. (\<exists>f. x = {y. \<exists>x\<in>A. y = f x} \<and> (\<forall>Y\<in>A. f Y \<in> Y)) \<and> \<Sqinter>{Some (f {y. \<exists>a\<in>x - {None}. y = the a}) |x. x \<in> A} = \<Sqinter>x"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   379
        apply (rule exI [of _ "{Some (f {y. \<exists>a\<in>x - {None}. y = the a}) | x . x\<in> A}"], safe)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   380
        by (rule exI [of _ "(\<lambda> Y . Some (f (the ` (Y - {None})))) "], safe, simp_all)
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   381
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   382
      {
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   383
        fix xb
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   384
        have "xb \<in> A \<Longrightarrow> (\<Sqinter>x\<in>{{ya. \<exists>x\<in>y - {None}. ya = the x} |y. y \<in> A}. f x) \<le> f {y. \<exists>x\<in>xb - {None}. y = the x}"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   385
          apply (rule INF_lower2 [of "{y. \<exists>x\<in>xb - {None}. y = the x}"])
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   386
          by blast+
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   387
      }
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   388
      from this have [simp]: "(\<Sqinter>x\<in>{the ` (y - {None}) |y. y \<in> A}. f x) \<le> the (\<Sqinter>Y\<in>A. Some (f (the ` (Y - {None}))))"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   389
        apply (simp add: Inf_option_def image_def Option.these_def)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   390
        by (rule Inf_greatest, clarsimp)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   391
      have [simp]: "the (\<Sqinter>Y\<in>A. Some (f (the ` (Y - {None})))) \<in> Option.these (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   392
        apply (auto simp add: Option.these_def)
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   393
        apply (rule imageI)
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   394
        apply auto
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   395
        using \<open>\<And>Y. Y \<in> A \<Longrightarrow> Some (f (the ` (Y - {None}))) \<in> Y\<close> apply blast
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   396
        apply (auto simp add: Some_INF [symmetric])
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   397
        done
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   398
      have "(\<Sqinter>x\<in>{the ` (y - {None}) |y. y \<in> A}. f x) \<le> \<Squnion>Option.these (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   399
        by (rule Sup_upper2 [of "the (Inf ((\<lambda> Y . Some (f (the ` (Y - {None})) )) ` A))"], simp_all)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   400
    }
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   401
    from this have X: "\<And> f . \<forall>Y. (\<exists>y. Y = the ` (y - {None}) \<and> y \<in> A) \<longrightarrow> f Y \<in> Y \<Longrightarrow>
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   402
      (\<Sqinter>x\<in>{the ` (y - {None}) |y. y \<in> A}. f x) \<le> \<Squnion>Option.these (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   403
      by blast
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   404
    
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   405
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   406
    have [simp]: "\<And> x . x\<in>{y - {None} |y. y \<in> A} \<Longrightarrow>  x \<noteq> {} \<and> x \<noteq> {None}"
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   407
      using F by fastforce
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   408
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   409
    have "(Inf (Sup `A)) = (Inf (Sup ` {y - {None} | y. y\<in>A}))"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   410
      by (subst A, simp)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   411
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   412
    also have "... = (\<Sqinter>x\<in>{y - {None} |y. y \<in> A}. if x = {} \<or> x = {None} then None else Some (\<Squnion>Option.these x))"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   413
      by (simp add: Sup_option_def)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   414
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   415
    also have "... = (\<Sqinter>x\<in>{y - {None} |y. y \<in> A}. Some (\<Squnion>Option.these x))"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   416
      using G by fastforce
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   417
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   418
    also have "... = Some (\<Sqinter>Option.these ((\<lambda>x. Some (\<Squnion>Option.these x)) ` {y - {None} |y. y \<in> A}))"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   419
      by (simp add: Inf_option_def, safe)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   420
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   421
    also have "... =  Some (\<Sqinter> ((\<lambda>x.  (\<Squnion>Option.these x)) ` {y - {None} |y. y \<in> A}))"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   422
      by (simp add: B)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   423
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   424
    also have "... = Some (Inf (Sup ` {the ` (y - {None}) |y. y \<in> A}))"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   425
      by (unfold C, simp)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   426
    thm Inf_Sup
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   427
    also have "... = Some (\<Squnion>x\<in>{f ` {the ` (y - {None}) |y. y \<in> A} |f. \<forall>Y. (\<exists>y. Y = the ` (y - {None}) \<and> y \<in> A) \<longrightarrow> f Y \<in> Y}. \<Sqinter>x) "
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   428
      by (simp add: Inf_Sup)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   429
  
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 68980
diff changeset
   430
    also have "... \<le> \<Squnion> (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
b021008c5397 removed legacy input syntax
haftmann
parents: 68980
diff changeset
   431
    proof (cases "\<Squnion> (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})")
67951
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   432
      case None
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   433
      then show ?thesis by (simp add: less_eq_option_def)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   434
    next
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   435
      case (Some a)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   436
      then show ?thesis
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   437
        apply simp
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   438
        apply (rule Sup_least, safe)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   439
        apply (simp add: Sup_option_def)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   440
        apply (cases "(\<forall>f. \<exists>Y\<in>A. f Y \<notin> Y) \<or> Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y} = {None}", simp_all)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   441
        by (drule X, simp)
655aa11359dc Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
Manuel Eberl <eberlm@in.tum.de>
parents: 67829
diff changeset
   442
    qed
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   443
    finally show ?thesis by simp
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   444
  qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   445
qed
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   446
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   447
instance option :: (complete_distrib_lattice) complete_distrib_lattice
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   448
  by (standard, simp add: option_Inf_Sup)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67091
diff changeset
   449
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   450
instance option :: (complete_linorder) complete_linorder ..
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   451
74334
ead56ad40e15 bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
wenzelm
parents: 69861
diff changeset
   452
unbundle no_lattice_syntax
49190
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   453
e1e1d427747d lattice instances for option type
haftmann
parents: 43815
diff changeset
   454
end