src/HOL/Complete_Lattices.thy
author wenzelm
Wed, 19 Aug 2015 16:21:10 +0200
changeset 60975 5f3d6e16ea78
parent 60758 d8d85a8172b5
child 61630 608520e0e8e2
permissions -rw-r--r--
avoid ambiguities on native Windows, such as / vs. /cygdrive/c/cygwin;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
     1
(*  Author:     Tobias Nipkow, Lawrence C Paulson and Markus Wenzel; Florian Haftmann, TU Muenchen *)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
     2
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
     3
section \<open>Complete lattices\<close>
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
     4
44860
56101fa00193 renamed theory Complete_Lattice to Complete_Lattices, in accordance with Lattices, Orderings etc.
haftmann
parents: 44845
diff changeset
     5
theory Complete_Lattices
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
     6
imports Fun
32139
e271a64f03ff moved complete_lattice &c. into separate theory
haftmann
parents: 32135
diff changeset
     7
begin
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
     8
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
     9
notation
34007
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 32879
diff changeset
    10
  less_eq (infix "\<sqsubseteq>" 50) and
46691
72d81e789106 tuned syntax declarations; tuned structure
haftmann
parents: 46689
diff changeset
    11
  less (infix "\<sqsubset>" 50)
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
    12
32139
e271a64f03ff moved complete_lattice &c. into separate theory
haftmann
parents: 32135
diff changeset
    13
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    14
subsection \<open>Syntactic infimum and supremum operations\<close>
32879
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    15
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    16
class Inf =
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    17
  fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    18
begin
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    19
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    20
definition INFIMUM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    21
  INF_def: "INFIMUM A f = \<Sqinter>(f ` A)"
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    22
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    23
lemma Inf_image_eq [simp]:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    24
  "\<Sqinter>(f ` A) = INFIMUM A f"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    25
  by (simp add: INF_def)
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    26
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    27
lemma INF_image [simp]:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    28
  "INFIMUM (f ` A) g = INFIMUM A (g \<circ> f)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    29
  by (simp only: INF_def image_comp)
54259
71c701dc5bf9 add SUP and INF for conditionally complete lattices
hoelzl
parents: 54257
diff changeset
    30
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    31
lemma INF_identity_eq [simp]:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    32
  "INFIMUM A (\<lambda>x. x) = \<Sqinter>A"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    33
  by (simp add: INF_def)
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    34
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    35
lemma INF_id_eq [simp]:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    36
  "INFIMUM A id = \<Sqinter>A"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    37
  by (simp add: id_def)
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    38
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    39
lemma INF_cong:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    40
  "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> INFIMUM A C = INFIMUM B D"
54259
71c701dc5bf9 add SUP and INF for conditionally complete lattices
hoelzl
parents: 54257
diff changeset
    41
  by (simp add: INF_def image_def)
71c701dc5bf9 add SUP and INF for conditionally complete lattices
hoelzl
parents: 54257
diff changeset
    42
56248
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
    43
lemma strong_INF_cong [cong]:
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
    44
  "A = B \<Longrightarrow> (\<And>x. x \<in> B =simp=> C x = D x) \<Longrightarrow> INFIMUM A C = INFIMUM B D"
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
    45
  unfolding simp_implies_def by (fact INF_cong)
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
    46
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    47
end
32879
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    48
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    49
class Sup =
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    50
  fixes Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900)
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    51
begin
32879
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    52
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    53
definition SUPREMUM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    54
  SUP_def: "SUPREMUM A f = \<Squnion>(f ` A)"
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    55
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    56
lemma Sup_image_eq [simp]:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    57
  "\<Squnion>(f ` A) = SUPREMUM A f"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    58
  by (simp add: SUP_def)
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    59
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    60
lemma SUP_image [simp]:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    61
  "SUPREMUM (f ` A) g = SUPREMUM A (g \<circ> f)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    62
  by (simp only: SUP_def image_comp)
54259
71c701dc5bf9 add SUP and INF for conditionally complete lattices
hoelzl
parents: 54257
diff changeset
    63
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    64
lemma SUP_identity_eq [simp]:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    65
  "SUPREMUM A (\<lambda>x. x) = \<Squnion>A"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    66
  by (simp add: SUP_def)
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    67
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    68
lemma SUP_id_eq [simp]:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    69
  "SUPREMUM A id = \<Squnion>A"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    70
  by (simp add: id_def)
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    71
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    72
lemma SUP_cong:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    73
  "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> SUPREMUM A C = SUPREMUM B D"
54259
71c701dc5bf9 add SUP and INF for conditionally complete lattices
hoelzl
parents: 54257
diff changeset
    74
  by (simp add: SUP_def image_def)
71c701dc5bf9 add SUP and INF for conditionally complete lattices
hoelzl
parents: 54257
diff changeset
    75
56248
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
    76
lemma strong_SUP_cong [cong]:
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
    77
  "A = B \<Longrightarrow> (\<And>x. x \<in> B =simp=> C x = D x) \<Longrightarrow> SUPREMUM A C = SUPREMUM B D"
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
    78
  unfolding simp_implies_def by (fact SUP_cong)
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
    79
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    80
end
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    81
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    82
text \<open>
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
    83
  Note: must use names @{const INFIMUM} and @{const SUPREMUM} here instead of
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    84
  @{text INF} and @{text SUP} to allow the following syntax coexist
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    85
  with the plain constant names.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    86
\<close>
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    87
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    88
syntax
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    89
  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3INF _./ _)" [0, 10] 10)
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    90
  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3INF _:_./ _)" [0, 0, 10] 10)
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    91
  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3SUP _./ _)" [0, 10] 10)
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    92
  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3SUP _:_./ _)" [0, 0, 10] 10)
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    93
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    94
syntax (xsymbols)
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    95
  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    96
  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    97
  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    98
  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    99
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
   100
translations
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
   101
  "INF x y. B"   == "INF x. INF y. B"
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   102
  "INF x. B"     == "CONST INFIMUM CONST UNIV (%x. B)"
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
   103
  "INF x. B"     == "INF x:CONST UNIV. B"
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   104
  "INF x:A. B"   == "CONST INFIMUM A (%x. B)"
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
   105
  "SUP x y. B"   == "SUP x. SUP y. B"
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   106
  "SUP x. B"     == "CONST SUPREMUM CONST UNIV (%x. B)"
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
   107
  "SUP x. B"     == "SUP x:CONST UNIV. B"
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   108
  "SUP x:A. B"   == "CONST SUPREMUM A (%x. B)"
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
   109
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   110
print_translation \<open>
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   111
  [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFIMUM} @{syntax_const "_INF"},
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   112
    Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPREMUM} @{syntax_const "_SUP"}]
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   113
\<close> -- \<open>to avoid eta-contraction of body\<close>
46691
72d81e789106 tuned syntax declarations; tuned structure
haftmann
parents: 46689
diff changeset
   114
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   115
subsection \<open>Abstract complete lattices\<close>
32139
e271a64f03ff moved complete_lattice &c. into separate theory
haftmann
parents: 32135
diff changeset
   116
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   117
text \<open>A complete lattice always has a bottom and a top,
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   118
so we include them into the following type class,
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   119
along with assumptions that define bottom and top
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   120
in terms of infimum and supremum.\<close>
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   121
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   122
class complete_lattice = lattice + Inf + Sup + bot + top +
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   123
  assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> x"
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   124
     and Inf_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> \<Sqinter>A"
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   125
  assumes Sup_upper: "x \<in> A \<Longrightarrow> x \<sqsubseteq> \<Squnion>A"
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   126
     and Sup_least: "(\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq> z) \<Longrightarrow> \<Squnion>A \<sqsubseteq> z"
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   127
  assumes Inf_empty [simp]: "\<Sqinter>{} = \<top>"
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   128
  assumes Sup_empty [simp]: "\<Squnion>{} = \<bottom>"
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   129
begin
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   130
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   131
subclass bounded_lattice
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   132
proof
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   133
  fix a
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   134
  show "\<bottom> \<le> a" by (auto intro: Sup_least simp only: Sup_empty [symmetric])
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   135
  show "a \<le> \<top>" by (auto intro: Inf_greatest simp only: Inf_empty [symmetric])
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   136
qed
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   137
32678
de1f7d4da21a added dual for complete lattice
haftmann
parents: 32642
diff changeset
   138
lemma dual_complete_lattice:
44845
5e51075cbd97 added syntactic classes for "inf" and "sup"
krauss
parents: 44322
diff changeset
   139
  "class.complete_lattice Sup Inf sup (op \<ge>) (op >) inf \<top> \<bottom>"
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   140
  by (auto intro!: class.complete_lattice.intro dual_lattice)
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   141
    (unfold_locales, (fact Inf_empty Sup_empty
34007
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 32879
diff changeset
   142
        Sup_upper Sup_least Inf_lower Inf_greatest)+)
32678
de1f7d4da21a added dual for complete lattice
haftmann
parents: 32642
diff changeset
   143
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   144
end
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   145
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   146
context complete_lattice
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   147
begin
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   148
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   149
lemma INF_foundation_dual:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   150
  "Sup.SUPREMUM Inf = INFIMUM"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   151
  by (simp add: fun_eq_iff Sup.SUP_def)
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   152
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   153
lemma SUP_foundation_dual:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   154
  "Inf.INFIMUM Sup = SUPREMUM"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   155
  by (simp add: fun_eq_iff Inf.INF_def)
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   156
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   157
lemma Sup_eqI:
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   158
  "(\<And>y. y \<in> A \<Longrightarrow> y \<le> x) \<Longrightarrow> (\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y) \<Longrightarrow> \<Squnion>A = x"
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   159
  by (blast intro: antisym Sup_least Sup_upper)
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   160
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   161
lemma Inf_eqI:
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   162
  "(\<And>i. i \<in> A \<Longrightarrow> x \<le> i) \<Longrightarrow> (\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> i) \<Longrightarrow> y \<le> x) \<Longrightarrow> \<Sqinter>A = x"
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   163
  by (blast intro: antisym Inf_greatest Inf_lower)
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   164
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   165
lemma SUP_eqI:
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   166
  "(\<And>i. i \<in> A \<Longrightarrow> f i \<le> x) \<Longrightarrow> (\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<le> y) \<Longrightarrow> x \<le> y) \<Longrightarrow> (\<Squnion>i\<in>A. f i) = x"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   167
  using Sup_eqI [of "f ` A" x] by auto
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   168
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   169
lemma INF_eqI:
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   170
  "(\<And>i. i \<in> A \<Longrightarrow> x \<le> f i) \<Longrightarrow> (\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<ge> y) \<Longrightarrow> x \<ge> y) \<Longrightarrow> (\<Sqinter>i\<in>A. f i) = x"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   171
  using Inf_eqI [of "f ` A" x] by auto
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   172
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   173
lemma INF_lower: "i \<in> A \<Longrightarrow> (\<Sqinter>i\<in>A. f i) \<sqsubseteq> f i"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   174
  using Inf_lower [of _ "f ` A"] by simp
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   175
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   176
lemma INF_greatest: "(\<And>i. i \<in> A \<Longrightarrow> u \<sqsubseteq> f i) \<Longrightarrow> u \<sqsubseteq> (\<Sqinter>i\<in>A. f i)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   177
  using Inf_greatest [of "f ` A"] by auto
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   178
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   179
lemma SUP_upper: "i \<in> A \<Longrightarrow> f i \<sqsubseteq> (\<Squnion>i\<in>A. f i)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   180
  using Sup_upper [of _ "f ` A"] by simp
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   181
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   182
lemma SUP_least: "(\<And>i. i \<in> A \<Longrightarrow> f i \<sqsubseteq> u) \<Longrightarrow> (\<Squnion>i\<in>A. f i) \<sqsubseteq> u"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   183
  using Sup_least [of "f ` A"] by auto
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   184
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   185
lemma Inf_lower2: "u \<in> A \<Longrightarrow> u \<sqsubseteq> v \<Longrightarrow> \<Sqinter>A \<sqsubseteq> v"
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   186
  using Inf_lower [of u A] by auto
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   187
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   188
lemma INF_lower2: "i \<in> A \<Longrightarrow> f i \<sqsubseteq> u \<Longrightarrow> (\<Sqinter>i\<in>A. f i) \<sqsubseteq> u"
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   189
  using INF_lower [of i A f] by auto
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   190
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   191
lemma Sup_upper2: "u \<in> A \<Longrightarrow> v \<sqsubseteq> u \<Longrightarrow> v \<sqsubseteq> \<Squnion>A"
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   192
  using Sup_upper [of u A] by auto
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   193
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   194
lemma SUP_upper2: "i \<in> A \<Longrightarrow> u \<sqsubseteq> f i \<Longrightarrow> u \<sqsubseteq> (\<Squnion>i\<in>A. f i)"
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   195
  using SUP_upper [of i A f] by auto
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   196
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   197
lemma le_Inf_iff: "b \<sqsubseteq> \<Sqinter>A \<longleftrightarrow> (\<forall>a\<in>A. b \<sqsubseteq> a)"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   198
  by (auto intro: Inf_greatest dest: Inf_lower)
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   199
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   200
lemma le_INF_iff: "u \<sqsubseteq> (\<Sqinter>i\<in>A. f i) \<longleftrightarrow> (\<forall>i\<in>A. u \<sqsubseteq> f i)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   201
  using le_Inf_iff [of _ "f ` A"] by simp
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   202
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   203
lemma Sup_le_iff: "\<Squnion>A \<sqsubseteq> b \<longleftrightarrow> (\<forall>a\<in>A. a \<sqsubseteq> b)"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   204
  by (auto intro: Sup_least dest: Sup_upper)
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   205
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   206
lemma SUP_le_iff: "(\<Squnion>i\<in>A. f i) \<sqsubseteq> u \<longleftrightarrow> (\<forall>i\<in>A. f i \<sqsubseteq> u)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   207
  using Sup_le_iff [of "f ` A"] by simp
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   208
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   209
lemma Inf_insert [simp]: "\<Sqinter>insert a A = a \<sqinter> \<Sqinter>A"
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   210
  by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower)
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   211
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   212
lemma INF_insert [simp]: "(\<Sqinter>x\<in>insert a A. f x) = f a \<sqinter> INFIMUM A f"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   213
  unfolding INF_def Inf_insert by simp
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   214
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   215
lemma Sup_insert [simp]: "\<Squnion>insert a A = a \<squnion> \<Squnion>A"
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   216
  by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper)
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   217
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   218
lemma SUP_insert [simp]: "(\<Squnion>x\<in>insert a A. f x) = f a \<squnion> SUPREMUM A f"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   219
  unfolding SUP_def Sup_insert by simp
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   220
44067
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   221
lemma INF_empty [simp]: "(\<Sqinter>x\<in>{}. f x) = \<top>"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   222
  by (simp add: INF_def)
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   223
44067
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   224
lemma SUP_empty [simp]: "(\<Squnion>x\<in>{}. f x) = \<bottom>"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   225
  by (simp add: SUP_def)
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   226
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   227
lemma Inf_UNIV [simp]:
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   228
  "\<Sqinter>UNIV = \<bottom>"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   229
  by (auto intro!: antisym Inf_lower)
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   230
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   231
lemma Sup_UNIV [simp]:
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   232
  "\<Squnion>UNIV = \<top>"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   233
  by (auto intro!: antisym Sup_upper)
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   234
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   235
lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<sqsubseteq> a}"
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   236
  by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   237
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   238
lemma Sup_Inf:  "\<Squnion>A = \<Sqinter>{b. \<forall>a \<in> A. a \<sqsubseteq> b}"
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   239
  by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   240
43899
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   241
lemma Inf_superset_mono: "B \<subseteq> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> \<Sqinter>B"
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   242
  by (auto intro: Inf_greatest Inf_lower)
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   243
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   244
lemma Sup_subset_mono: "A \<subseteq> B \<Longrightarrow> \<Squnion>A \<sqsubseteq> \<Squnion>B"
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   245
  by (auto intro: Sup_least Sup_upper)
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   246
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   247
lemma Inf_mono:
41971
a54e8e95fe96 add lemmas for SUP and INF
hoelzl
parents: 41082
diff changeset
   248
  assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<sqsubseteq> b"
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   249
  shows "\<Sqinter>A \<sqsubseteq> \<Sqinter>B"
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   250
proof (rule Inf_greatest)
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   251
  fix b assume "b \<in> B"
41971
a54e8e95fe96 add lemmas for SUP and INF
hoelzl
parents: 41082
diff changeset
   252
  with assms obtain a where "a \<in> A" and "a \<sqsubseteq> b" by blast
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   253
  from \<open>a \<in> A\<close> have "\<Sqinter>A \<sqsubseteq> a" by (rule Inf_lower)
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   254
  with \<open>a \<sqsubseteq> b\<close> show "\<Sqinter>A \<sqsubseteq> b" by auto
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   255
qed
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   256
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   257
lemma INF_mono:
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   258
  "(\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<sqsubseteq> g m) \<Longrightarrow> (\<Sqinter>n\<in>A. f n) \<sqsubseteq> (\<Sqinter>n\<in>B. g n)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   259
  using Inf_mono [of "g ` B" "f ` A"] by auto
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   260
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   261
lemma Sup_mono:
41971
a54e8e95fe96 add lemmas for SUP and INF
hoelzl
parents: 41082
diff changeset
   262
  assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<sqsubseteq> b"
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   263
  shows "\<Squnion>A \<sqsubseteq> \<Squnion>B"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   264
proof (rule Sup_least)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   265
  fix a assume "a \<in> A"
41971
a54e8e95fe96 add lemmas for SUP and INF
hoelzl
parents: 41082
diff changeset
   266
  with assms obtain b where "b \<in> B" and "a \<sqsubseteq> b" by blast
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   267
  from \<open>b \<in> B\<close> have "b \<sqsubseteq> \<Squnion>B" by (rule Sup_upper)
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   268
  with \<open>a \<sqsubseteq> b\<close> show "a \<sqsubseteq> \<Squnion>B" by auto
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   269
qed
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   270
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   271
lemma SUP_mono:
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   272
  "(\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<sqsubseteq> g m) \<Longrightarrow> (\<Squnion>n\<in>A. f n) \<sqsubseteq> (\<Squnion>n\<in>B. g n)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   273
  using Sup_mono [of "f ` A" "g ` B"] by auto
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   274
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   275
lemma INF_superset_mono:
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   276
  "B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<sqsubseteq> g x) \<Longrightarrow> (\<Sqinter>x\<in>A. f x) \<sqsubseteq> (\<Sqinter>x\<in>B. g x)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   277
  -- \<open>The last inclusion is POSITIVE!\<close>
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   278
  by (blast intro: INF_mono dest: subsetD)
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   279
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   280
lemma SUP_subset_mono:
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   281
  "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<sqsubseteq> g x) \<Longrightarrow> (\<Squnion>x\<in>A. f x) \<sqsubseteq> (\<Squnion>x\<in>B. g x)"
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   282
  by (blast intro: SUP_mono dest: subsetD)
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   283
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   284
lemma Inf_less_eq:
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   285
  assumes "\<And>v. v \<in> A \<Longrightarrow> v \<sqsubseteq> u"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   286
    and "A \<noteq> {}"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   287
  shows "\<Sqinter>A \<sqsubseteq> u"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   288
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   289
  from \<open>A \<noteq> {}\<close> obtain v where "v \<in> A" by blast
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   290
  moreover from \<open>v \<in> A\<close> assms(1) have "v \<sqsubseteq> u" by blast
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   291
  ultimately show ?thesis by (rule Inf_lower2)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   292
qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   293
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   294
lemma less_eq_Sup:
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   295
  assumes "\<And>v. v \<in> A \<Longrightarrow> u \<sqsubseteq> v"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   296
    and "A \<noteq> {}"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   297
  shows "u \<sqsubseteq> \<Squnion>A"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   298
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   299
  from \<open>A \<noteq> {}\<close> obtain v where "v \<in> A" by blast
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   300
  moreover from \<open>v \<in> A\<close> assms(1) have "u \<sqsubseteq> v" by blast
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   301
  ultimately show ?thesis by (rule Sup_upper2)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   302
qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   303
56212
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 56166
diff changeset
   304
lemma SUP_eq:
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   305
  assumes "\<And>i. i \<in> A \<Longrightarrow> \<exists>j\<in>B. f i \<le> g j"
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   306
  assumes "\<And>j. j \<in> B \<Longrightarrow> \<exists>i\<in>A. g j \<le> f i"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   307
  shows "(\<Squnion>i\<in>A. f i) = (\<Squnion>j\<in>B. g j)"
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   308
  by (intro antisym SUP_least) (blast intro: SUP_upper2 dest: assms)+
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   309
56212
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 56166
diff changeset
   310
lemma INF_eq:
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   311
  assumes "\<And>i. i \<in> A \<Longrightarrow> \<exists>j\<in>B. f i \<ge> g j"
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   312
  assumes "\<And>j. j \<in> B \<Longrightarrow> \<exists>i\<in>A. g j \<ge> f i"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   313
  shows "(\<Sqinter>i\<in>A. f i) = (\<Sqinter>j\<in>B. g j)"
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   314
  by (intro antisym INF_greatest) (blast intro: INF_lower2 dest: assms)+
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   315
43899
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   316
lemma less_eq_Inf_inter: "\<Sqinter>A \<squnion> \<Sqinter>B \<sqsubseteq> \<Sqinter>(A \<inter> B)"
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   317
  by (auto intro: Inf_greatest Inf_lower)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   318
43899
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   319
lemma Sup_inter_less_eq: "\<Squnion>(A \<inter> B) \<sqsubseteq> \<Squnion>A \<sqinter> \<Squnion>B "
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   320
  by (auto intro: Sup_least Sup_upper)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   321
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   322
lemma Inf_union_distrib: "\<Sqinter>(A \<union> B) = \<Sqinter>A \<sqinter> \<Sqinter>B"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   323
  by (rule antisym) (auto intro: Inf_greatest Inf_lower le_infI1 le_infI2)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   324
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   325
lemma INF_union:
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   326
  "(\<Sqinter>i \<in> A \<union> B. M i) = (\<Sqinter>i \<in> A. M i) \<sqinter> (\<Sqinter>i\<in>B. M i)"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   327
  by (auto intro!: antisym INF_mono intro: le_infI1 le_infI2 INF_greatest INF_lower)
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   328
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   329
lemma Sup_union_distrib: "\<Squnion>(A \<union> B) = \<Squnion>A \<squnion> \<Squnion>B"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   330
  by (rule antisym) (auto intro: Sup_least Sup_upper le_supI1 le_supI2)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   331
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   332
lemma SUP_union:
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   333
  "(\<Squnion>i \<in> A \<union> B. M i) = (\<Squnion>i \<in> A. M i) \<squnion> (\<Squnion>i\<in>B. M i)"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   334
  by (auto intro!: antisym SUP_mono intro: le_supI1 le_supI2 SUP_least SUP_upper)
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   335
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   336
lemma INF_inf_distrib: "(\<Sqinter>a\<in>A. f a) \<sqinter> (\<Sqinter>a\<in>A. g a) = (\<Sqinter>a\<in>A. f a \<sqinter> g a)"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   337
  by (rule antisym) (rule INF_greatest, auto intro: le_infI1 le_infI2 INF_lower INF_mono)
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   338
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   339
lemma SUP_sup_distrib: "(\<Squnion>a\<in>A. f a) \<squnion> (\<Squnion>a\<in>A. g a) = (\<Squnion>a\<in>A. f a \<squnion> g a)" (is "?L = ?R")
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   340
proof (rule antisym)
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   341
  show "?L \<le> ?R" by (auto intro: le_supI1 le_supI2 SUP_upper SUP_mono)
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   342
next
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   343
  show "?R \<le> ?L" by (rule SUP_least) (auto intro: le_supI1 le_supI2 SUP_upper)
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   344
qed
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   345
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   346
lemma Inf_top_conv [simp]:
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   347
  "\<Sqinter>A = \<top> \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   348
  "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   349
proof -
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   350
  show "\<Sqinter>A = \<top> \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   351
  proof
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   352
    assume "\<forall>x\<in>A. x = \<top>"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   353
    then have "A = {} \<or> A = {\<top>}" by auto
44919
482f1807976e tune proofs
noschinl
parents: 44918
diff changeset
   354
    then show "\<Sqinter>A = \<top>" by auto
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   355
  next
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   356
    assume "\<Sqinter>A = \<top>"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   357
    show "\<forall>x\<in>A. x = \<top>"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   358
    proof (rule ccontr)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   359
      assume "\<not> (\<forall>x\<in>A. x = \<top>)"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   360
      then obtain x where "x \<in> A" and "x \<noteq> \<top>" by blast
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   361
      then obtain B where "A = insert x B" by blast
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   362
      with \<open>\<Sqinter>A = \<top>\<close> \<open>x \<noteq> \<top>\<close> show False by simp
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   363
    qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   364
  qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   365
  then show "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)" by auto
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   366
qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   367
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   368
lemma INF_top_conv [simp]:
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   369
  "(\<Sqinter>x\<in>A. B x) = \<top> \<longleftrightarrow> (\<forall>x\<in>A. B x = \<top>)"
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   370
  "\<top> = (\<Sqinter>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = \<top>)"
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   371
  using Inf_top_conv [of "B ` A"] by simp_all
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   372
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   373
lemma Sup_bot_conv [simp]:
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   374
  "\<Squnion>A = \<bottom> \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)" (is ?P)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   375
  "\<bottom> = \<Squnion>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)" (is ?Q)
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   376
  using dual_complete_lattice
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   377
  by (rule complete_lattice.Inf_top_conv)+
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   378
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   379
lemma SUP_bot_conv [simp]:
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   380
 "(\<Squnion>x\<in>A. B x) = \<bottom> \<longleftrightarrow> (\<forall>x\<in>A. B x = \<bottom>)"
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   381
 "\<bottom> = (\<Squnion>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = \<bottom>)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   382
  using Sup_bot_conv [of "B ` A"] by simp_all
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   383
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   384
lemma INF_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Sqinter>i\<in>A. f) = f"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   385
  by (auto intro: antisym INF_lower INF_greatest)
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   386
43870
92129f505125 structuring duals together
haftmann
parents: 43868
diff changeset
   387
lemma SUP_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Squnion>i\<in>A. f) = f"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   388
  by (auto intro: antisym SUP_upper SUP_least)
43870
92129f505125 structuring duals together
haftmann
parents: 43868
diff changeset
   389
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   390
lemma INF_top [simp]: "(\<Sqinter>x\<in>A. \<top>) = \<top>"
44921
58eef4843641 tuned proofs
huffman
parents: 44920
diff changeset
   391
  by (cases "A = {}") simp_all
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
   392
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   393
lemma SUP_bot [simp]: "(\<Squnion>x\<in>A. \<bottom>) = \<bottom>"
44921
58eef4843641 tuned proofs
huffman
parents: 44920
diff changeset
   394
  by (cases "A = {}") simp_all
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
   395
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   396
lemma INF_commute: "(\<Sqinter>i\<in>A. \<Sqinter>j\<in>B. f i j) = (\<Sqinter>j\<in>B. \<Sqinter>i\<in>A. f i j)"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   397
  by (iprover intro: INF_lower INF_greatest order_trans antisym)
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   398
43870
92129f505125 structuring duals together
haftmann
parents: 43868
diff changeset
   399
lemma SUP_commute: "(\<Squnion>i\<in>A. \<Squnion>j\<in>B. f i j) = (\<Squnion>j\<in>B. \<Squnion>i\<in>A. f i j)"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   400
  by (iprover intro: SUP_upper SUP_least order_trans antisym)
43870
92129f505125 structuring duals together
haftmann
parents: 43868
diff changeset
   401
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   402
lemma INF_absorb:
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   403
  assumes "k \<in> I"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   404
  shows "A k \<sqinter> (\<Sqinter>i\<in>I. A i) = (\<Sqinter>i\<in>I. A i)"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   405
proof -
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   406
  from assms obtain J where "I = insert k J" by blast
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   407
  then show ?thesis by simp
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   408
qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   409
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   410
lemma SUP_absorb:
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   411
  assumes "k \<in> I"
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   412
  shows "A k \<squnion> (\<Squnion>i\<in>I. A i) = (\<Squnion>i\<in>I. A i)"
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   413
proof -
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   414
  from assms obtain J where "I = insert k J" by blast
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   415
  then show ?thesis by simp
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   416
qed
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   417
57448
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   418
lemma INF_inf_const1:
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   419
  "I \<noteq> {} \<Longrightarrow> (INF i:I. inf x (f i)) = inf x (INF i:I. f i)"
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   420
  by (intro antisym INF_greatest inf_mono order_refl INF_lower)
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   421
     (auto intro: INF_lower2 le_infI2 intro!: INF_mono)
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   422
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   423
lemma INF_inf_const2:
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   424
  "I \<noteq> {} \<Longrightarrow> (INF i:I. inf (f i) x) = inf (INF i:I. f i) x"
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   425
  using INF_inf_const1[of I x f] by (simp add: inf_commute)
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   426
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   427
lemma INF_constant:
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   428
  "(\<Sqinter>y\<in>A. c) = (if A = {} then \<top> else c)"
44921
58eef4843641 tuned proofs
huffman
parents: 44920
diff changeset
   429
  by simp
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   430
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   431
lemma SUP_constant:
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   432
  "(\<Squnion>y\<in>A. c) = (if A = {} then \<bottom> else c)"
44921
58eef4843641 tuned proofs
huffman
parents: 44920
diff changeset
   433
  by simp
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   434
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   435
lemma less_INF_D:
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   436
  assumes "y < (\<Sqinter>i\<in>A. f i)" "i \<in> A" shows "y < f i"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   437
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   438
  note \<open>y < (\<Sqinter>i\<in>A. f i)\<close>
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   439
  also have "(\<Sqinter>i\<in>A. f i) \<le> f i" using \<open>i \<in> A\<close>
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   440
    by (rule INF_lower)
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   441
  finally show "y < f i" .
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   442
qed
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   443
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   444
lemma SUP_lessD:
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   445
  assumes "(\<Squnion>i\<in>A. f i) < y" "i \<in> A" shows "f i < y"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   446
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   447
  have "f i \<le> (\<Squnion>i\<in>A. f i)" using \<open>i \<in> A\<close>
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   448
    by (rule SUP_upper)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   449
  also note \<open>(\<Squnion>i\<in>A. f i) < y\<close>
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   450
  finally show "f i < y" .
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   451
qed
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   452
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   453
lemma INF_UNIV_bool_expand:
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   454
  "(\<Sqinter>b. A b) = A True \<sqinter> A False"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   455
  by (simp add: UNIV_bool inf_commute)
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   456
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   457
lemma SUP_UNIV_bool_expand:
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   458
  "(\<Squnion>b. A b) = A True \<squnion> A False"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   459
  by (simp add: UNIV_bool sup_commute)
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   460
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   461
lemma Inf_le_Sup: "A \<noteq> {} \<Longrightarrow> Inf A \<le> Sup A"
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   462
  by (blast intro: Sup_upper2 Inf_lower ex_in_conv)
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   463
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   464
lemma INF_le_SUP: "A \<noteq> {} \<Longrightarrow> INFIMUM A f \<le> SUPREMUM A f"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   465
  using Inf_le_Sup [of "f ` A"] by simp
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   466
54414
72949fae4f81 add equalities for SUP and INF over constant functions
hoelzl
parents: 54259
diff changeset
   467
lemma INF_eq_const:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   468
  "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i = x) \<Longrightarrow> INFIMUM I f = x"
54414
72949fae4f81 add equalities for SUP and INF over constant functions
hoelzl
parents: 54259
diff changeset
   469
  by (auto intro: INF_eqI)
72949fae4f81 add equalities for SUP and INF over constant functions
hoelzl
parents: 54259
diff changeset
   470
56248
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   471
lemma SUP_eq_const:
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   472
  "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i = x) \<Longrightarrow> SUPREMUM I f = x"
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   473
  by (auto intro: SUP_eqI)
54414
72949fae4f81 add equalities for SUP and INF over constant functions
hoelzl
parents: 54259
diff changeset
   474
72949fae4f81 add equalities for SUP and INF over constant functions
hoelzl
parents: 54259
diff changeset
   475
lemma INF_eq_iff:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   476
  "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<le> c) \<Longrightarrow> (INFIMUM I f = c) \<longleftrightarrow> (\<forall>i\<in>I. f i = c)"
56248
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   477
  using INF_eq_const [of I f c] INF_lower [of _ I f]
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   478
  by (auto intro: antisym cong del: strong_INF_cong)
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   479
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   480
lemma SUP_eq_iff:
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   481
  "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> c \<le> f i) \<Longrightarrow> (SUPREMUM I f = c) \<longleftrightarrow> (\<forall>i\<in>I. f i = c)"
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   482
  using SUP_eq_const [of I f c] SUP_upper [of _ I f]
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   483
  by (auto intro: antisym cong del: strong_SUP_cong)
54414
72949fae4f81 add equalities for SUP and INF over constant functions
hoelzl
parents: 54259
diff changeset
   484
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   485
end
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   486
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   487
class complete_distrib_lattice = complete_lattice +
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   488
  assumes sup_Inf: "a \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> b)"
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   489
  assumes inf_Sup: "a \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)"
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   490
begin
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   491
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   492
lemma sup_INF:
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   493
  "a \<squnion> (\<Sqinter>b\<in>B. f b) = (\<Sqinter>b\<in>B. a \<squnion> f b)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   494
  by (simp only: INF_def sup_Inf image_image)
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   495
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   496
lemma inf_SUP:
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   497
  "a \<sqinter> (\<Squnion>b\<in>B. f b) = (\<Squnion>b\<in>B. a \<sqinter> f b)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   498
  by (simp only: SUP_def inf_Sup image_image)
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   499
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
   500
lemma dual_complete_distrib_lattice:
44845
5e51075cbd97 added syntactic classes for "inf" and "sup"
krauss
parents: 44322
diff changeset
   501
  "class.complete_distrib_lattice Sup Inf sup (op \<ge>) (op >) inf \<top> \<bottom>"
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   502
  apply (rule class.complete_distrib_lattice.intro)
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   503
  apply (fact dual_complete_lattice)
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   504
  apply (rule class.complete_distrib_lattice_axioms.intro)
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
   505
  apply (simp_all only: INF_foundation_dual SUP_foundation_dual inf_Sup sup_Inf)
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
   506
  done
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   507
44322
43b465f4c480 more concise definition for Inf, Sup on bool
haftmann
parents: 44104
diff changeset
   508
subclass distrib_lattice proof
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   509
  fix a b c
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   510
  from sup_Inf have "a \<squnion> \<Sqinter>{b, c} = (\<Sqinter>d\<in>{b, c}. a \<squnion> d)" .
44919
482f1807976e tune proofs
noschinl
parents: 44918
diff changeset
   511
  then show "a \<squnion> b \<sqinter> c = (a \<squnion> b) \<sqinter> (a \<squnion> c)" by (simp add: INF_def)
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   512
qed
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   513
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   514
lemma Inf_sup:
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   515
  "\<Sqinter>B \<squnion> a = (\<Sqinter>b\<in>B. b \<squnion> a)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   516
  by (simp add: sup_Inf sup_commute)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   517
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   518
lemma Sup_inf:
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   519
  "\<Squnion>B \<sqinter> a = (\<Squnion>b\<in>B. b \<sqinter> a)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   520
  by (simp add: inf_Sup inf_commute)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   521
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   522
lemma INF_sup: 
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   523
  "(\<Sqinter>b\<in>B. f b) \<squnion> a = (\<Sqinter>b\<in>B. f b \<squnion> a)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   524
  by (simp add: sup_INF sup_commute)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   525
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   526
lemma SUP_inf:
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   527
  "(\<Squnion>b\<in>B. f b) \<sqinter> a = (\<Squnion>b\<in>B. f b \<sqinter> a)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   528
  by (simp add: inf_SUP inf_commute)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   529
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   530
lemma Inf_sup_eq_top_iff:
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   531
  "(\<Sqinter>B \<squnion> a = \<top>) \<longleftrightarrow> (\<forall>b\<in>B. b \<squnion> a = \<top>)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   532
  by (simp only: Inf_sup INF_top_conv)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   533
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   534
lemma Sup_inf_eq_bot_iff:
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   535
  "(\<Squnion>B \<sqinter> a = \<bottom>) \<longleftrightarrow> (\<forall>b\<in>B. b \<sqinter> a = \<bottom>)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   536
  by (simp only: Sup_inf SUP_bot_conv)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   537
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   538
lemma INF_sup_distrib2:
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   539
  "(\<Sqinter>a\<in>A. f a) \<squnion> (\<Sqinter>b\<in>B. g b) = (\<Sqinter>a\<in>A. \<Sqinter>b\<in>B. f a \<squnion> g b)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   540
  by (subst INF_commute) (simp add: sup_INF INF_sup)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   541
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   542
lemma SUP_inf_distrib2:
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   543
  "(\<Squnion>a\<in>A. f a) \<sqinter> (\<Squnion>b\<in>B. g b) = (\<Squnion>a\<in>A. \<Squnion>b\<in>B. f a \<sqinter> g b)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   544
  by (subst SUP_commute) (simp add: inf_SUP SUP_inf)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   545
56074
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   546
context
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   547
  fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   548
  assumes "mono f"
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   549
begin
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   550
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   551
lemma mono_Inf:
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   552
  shows "f (\<Sqinter>A) \<le> (\<Sqinter>x\<in>A. f x)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   553
  using \<open>mono f\<close> by (auto intro: complete_lattice_class.INF_greatest Inf_lower dest: monoD)
56074
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   554
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   555
lemma mono_Sup:
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   556
  shows "(\<Squnion>x\<in>A. f x) \<le> f (\<Squnion>A)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   557
  using \<open>mono f\<close> by (auto intro: complete_lattice_class.SUP_least Sup_upper dest: monoD)
56074
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   558
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 58889
diff changeset
   559
lemma mono_INF:
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 58889
diff changeset
   560
  "f (INF i : I. A i) \<le> (INF x : I. f (A x))"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   561
  by (intro complete_lattice_class.INF_greatest monoD[OF \<open>mono f\<close>] INF_lower)
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 58889
diff changeset
   562
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 58889
diff changeset
   563
lemma mono_SUP:
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 58889
diff changeset
   564
  "(SUP x : I. f (A x)) \<le> f (SUP i : I. A i)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   565
  by (intro complete_lattice_class.SUP_least monoD[OF \<open>mono f\<close>] SUP_upper)
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 58889
diff changeset
   566
56074
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   567
end
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   568
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   569
end
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   570
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
   571
class complete_boolean_algebra = boolean_algebra + complete_distrib_lattice
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   572
begin
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   573
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   574
lemma dual_complete_boolean_algebra:
44845
5e51075cbd97 added syntactic classes for "inf" and "sup"
krauss
parents: 44322
diff changeset
   575
  "class.complete_boolean_algebra Sup Inf sup (op \<ge>) (op >) inf \<top> \<bottom> (\<lambda>x y. x \<squnion> - y) uminus"
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
   576
  by (rule class.complete_boolean_algebra.intro, rule dual_complete_distrib_lattice, rule dual_boolean_algebra)
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   577
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   578
lemma uminus_Inf:
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   579
  "- (\<Sqinter>A) = \<Squnion>(uminus ` A)"
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   580
proof (rule antisym)
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   581
  show "- \<Sqinter>A \<le> \<Squnion>(uminus ` A)"
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   582
    by (rule compl_le_swap2, rule Inf_greatest, rule compl_le_swap2, rule Sup_upper) simp
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   583
  show "\<Squnion>(uminus ` A) \<le> - \<Sqinter>A"
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   584
    by (rule Sup_least, rule compl_le_swap1, rule Inf_lower) auto
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   585
qed
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   586
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   587
lemma uminus_INF: "- (\<Sqinter>x\<in>A. B x) = (\<Squnion>x\<in>A. - B x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   588
  by (simp only: INF_def SUP_def uminus_Inf image_image)
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   589
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   590
lemma uminus_Sup:
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   591
  "- (\<Squnion>A) = \<Sqinter>(uminus ` A)"
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   592
proof -
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   593
  have "\<Squnion>A = - \<Sqinter>(uminus ` A)" by (simp add: image_image uminus_INF)
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   594
  then show ?thesis by simp
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   595
qed
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   596
  
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   597
lemma uminus_SUP: "- (\<Squnion>x\<in>A. B x) = (\<Sqinter>x\<in>A. - B x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   598
  by (simp only: INF_def SUP_def uminus_Sup image_image)
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   599
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   600
end
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   601
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   602
class complete_linorder = linorder + complete_lattice
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   603
begin
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   604
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   605
lemma dual_complete_linorder:
44845
5e51075cbd97 added syntactic classes for "inf" and "sup"
krauss
parents: 44322
diff changeset
   606
  "class.complete_linorder Sup Inf sup (op \<ge>) (op >) inf \<top> \<bottom>"
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   607
  by (rule class.complete_linorder.intro, rule dual_complete_lattice, rule dual_linorder)
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   608
51386
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   609
lemma complete_linorder_inf_min: "inf = min"
51540
eea5c4ca4a0e explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
haftmann
parents: 51489
diff changeset
   610
  by (auto intro: antisym simp add: min_def fun_eq_iff)
51386
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   611
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   612
lemma complete_linorder_sup_max: "sup = max"
51540
eea5c4ca4a0e explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
haftmann
parents: 51489
diff changeset
   613
  by (auto intro: antisym simp add: max_def fun_eq_iff)
51386
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   614
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   615
lemma Inf_less_iff:
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   616
  "\<Sqinter>S \<sqsubset> a \<longleftrightarrow> (\<exists>x\<in>S. x \<sqsubset> a)"
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   617
  unfolding not_le [symmetric] le_Inf_iff by auto
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   618
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   619
lemma INF_less_iff:
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   620
  "(\<Sqinter>i\<in>A. f i) \<sqsubset> a \<longleftrightarrow> (\<exists>x\<in>A. f x \<sqsubset> a)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   621
  using Inf_less_iff [of "f ` A"] by simp
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   622
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   623
lemma less_Sup_iff:
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   624
  "a \<sqsubset> \<Squnion>S \<longleftrightarrow> (\<exists>x\<in>S. a \<sqsubset> x)"
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   625
  unfolding not_le [symmetric] Sup_le_iff by auto
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   626
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   627
lemma less_SUP_iff:
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   628
  "a \<sqsubset> (\<Squnion>i\<in>A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a \<sqsubset> f x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   629
  using less_Sup_iff [of _ "f ` A"] by simp
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   630
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   631
lemma Sup_eq_top_iff [simp]:
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   632
  "\<Squnion>A = \<top> \<longleftrightarrow> (\<forall>x<\<top>. \<exists>i\<in>A. x < i)"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   633
proof
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   634
  assume *: "\<Squnion>A = \<top>"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   635
  show "(\<forall>x<\<top>. \<exists>i\<in>A. x < i)" unfolding * [symmetric]
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   636
  proof (intro allI impI)
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   637
    fix x assume "x < \<Squnion>A" then show "\<exists>i\<in>A. x < i"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   638
      unfolding less_Sup_iff by auto
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   639
  qed
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   640
next
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   641
  assume *: "\<forall>x<\<top>. \<exists>i\<in>A. x < i"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   642
  show "\<Squnion>A = \<top>"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   643
  proof (rule ccontr)
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   644
    assume "\<Squnion>A \<noteq> \<top>"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   645
    with top_greatest [of "\<Squnion>A"]
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   646
    have "\<Squnion>A < \<top>" unfolding le_less by auto
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   647
    then have "\<Squnion>A < \<Squnion>A"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   648
      using * unfolding less_Sup_iff by auto
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   649
    then show False by auto
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   650
  qed
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   651
qed
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   652
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   653
lemma SUP_eq_top_iff [simp]:
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   654
  "(\<Squnion>i\<in>A. f i) = \<top> \<longleftrightarrow> (\<forall>x<\<top>. \<exists>i\<in>A. x < f i)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   655
  using Sup_eq_top_iff [of "f ` A"] by simp
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   656
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   657
lemma Inf_eq_bot_iff [simp]:
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   658
  "\<Sqinter>A = \<bottom> \<longleftrightarrow> (\<forall>x>\<bottom>. \<exists>i\<in>A. i < x)"
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   659
  using dual_complete_linorder
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   660
  by (rule complete_linorder.Sup_eq_top_iff)
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   661
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   662
lemma INF_eq_bot_iff [simp]:
43967
610efb6bda1f more coherent structure in and across theories
haftmann
parents: 43944
diff changeset
   663
  "(\<Sqinter>i\<in>A. f i) = \<bottom> \<longleftrightarrow> (\<forall>x>\<bottom>. \<exists>i\<in>A. f i < x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   664
  using Inf_eq_bot_iff [of "f ` A"] by simp
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   665
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   666
lemma Inf_le_iff: "\<Sqinter>A \<le> x \<longleftrightarrow> (\<forall>y>x. \<exists>a\<in>A. y > a)"
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   667
proof safe
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   668
  fix y assume "x \<ge> \<Sqinter>A" "y > x"
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   669
  then have "y > \<Sqinter>A" by auto
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   670
  then show "\<exists>a\<in>A. y > a"
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   671
    unfolding Inf_less_iff .
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   672
qed (auto elim!: allE[of _ "\<Sqinter>A"] simp add: not_le[symmetric] Inf_lower)
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   673
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   674
lemma INF_le_iff:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   675
  "INFIMUM A f \<le> x \<longleftrightarrow> (\<forall>y>x. \<exists>i\<in>A. y > f i)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   676
  using Inf_le_iff [of "f ` A"] by simp
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   677
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   678
lemma le_Sup_iff: "x \<le> \<Squnion>A \<longleftrightarrow> (\<forall>y<x. \<exists>a\<in>A. y < a)"
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   679
proof safe
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   680
  fix y assume "x \<le> \<Squnion>A" "y < x"
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   681
  then have "y < \<Squnion>A" by auto
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   682
  then show "\<exists>a\<in>A. y < a"
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   683
    unfolding less_Sup_iff .
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   684
qed (auto elim!: allE[of _ "\<Squnion>A"] simp add: not_le[symmetric] Sup_upper)
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   685
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   686
lemma le_SUP_iff: "x \<le> SUPREMUM A f \<longleftrightarrow> (\<forall>y<x. \<exists>i\<in>A. y < f i)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   687
  using le_Sup_iff [of _ "f ` A"] by simp
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   688
51386
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   689
subclass complete_distrib_lattice
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   690
proof
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   691
  fix a and B
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   692
  show "a \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> b)" and "a \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)"
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   693
    by (safe intro!: INF_eqI [symmetric] sup_mono Inf_lower SUP_eqI [symmetric] inf_mono Sup_upper)
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   694
      (auto simp: not_less [symmetric] Inf_less_iff less_Sup_iff
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   695
        le_max_iff_disj complete_linorder_sup_max min_le_iff_disj complete_linorder_inf_min)
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   696
qed
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   697
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   698
end
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   699
51341
8c10293e7ea7 complete_linorder is also a complete_distrib_lattice
hoelzl
parents: 51328
diff changeset
   700
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   701
subsection \<open>Complete lattice on @{typ bool}\<close>
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   702
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   703
instantiation bool :: complete_lattice
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   704
begin
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   705
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   706
definition
46154
5115e47a7752 use Inf/Sup_bool_def/apply as code equations
haftmann
parents: 46036
diff changeset
   707
  [simp, code]: "\<Sqinter>A \<longleftrightarrow> False \<notin> A"
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   708
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   709
definition
46154
5115e47a7752 use Inf/Sup_bool_def/apply as code equations
haftmann
parents: 46036
diff changeset
   710
  [simp, code]: "\<Squnion>A \<longleftrightarrow> True \<in> A"
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   711
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   712
instance proof
44322
43b465f4c480 more concise definition for Inf, Sup on bool
haftmann
parents: 44104
diff changeset
   713
qed (auto intro: bool_induct)
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   714
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   715
end
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   716
49905
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   717
lemma not_False_in_image_Ball [simp]:
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   718
  "False \<notin> P ` A \<longleftrightarrow> Ball A P"
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   719
  by auto
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   720
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   721
lemma True_in_image_Bex [simp]:
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   722
  "True \<in> P ` A \<longleftrightarrow> Bex A P"
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   723
  by auto
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   724
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   725
lemma INF_bool_eq [simp]:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   726
  "INFIMUM = Ball"
49905
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   727
  by (simp add: fun_eq_iff INF_def)
32120
53a21a5e6889 attempt for more concise setup of non-etacontracting binders
haftmann
parents: 32117
diff changeset
   728
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   729
lemma SUP_bool_eq [simp]:
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   730
  "SUPREMUM = Bex"
49905
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   731
  by (simp add: fun_eq_iff SUP_def)
32120
53a21a5e6889 attempt for more concise setup of non-etacontracting binders
haftmann
parents: 32117
diff changeset
   732
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
   733
instance bool :: complete_boolean_algebra proof
44322
43b465f4c480 more concise definition for Inf, Sup on bool
haftmann
parents: 44104
diff changeset
   734
qed (auto intro: bool_induct)
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   735
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   736
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   737
subsection \<open>Complete lattice on @{typ "_ \<Rightarrow> _"}\<close>
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   738
57197
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   739
instantiation "fun" :: (type, Inf) Inf
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   740
begin
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   741
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   742
definition
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   743
  "\<Sqinter>A = (\<lambda>x. \<Sqinter>f\<in>A. f x)"
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   744
46882
6242b4bc05bc tuned simpset
noschinl
parents: 46693
diff changeset
   745
lemma Inf_apply [simp, code]:
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   746
  "(\<Sqinter>A) x = (\<Sqinter>f\<in>A. f x)"
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   747
  by (simp add: Inf_fun_def)
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   748
57197
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   749
instance ..
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   750
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   751
end
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   752
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   753
instantiation "fun" :: (type, Sup) Sup
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   754
begin
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   755
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   756
definition
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   757
  "\<Squnion>A = (\<lambda>x. \<Squnion>f\<in>A. f x)"
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   758
46882
6242b4bc05bc tuned simpset
noschinl
parents: 46693
diff changeset
   759
lemma Sup_apply [simp, code]:
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   760
  "(\<Squnion>A) x = (\<Squnion>f\<in>A. f x)"
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   761
  by (simp add: Sup_fun_def)
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   762
57197
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   763
instance ..
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   764
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   765
end
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   766
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   767
instantiation "fun" :: (type, complete_lattice) complete_lattice
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   768
begin
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   769
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   770
instance proof
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   771
qed (auto simp add: le_fun_def intro: INF_lower INF_greatest SUP_upper SUP_least)
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   772
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   773
end
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   774
46882
6242b4bc05bc tuned simpset
noschinl
parents: 46693
diff changeset
   775
lemma INF_apply [simp]:
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   776
  "(\<Sqinter>y\<in>A. f y) x = (\<Sqinter>y\<in>A. f y x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   777
  using Inf_apply [of "f ` A"] by (simp add: comp_def)
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   778
46882
6242b4bc05bc tuned simpset
noschinl
parents: 46693
diff changeset
   779
lemma SUP_apply [simp]:
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   780
  "(\<Squnion>y\<in>A. f y) x = (\<Squnion>y\<in>A. f y x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   781
  using Sup_apply [of "f ` A"] by (simp add: comp_def)
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   782
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   783
instance "fun" :: (type, complete_distrib_lattice) complete_distrib_lattice proof
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   784
qed (auto simp add: INF_def SUP_def inf_Sup sup_Inf fun_eq_iff image_image
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   785
  simp del: Inf_image_eq Sup_image_eq)
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   786
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   787
instance "fun" :: (type, complete_boolean_algebra) complete_boolean_algebra ..
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   788
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   789
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   790
subsection \<open>Complete lattice on unary and binary predicates\<close>
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   791
56742
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   792
lemma Inf1_I: 
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   793
  "(\<And>P. P \<in> A \<Longrightarrow> P a) \<Longrightarrow> (\<Sqinter>A) a"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   794
  by auto
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   795
56742
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   796
lemma INF1_I:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   797
  "(\<And>x. x \<in> A \<Longrightarrow> B x b) \<Longrightarrow> (\<Sqinter>x\<in>A. B x) b"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   798
  by simp
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   799
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   800
lemma INF2_I:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   801
  "(\<And>x. x \<in> A \<Longrightarrow> B x b c) \<Longrightarrow> (\<Sqinter>x\<in>A. B x) b c"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   802
  by simp
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   803
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   804
lemma Inf2_I: 
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   805
  "(\<And>r. r \<in> A \<Longrightarrow> r a b) \<Longrightarrow> (\<Sqinter>A) a b"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   806
  by auto
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   807
56742
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   808
lemma Inf1_D:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   809
  "(\<Sqinter>A) a \<Longrightarrow> P \<in> A \<Longrightarrow> P a"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   810
  by auto
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   811
56742
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   812
lemma INF1_D:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   813
  "(\<Sqinter>x\<in>A. B x) b \<Longrightarrow> a \<in> A \<Longrightarrow> B a b"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   814
  by simp
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   815
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   816
lemma Inf2_D:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   817
  "(\<Sqinter>A) a b \<Longrightarrow> r \<in> A \<Longrightarrow> r a b"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   818
  by auto
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   819
56742
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   820
lemma INF2_D:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   821
  "(\<Sqinter>x\<in>A. B x) b c \<Longrightarrow> a \<in> A \<Longrightarrow> B a b c"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   822
  by simp
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   823
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   824
lemma Inf1_E:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   825
  assumes "(\<Sqinter>A) a"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   826
  obtains "P a" | "P \<notin> A"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   827
  using assms by auto
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   828
56742
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   829
lemma INF1_E:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   830
  assumes "(\<Sqinter>x\<in>A. B x) b"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   831
  obtains "B a b" | "a \<notin> A"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   832
  using assms by auto
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   833
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   834
lemma Inf2_E:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   835
  assumes "(\<Sqinter>A) a b"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   836
  obtains "r a b" | "r \<notin> A"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   837
  using assms by auto
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   838
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   839
lemma INF2_E:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   840
  assumes "(\<Sqinter>x\<in>A. B x) b c"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   841
  obtains "B a b c" | "a \<notin> A"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   842
  using assms by auto
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   843
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   844
lemma Sup1_I:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   845
  "P \<in> A \<Longrightarrow> P a \<Longrightarrow> (\<Squnion>A) a"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   846
  by auto
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   847
56742
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   848
lemma SUP1_I:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   849
  "a \<in> A \<Longrightarrow> B a b \<Longrightarrow> (\<Squnion>x\<in>A. B x) b"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   850
  by auto
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   851
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   852
lemma Sup2_I:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   853
  "r \<in> A \<Longrightarrow> r a b \<Longrightarrow> (\<Squnion>A) a b"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   854
  by auto
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   855
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   856
lemma SUP2_I:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   857
  "a \<in> A \<Longrightarrow> B a b c \<Longrightarrow> (\<Squnion>x\<in>A. B x) b c"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   858
  by auto
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   859
56742
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   860
lemma Sup1_E:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   861
  assumes "(\<Squnion>A) a"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   862
  obtains P where "P \<in> A" and "P a"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   863
  using assms by auto
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   864
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   865
lemma SUP1_E:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   866
  assumes "(\<Squnion>x\<in>A. B x) b"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   867
  obtains x where "x \<in> A" and "B x b"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   868
  using assms by auto
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   869
56742
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   870
lemma Sup2_E:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   871
  assumes "(\<Squnion>A) a b"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   872
  obtains r where "r \<in> A" "r a b"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   873
  using assms by auto
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   874
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   875
lemma SUP2_E:
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   876
  assumes "(\<Squnion>x\<in>A. B x) b c"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   877
  obtains x where "x \<in> A" "B x b c"
678a52e676b6 more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents: 56741
diff changeset
   878
  using assms by auto
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   879
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   880
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   881
subsection \<open>Complete lattice on @{typ "_ set"}\<close>
46631
2c5c003cee35 moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
haftmann
parents: 46557
diff changeset
   882
45960
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   883
instantiation "set" :: (type) complete_lattice
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   884
begin
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   885
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   886
definition
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   887
  "\<Sqinter>A = {x. \<Sqinter>((\<lambda>B. x \<in> B) ` A)}"
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   888
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   889
definition
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   890
  "\<Squnion>A = {x. \<Squnion>((\<lambda>B. x \<in> B) ` A)}"
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   891
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   892
instance proof
51386
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   893
qed (auto simp add: less_eq_set_def Inf_set_def Sup_set_def le_fun_def)
45960
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   894
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   895
end
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   896
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   897
instance "set" :: (type) complete_boolean_algebra
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   898
proof
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   899
qed (auto simp add: INF_def SUP_def Inf_set_def Sup_set_def image_def)
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   900
  
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   901
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   902
subsubsection \<open>Inter\<close>
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   903
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   904
abbreviation Inter :: "'a set set \<Rightarrow> 'a set" where
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   905
  "Inter S \<equiv> \<Sqinter>S"
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   906
  
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   907
notation (xsymbols)
52141
eff000cab70f weaker precendence of syntax for big intersection and union on sets
haftmann
parents: 51540
diff changeset
   908
  Inter  ("\<Inter>_" [900] 900)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   909
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   910
lemma Inter_eq:
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   911
  "\<Inter>A = {x. \<forall>B \<in> A. x \<in> B}"
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   912
proof (rule set_eqI)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   913
  fix x
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   914
  have "(\<forall>Q\<in>{P. \<exists>B\<in>A. P \<longleftrightarrow> x \<in> B}. Q) \<longleftrightarrow> (\<forall>B\<in>A. x \<in> B)"
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   915
    by auto
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   916
  then show "x \<in> \<Inter>A \<longleftrightarrow> x \<in> {x. \<forall>B \<in> A. x \<in> B}"
45960
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   917
    by (simp add: Inf_set_def image_def)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   918
qed
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   919
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   920
lemma Inter_iff [simp]: "A \<in> \<Inter>C \<longleftrightarrow> (\<forall>X\<in>C. A \<in> X)"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   921
  by (unfold Inter_eq) blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   922
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   923
lemma InterI [intro!]: "(\<And>X. X \<in> C \<Longrightarrow> A \<in> X) \<Longrightarrow> A \<in> \<Inter>C"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   924
  by (simp add: Inter_eq)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   925
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   926
text \<open>
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   927
  \medskip A ``destruct'' rule -- every @{term X} in @{term C}
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   928
  contains @{term A} as an element, but @{prop "A \<in> X"} can hold when
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   929
  @{prop "X \<in> C"} does not!  This rule is analogous to @{text spec}.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   930
\<close>
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   931
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   932
lemma InterD [elim, Pure.elim]: "A \<in> \<Inter>C \<Longrightarrow> X \<in> C \<Longrightarrow> A \<in> X"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   933
  by auto
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   934
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   935
lemma InterE [elim]: "A \<in> \<Inter>C \<Longrightarrow> (X \<notin> C \<Longrightarrow> R) \<Longrightarrow> (A \<in> X \<Longrightarrow> R) \<Longrightarrow> R"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   936
  -- \<open>``Classical'' elimination rule -- does not require proving
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   937
    @{prop "X \<in> C"}.\<close>
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   938
  by (unfold Inter_eq) blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   939
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   940
lemma Inter_lower: "B \<in> A \<Longrightarrow> \<Inter>A \<subseteq> B"
43740
3316e6831801 more succinct proofs
haftmann
parents: 43739
diff changeset
   941
  by (fact Inf_lower)
3316e6831801 more succinct proofs
haftmann
parents: 43739
diff changeset
   942
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   943
lemma Inter_subset:
43755
5e4a595e63fc tuned notation
haftmann
parents: 43754
diff changeset
   944
  "(\<And>X. X \<in> A \<Longrightarrow> X \<subseteq> B) \<Longrightarrow> A \<noteq> {} \<Longrightarrow> \<Inter>A \<subseteq> B"
43740
3316e6831801 more succinct proofs
haftmann
parents: 43739
diff changeset
   945
  by (fact Inf_less_eq)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   946
43755
5e4a595e63fc tuned notation
haftmann
parents: 43754
diff changeset
   947
lemma Inter_greatest: "(\<And>X. X \<in> A \<Longrightarrow> C \<subseteq> X) \<Longrightarrow> C \<subseteq> Inter A"
43740
3316e6831801 more succinct proofs
haftmann
parents: 43739
diff changeset
   948
  by (fact Inf_greatest)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   949
44067
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   950
lemma Inter_empty: "\<Inter>{} = UNIV"
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   951
  by (fact Inf_empty) (* already simp *)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   952
44067
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   953
lemma Inter_UNIV: "\<Inter>UNIV = {}"
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   954
  by (fact Inf_UNIV) (* already simp *)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   955
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   956
lemma Inter_insert: "\<Inter>(insert a B) = a \<inter> \<Inter>B"
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   957
  by (fact Inf_insert) (* already simp *)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   958
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   959
lemma Inter_Un_subset: "\<Inter>A \<union> \<Inter>B \<subseteq> \<Inter>(A \<inter> B)"
43899
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   960
  by (fact less_eq_Inf_inter)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   961
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   962
lemma Inter_Un_distrib: "\<Inter>(A \<union> B) = \<Inter>A \<inter> \<Inter>B"
43756
15ea1a07a8cf tuned proofs
haftmann
parents: 43755
diff changeset
   963
  by (fact Inf_union_distrib)
15ea1a07a8cf tuned proofs
haftmann
parents: 43755
diff changeset
   964
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   965
lemma Inter_UNIV_conv [simp]:
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   966
  "\<Inter>A = UNIV \<longleftrightarrow> (\<forall>x\<in>A. x = UNIV)"
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   967
  "UNIV = \<Inter>A \<longleftrightarrow> (\<forall>x\<in>A. x = UNIV)"
43801
097732301fc0 more generalization towards complete lattices
haftmann
parents: 43756
diff changeset
   968
  by (fact Inf_top_conv)+
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   969
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   970
lemma Inter_anti_mono: "B \<subseteq> A \<Longrightarrow> \<Inter>A \<subseteq> \<Inter>B"
43899
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   971
  by (fact Inf_superset_mono)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   972
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   973
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   974
subsubsection \<open>Intersections of families\<close>
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   975
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   976
abbreviation INTER :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
   977
  "INTER \<equiv> INFIMUM"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   978
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   979
text \<open>
43872
6b917e5877d2 more consistent theorem names
haftmann
parents: 43871
diff changeset
   980
  Note: must use name @{const INTER} here instead of @{text INT}
6b917e5877d2 more consistent theorem names
haftmann
parents: 43871
diff changeset
   981
  to allow the following syntax coexist with the plain constant name.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   982
\<close>
43872
6b917e5877d2 more consistent theorem names
haftmann
parents: 43871
diff changeset
   983
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   984
syntax
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   985
  "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3INT _./ _)" [0, 10] 10)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   986
  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3INT _:_./ _)" [0, 0, 10] 10)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   987
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   988
syntax (xsymbols)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   989
  "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>_./ _)" [0, 10] 10)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   990
  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>_\<in>_./ _)" [0, 0, 10] 10)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   991
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   992
syntax (latex output)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   993
  "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   994
  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   995
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   996
translations
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   997
  "INT x y. B"  == "INT x. INT y. B"
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   998
  "INT x. B"    == "CONST INTER CONST UNIV (%x. B)"
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   999
  "INT x. B"    == "INT x:CONST UNIV. B"
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1000
  "INT x:A. B"  == "CONST INTER A (%x. B)"
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1001
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1002
print_translation \<open>
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 41971
diff changeset
  1003
  [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INTER} @{syntax_const "_INTER"}]
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1004
\<close> -- \<open>to avoid eta-contraction of body\<close>
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1005
44085
a65e26f1427b move legacy candiates to bottom; marked candidates for default simp rules
haftmann
parents: 44084
diff changeset
  1006
lemma INTER_eq:
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1007
  "(\<Inter>x\<in>A. B x) = {y. \<forall>x\<in>A. y \<in> B x}"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1008
  by (auto intro!: INF_eqI)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1009
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1010
lemma Inter_image_eq:
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1011
  "\<Inter>(B ` A) = (\<Inter>x\<in>A. B x)"
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1012
  by (fact Inf_image_eq)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1013
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1014
lemma INT_iff [simp]: "b \<in> (\<Inter>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. b \<in> B x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1015
  using Inter_iff [of _ "B ` A"] by simp
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1016
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1017
lemma INT_I [intro!]: "(\<And>x. x \<in> A \<Longrightarrow> b \<in> B x) \<Longrightarrow> b \<in> (\<Inter>x\<in>A. B x)"
44085
a65e26f1427b move legacy candiates to bottom; marked candidates for default simp rules
haftmann
parents: 44084
diff changeset
  1018
  by (auto simp add: INF_def image_def)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1019
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1020
lemma INT_D [elim, Pure.elim]: "b \<in> (\<Inter>x\<in>A. B x) \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> B a"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1021
  by auto
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1022
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1023
lemma INT_E [elim]: "b \<in> (\<Inter>x\<in>A. B x) \<Longrightarrow> (b \<in> B a \<Longrightarrow> R) \<Longrightarrow> (a \<notin> A \<Longrightarrow> R) \<Longrightarrow> R"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1024
  -- \<open>"Classical" elimination -- by the Excluded Middle on @{prop "a\<in>A"}.\<close>
44085
a65e26f1427b move legacy candiates to bottom; marked candidates for default simp rules
haftmann
parents: 44084
diff changeset
  1025
  by (auto simp add: INF_def image_def)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1026
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1027
lemma Collect_ball_eq: "{x. \<forall>y\<in>A. P x y} = (\<Inter>y\<in>A. {x. P x y})"
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1028
  by blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1029
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1030
lemma Collect_all_eq: "{x. \<forall>y. P x y} = (\<Inter>y. {x. P x y})"
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1031
  by blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1032
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1033
lemma INT_lower: "a \<in> A \<Longrightarrow> (\<Inter>x\<in>A. B x) \<subseteq> B a"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
  1034
  by (fact INF_lower)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1035
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1036
lemma INT_greatest: "(\<And>x. x \<in> A \<Longrightarrow> C \<subseteq> B x) \<Longrightarrow> C \<subseteq> (\<Inter>x\<in>A. B x)"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
  1037
  by (fact INF_greatest)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1038
44067
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
  1039
lemma INT_empty: "(\<Inter>x\<in>{}. B x) = UNIV"
44085
a65e26f1427b move legacy candiates to bottom; marked candidates for default simp rules
haftmann
parents: 44084
diff changeset
  1040
  by (fact INF_empty)
43854
f1d23df1adde generalized some lemmas
haftmann
parents: 43853
diff changeset
  1041
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1042
lemma INT_absorb: "k \<in> I \<Longrightarrow> A k \<inter> (\<Inter>i\<in>I. A i) = (\<Inter>i\<in>I. A i)"
43872
6b917e5877d2 more consistent theorem names
haftmann
parents: 43871
diff changeset
  1043
  by (fact INF_absorb)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1044
43854
f1d23df1adde generalized some lemmas
haftmann
parents: 43853
diff changeset
  1045
lemma INT_subset_iff: "B \<subseteq> (\<Inter>i\<in>I. A i) \<longleftrightarrow> (\<forall>i\<in>I. B \<subseteq> A i)"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1046
  by (fact le_INF_iff)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1047
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1048
lemma INT_insert [simp]: "(\<Inter>x \<in> insert a A. B x) = B a \<inter> INTER A B"
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1049
  by (fact INF_insert)
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1050
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1051
lemma INT_Un: "(\<Inter>i \<in> A \<union> B. M i) = (\<Inter>i \<in> A. M i) \<inter> (\<Inter>i\<in>B. M i)"
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1052
  by (fact INF_union)
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1053
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1054
lemma INT_insert_distrib:
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1055
  "u \<in> A \<Longrightarrow> (\<Inter>x\<in>A. insert a (B x)) = insert a (\<Inter>x\<in>A. B x)"
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1056
  by blast
43854
f1d23df1adde generalized some lemmas
haftmann
parents: 43853
diff changeset
  1057
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1058
lemma INT_constant [simp]: "(\<Inter>y\<in>A. c) = (if A = {} then UNIV else c)"
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1059
  by (fact INF_constant)
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1060
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1061
lemma INTER_UNIV_conv:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1062
 "(UNIV = (\<Inter>x\<in>A. B x)) = (\<forall>x\<in>A. B x = UNIV)"
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1063
 "((\<Inter>x\<in>A. B x) = UNIV) = (\<forall>x\<in>A. B x = UNIV)"
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1064
  by (fact INF_top_conv)+ (* already simp *)
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1065
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1066
lemma INT_bool_eq: "(\<Inter>b. A b) = A True \<inter> A False"
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
  1067
  by (fact INF_UNIV_bool_expand)
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1068
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
  1069
lemma INT_anti_mono:
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
  1070
  "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<subseteq> g x) \<Longrightarrow> (\<Inter>x\<in>B. f x) \<subseteq> (\<Inter>x\<in>A. g x)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1071
  -- \<open>The last inclusion is POSITIVE!\<close>
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
  1072
  by (fact INF_superset_mono)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1073
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1074
lemma Pow_INT_eq: "Pow (\<Inter>x\<in>A. B x) = (\<Inter>x\<in>A. Pow (B x))"
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1075
  by blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1076
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1077
lemma vimage_INT: "f -` (\<Inter>x\<in>A. B x) = (\<Inter>x\<in>A. f -` B x)"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1078
  by blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1079
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
  1080
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1081
subsubsection \<open>Union\<close>
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1082
32587
caa5ada96a00 Inter and Union are mere abbreviations for Inf and Sup
haftmann
parents: 32436
diff changeset
  1083
abbreviation Union :: "'a set set \<Rightarrow> 'a set" where
caa5ada96a00 Inter and Union are mere abbreviations for Inf and Sup
haftmann
parents: 32436
diff changeset
  1084
  "Union S \<equiv> \<Squnion>S"
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1085
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1086
notation (xsymbols)
52141
eff000cab70f weaker precendence of syntax for big intersection and union on sets
haftmann
parents: 51540
diff changeset
  1087
  Union  ("\<Union>_" [900] 900)
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1088
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1089
lemma Union_eq:
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1090
  "\<Union>A = {x. \<exists>B \<in> A. x \<in> B}"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 38705
diff changeset
  1091
proof (rule set_eqI)
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1092
  fix x
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1093
  have "(\<exists>Q\<in>{P. \<exists>B\<in>A. P \<longleftrightarrow> x \<in> B}. Q) \<longleftrightarrow> (\<exists>B\<in>A. x \<in> B)"
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1094
    by auto
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1095
  then show "x \<in> \<Union>A \<longleftrightarrow> x \<in> {x. \<exists>B\<in>A. x \<in> B}"
45960
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
  1096
    by (simp add: Sup_set_def image_def)
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1097
qed
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1098
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1099
lemma Union_iff [simp]:
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1100
  "A \<in> \<Union>C \<longleftrightarrow> (\<exists>X\<in>C. A\<in>X)"
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1101
  by (unfold Union_eq) blast
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1102
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1103
lemma UnionI [intro]:
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1104
  "X \<in> C \<Longrightarrow> A \<in> X \<Longrightarrow> A \<in> \<Union>C"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1105
  -- \<open>The order of the premises presupposes that @{term C} is rigid;
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1106
    @{term A} may be flexible.\<close>
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1107
  by auto
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1108
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1109
lemma UnionE [elim!]:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1110
  "A \<in> \<Union>C \<Longrightarrow> (\<And>X. A \<in> X \<Longrightarrow> X \<in> C \<Longrightarrow> R) \<Longrightarrow> R"
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1111
  by auto
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1112
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1113
lemma Union_upper: "B \<in> A \<Longrightarrow> B \<subseteq> \<Union>A"
43901
3ab6c30d256d proof tuning
haftmann
parents: 43900
diff changeset
  1114
  by (fact Sup_upper)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1115
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1116
lemma Union_least: "(\<And>X. X \<in> A \<Longrightarrow> X \<subseteq> C) \<Longrightarrow> \<Union>A \<subseteq> C"
43901
3ab6c30d256d proof tuning
haftmann
parents: 43900
diff changeset
  1117
  by (fact Sup_least)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1118
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1119
lemma Union_empty: "\<Union>{} = {}"
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1120
  by (fact Sup_empty) (* already simp *)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1121
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1122
lemma Union_UNIV: "\<Union>UNIV = UNIV"
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1123
  by (fact Sup_UNIV) (* already simp *)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1124
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1125
lemma Union_insert: "\<Union>insert a B = a \<union> \<Union>B"
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1126
  by (fact Sup_insert) (* already simp *)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1127
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1128
lemma Union_Un_distrib [simp]: "\<Union>(A \<union> B) = \<Union>A \<union> \<Union>B"
43901
3ab6c30d256d proof tuning
haftmann
parents: 43900
diff changeset
  1129
  by (fact Sup_union_distrib)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1130
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1131
lemma Union_Int_subset: "\<Union>(A \<inter> B) \<subseteq> \<Union>A \<inter> \<Union>B"
43901
3ab6c30d256d proof tuning
haftmann
parents: 43900
diff changeset
  1132
  by (fact Sup_inter_less_eq)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1133
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1134
lemma Union_empty_conv: "(\<Union>A = {}) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1135
  by (fact Sup_bot_conv) (* already simp *)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1136
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1137
lemma empty_Union_conv: "({} = \<Union>A) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1138
  by (fact Sup_bot_conv) (* already simp *)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1139
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1140
lemma subset_Pow_Union: "A \<subseteq> Pow (\<Union>A)"
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1141
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1142
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1143
lemma Union_Pow_eq [simp]: "\<Union>(Pow A) = A"
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1144
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1145
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1146
lemma Union_mono: "A \<subseteq> B \<Longrightarrow> \<Union>A \<subseteq> \<Union>B"
43901
3ab6c30d256d proof tuning
haftmann
parents: 43900
diff changeset
  1147
  by (fact Sup_subset_mono)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1148
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1149
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1150
subsubsection \<open>Unions of families\<close>
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1151
32606
b5c3a8a75772 INTER and UNION are mere abbreviations for INFI and SUPR
haftmann
parents: 32587
diff changeset
  1152
abbreviation UNION :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56212
diff changeset
  1153
  "UNION \<equiv> SUPREMUM"
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1154
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1155
text \<open>
43872
6b917e5877d2 more consistent theorem names
haftmann
parents: 43871
diff changeset
  1156
  Note: must use name @{const UNION} here instead of @{text UN}
6b917e5877d2 more consistent theorem names
haftmann
parents: 43871
diff changeset
  1157
  to allow the following syntax coexist with the plain constant name.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1158
\<close>
43872
6b917e5877d2 more consistent theorem names
haftmann
parents: 43871
diff changeset
  1159
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1160
syntax
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34007
diff changeset
  1161
  "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3UN _./ _)" [0, 10] 10)
36364
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 35828
diff changeset
  1162
  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3UN _:_./ _)" [0, 0, 10] 10)
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1163
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1164
syntax (xsymbols)
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34007
diff changeset
  1165
  "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>_./ _)" [0, 10] 10)
36364
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 35828
diff changeset
  1166
  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>_\<in>_./ _)" [0, 0, 10] 10)
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1167
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1168
syntax (latex output)
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34007
diff changeset
  1169
  "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
36364
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 35828
diff changeset
  1170
  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1171
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1172
translations
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1173
  "UN x y. B"   == "UN x. UN y. B"
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1174
  "UN x. B"     == "CONST UNION CONST UNIV (%x. B)"
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1175
  "UN x. B"     == "UN x:CONST UNIV. B"
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1176
  "UN x:A. B"   == "CONST UNION A (%x. B)"
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1177
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1178
text \<open>
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1179
  Note the difference between ordinary xsymbol syntax of indexed
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
  1180
  unions and intersections (e.g.\ @{text"\<Union>a\<^sub>1\<in>A\<^sub>1. B"})
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
  1181
  and their \LaTeX\ rendition: @{term"\<Union>a\<^sub>1\<in>A\<^sub>1. B"}. The
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1182
  former does not make the index expression a subscript of the
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1183
  union/intersection symbol because this leads to problems with nested
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1184
  subscripts in Proof General.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1185
\<close>
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1186
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1187
print_translation \<open>
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 41971
diff changeset
  1188
  [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax UNION} @{syntax_const "_UNION"}]
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1189
\<close> -- \<open>to avoid eta-contraction of body\<close>
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1190
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1191
lemma UNION_eq:
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1192
  "(\<Union>x\<in>A. B x) = {y. \<exists>x\<in>A. y \<in> B x}"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1193
  by (auto intro!: SUP_eqI)
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1194
45960
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
  1195
lemma bind_UNION [code]:
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
  1196
  "Set.bind A f = UNION A f"
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
  1197
  by (simp add: bind_def UNION_eq)
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
  1198
46036
6a86cc88b02f fundamental theorems on Set.bind
haftmann
parents: 45960
diff changeset
  1199
lemma member_bind [simp]:
6a86cc88b02f fundamental theorems on Set.bind
haftmann
parents: 45960
diff changeset
  1200
  "x \<in> Set.bind P f \<longleftrightarrow> x \<in> UNION P f "
6a86cc88b02f fundamental theorems on Set.bind
haftmann
parents: 45960
diff changeset
  1201
  by (simp add: bind_UNION)
6a86cc88b02f fundamental theorems on Set.bind
haftmann
parents: 45960
diff changeset
  1202
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1203
lemma Union_image_eq:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1204
  "\<Union>(B ` A) = (\<Union>x\<in>A. B x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1205
  by (fact Sup_image_eq)
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1206
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60307
diff changeset
  1207
lemma Union_SetCompr_eq: "\<Union>{f x| x. P x} = {a. \<exists>x. P x \<and> a \<in> f x}"
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60172
diff changeset
  1208
  by blast
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60172
diff changeset
  1209
46036
6a86cc88b02f fundamental theorems on Set.bind
haftmann
parents: 45960
diff changeset
  1210
lemma UN_iff [simp]: "b \<in> (\<Union>x\<in>A. B x) \<longleftrightarrow> (\<exists>x\<in>A. b \<in> B x)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1211
  using Union_iff [of _ "B ` A"] by simp
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1212
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1213
lemma UN_I [intro]: "a \<in> A \<Longrightarrow> b \<in> B a \<Longrightarrow> b \<in> (\<Union>x\<in>A. B x)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1214
  -- \<open>The order of the premises presupposes that @{term A} is rigid;
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1215
    @{term b} may be flexible.\<close>
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1216
  by auto
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1217
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1218
lemma UN_E [elim!]: "b \<in> (\<Union>x\<in>A. B x) \<Longrightarrow> (\<And>x. x\<in>A \<Longrightarrow> b \<in> B x \<Longrightarrow> R) \<Longrightarrow> R"
44085
a65e26f1427b move legacy candiates to bottom; marked candidates for default simp rules
haftmann
parents: 44084
diff changeset
  1219
  by (auto simp add: SUP_def image_def)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
  1220
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1221
lemma image_eq_UN: "f ` A = (\<Union>x\<in>A. {f x})"
32077
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1222
  by blast
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1223
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1224
lemma UN_upper: "a \<in> A \<Longrightarrow> B a \<subseteq> (\<Union>x\<in>A. B x)"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
  1225
  by (fact SUP_upper)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1226
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1227
lemma UN_least: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C) \<Longrightarrow> (\<Union>x\<in>A. B x) \<subseteq> C"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
  1228
  by (fact SUP_least)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1229
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1230
lemma Collect_bex_eq: "{x. \<exists>y\<in>A. P x y} = (\<Union>y\<in>A. {x. P x y})"
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1231
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1232
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1233
lemma UN_insert_distrib: "u \<in> A \<Longrightarrow> (\<Union>x\<in>A. insert a (B x)) = insert a (\<Union>x\<in>A. B x)"
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1234
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1235
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1236
lemma UN_empty: "(\<Union>x\<in>{}. B x) = {}"
44085
a65e26f1427b move legacy candiates to bottom; marked candidates for default simp rules
haftmann
parents: 44084
diff changeset
  1237
  by (fact SUP_empty)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1238
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1239
lemma UN_empty2: "(\<Union>x\<in>A. {}) = {}"
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1240
  by (fact SUP_bot) (* already simp *)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1241
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1242
lemma UN_absorb: "k \<in> I \<Longrightarrow> A k \<union> (\<Union>i\<in>I. A i) = (\<Union>i\<in>I. A i)"
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
  1243
  by (fact SUP_absorb)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1244
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1245
lemma UN_insert [simp]: "(\<Union>x\<in>insert a A. B x) = B a \<union> UNION A B"
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
  1246
  by (fact SUP_insert)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1247
44085
a65e26f1427b move legacy candiates to bottom; marked candidates for default simp rules
haftmann
parents: 44084
diff changeset
  1248
lemma UN_Un [simp]: "(\<Union>i \<in> A \<union> B. M i) = (\<Union>i\<in>A. M i) \<union> (\<Union>i\<in>B. M i)"
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
  1249
  by (fact SUP_union)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1250
43967
610efb6bda1f more coherent structure in and across theories
haftmann
parents: 43944
diff changeset
  1251
lemma UN_UN_flatten: "(\<Union>x \<in> (\<Union>y\<in>A. B y). C x) = (\<Union>y\<in>A. \<Union>x\<in>B y. C x)"
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1252
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1253
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1254
lemma UN_subset_iff: "((\<Union>i\<in>I. A i) \<subseteq> B) = (\<forall>i\<in>I. A i \<subseteq> B)"
35629
57f1a5e93b6b add some lemmas about complete lattices
huffman
parents: 35115
diff changeset
  1255
  by (fact SUP_le_iff)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1256
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1257
lemma UN_constant [simp]: "(\<Union>y\<in>A. c) = (if A = {} then {} else c)"
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
  1258
  by (fact SUP_constant)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1259
43944
b1b436f75070 dropped errorneous hint
haftmann
parents: 43943
diff changeset
  1260
lemma image_Union: "f ` \<Union>S = (\<Union>x\<in>S. f ` x)"
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1261
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1262
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1263
lemma UNION_empty_conv:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1264
  "{} = (\<Union>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = {})"
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1265
  "(\<Union>x\<in>A. B x) = {} \<longleftrightarrow> (\<forall>x\<in>A. B x = {})"
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1266
  by (fact SUP_bot_conv)+ (* already simp *)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1267
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1268
lemma Collect_ex_eq: "{x. \<exists>y. P x y} = (\<Union>y. {x. P x y})"
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1269
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1270
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
  1271
lemma ball_UN: "(\<forall>z \<in> UNION A B. P z) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>z \<in> B x. P z)"
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1272
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1273
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
  1274
lemma bex_UN: "(\<exists>z \<in> UNION A B. P z) \<longleftrightarrow> (\<exists>x\<in>A. \<exists>z\<in>B x. P z)"
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1275
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1276
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1277
lemma Un_eq_UN: "A \<union> B = (\<Union>b. if b then A else B)"
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1278
  by (auto simp add: split_if_mem2)
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1279
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1280
lemma UN_bool_eq: "(\<Union>b. A b) = (A True \<union> A False)"
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
  1281
  by (fact SUP_UNIV_bool_expand)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1282
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1283
lemma UN_Pow_subset: "(\<Union>x\<in>A. Pow (B x)) \<subseteq> Pow (\<Union>x\<in>A. B x)"
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1284
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1285
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1286
lemma UN_mono:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1287
  "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<subseteq> g x) \<Longrightarrow>
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1288
    (\<Union>x\<in>A. f x) \<subseteq> (\<Union>x\<in>B. g x)"
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
  1289
  by (fact SUP_subset_mono)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1290
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1291
lemma vimage_Union: "f -` (\<Union>A) = (\<Union>X\<in>A. f -` X)"
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1292
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1293
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1294
lemma vimage_UN: "f -` (\<Union>x\<in>A. B x) = (\<Union>x\<in>A. f -` B x)"
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1295
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1296
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1297
lemma vimage_eq_UN: "f -` B = (\<Union>y\<in>B. f -` {y})"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1298
  -- \<open>NOT suitable for rewriting\<close>
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1299
  by blast
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1300
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1301
lemma image_UN: "f ` UNION A B = (\<Union>x\<in>A. f ` B x)"
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1302
  by blast
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1303
45013
05031b71a89a official status for UN_singleton
haftmann
parents: 44995
diff changeset
  1304
lemma UN_singleton [simp]: "(\<Union>x\<in>A. {x}) = A"
05031b71a89a official status for UN_singleton
haftmann
parents: 44995
diff changeset
  1305
  by blast
05031b71a89a official status for UN_singleton
haftmann
parents: 44995
diff changeset
  1306
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1307
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1308
subsubsection \<open>Distributive laws\<close>
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1309
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1310
lemma Int_Union: "A \<inter> \<Union>B = (\<Union>C\<in>B. A \<inter> C)"
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
  1311
  by (fact inf_Sup)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1312
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1313
lemma Un_Inter: "A \<union> \<Inter>B = (\<Inter>C\<in>B. A \<union> C)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1314
  by (fact sup_Inf)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1315
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1316
lemma Int_Union2: "\<Union>B \<inter> A = (\<Union>C\<in>B. C \<inter> A)"
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1317
  by (fact Sup_inf)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1318
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1319
lemma INT_Int_distrib: "(\<Inter>i\<in>I. A i \<inter> B i) = (\<Inter>i\<in>I. A i) \<inter> (\<Inter>i\<in>I. B i)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1320
  by (rule sym) (rule INF_inf_distrib)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1321
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1322
lemma UN_Un_distrib: "(\<Union>i\<in>I. A i \<union> B i) = (\<Union>i\<in>I. A i) \<union> (\<Union>i\<in>I. B i)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1323
  by (rule sym) (rule SUP_sup_distrib)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1324
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1325
lemma Int_Inter_image: "(\<Inter>x\<in>C. A x \<inter> B x) = \<Inter>(A ` C) \<inter> \<Inter>(B ` C)" -- \<open>FIXME drop\<close>
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1326
  by (simp add: INT_Int_distrib)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1327
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1328
lemma Un_Union_image: "(\<Union>x\<in>C. A x \<union> B x) = \<Union>(A ` C) \<union> \<Union>(B ` C)" -- \<open>FIXME drop\<close>
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1329
  -- \<open>Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5:\<close>
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1330
  -- \<open>Union of a family of unions\<close>
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1331
  by (simp add: UN_Un_distrib)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1332
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1333
lemma Un_INT_distrib: "B \<union> (\<Inter>i\<in>I. A i) = (\<Inter>i\<in>I. B \<union> A i)"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1334
  by (fact sup_INF)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1335
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1336
lemma Int_UN_distrib: "B \<inter> (\<Union>i\<in>I. A i) = (\<Union>i\<in>I. B \<inter> A i)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1337
  -- \<open>Halmos, Naive Set Theory, page 35.\<close>
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1338
  by (fact inf_SUP)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1339
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1340
lemma Int_UN_distrib2: "(\<Union>i\<in>I. A i) \<inter> (\<Union>j\<in>J. B j) = (\<Union>i\<in>I. \<Union>j\<in>J. A i \<inter> B j)"
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1341
  by (fact SUP_inf_distrib2)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1342
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1343
lemma Un_INT_distrib2: "(\<Inter>i\<in>I. A i) \<union> (\<Inter>j\<in>J. B j) = (\<Inter>i\<in>I. \<Inter>j\<in>J. A i \<union> B j)"
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1344
  by (fact INF_sup_distrib2)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1345
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1346
lemma Union_disjoint: "(\<Union>C \<inter> A = {}) \<longleftrightarrow> (\<forall>B\<in>C. B \<inter> A = {})"
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1347
  by (fact Sup_inf_eq_bot_iff)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1348
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1349
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1350
subsection \<open>Injections and bijections\<close>
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1351
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1352
lemma inj_on_Inter:
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1353
  "S \<noteq> {} \<Longrightarrow> (\<And>A. A \<in> S \<Longrightarrow> inj_on f A) \<Longrightarrow> inj_on f (\<Inter>S)"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1354
  unfolding inj_on_def by blast
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1355
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1356
lemma inj_on_INTER:
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1357
  "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> inj_on f (A i)) \<Longrightarrow> inj_on f (\<Inter>i \<in> I. A i)"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1358
  unfolding inj_on_def by blast
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1359
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1360
lemma inj_on_UNION_chain:
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1361
  assumes CH: "\<And> i j. \<lbrakk>i \<in> I; j \<in> I\<rbrakk> \<Longrightarrow> A i \<le> A j \<or> A j \<le> A i" and
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1362
         INJ: "\<And> i. i \<in> I \<Longrightarrow> inj_on f (A i)"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60307
diff changeset
  1363
  shows "inj_on f (\<Union>i \<in> I. A i)"
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1364
proof -
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1365
  {
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1366
    fix i j x y
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1367
    assume *: "i \<in> I" "j \<in> I" and **: "x \<in> A i" "y \<in> A j"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1368
      and ***: "f x = f y"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1369
    have "x = y"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1370
    proof -
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1371
      {
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1372
        assume "A i \<le> A j"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1373
        with ** have "x \<in> A j" by auto
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1374
        with INJ * ** *** have ?thesis
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1375
        by(auto simp add: inj_on_def)
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1376
      }
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1377
      moreover
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1378
      {
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1379
        assume "A j \<le> A i"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1380
        with ** have "y \<in> A i" by auto
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1381
        with INJ * ** *** have ?thesis
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1382
        by(auto simp add: inj_on_def)
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1383
      }
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1384
      ultimately show ?thesis using CH * by blast
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1385
    qed
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1386
  }
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1387
  then show ?thesis by (unfold inj_on_def UNION_eq) auto
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1388
qed
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1389
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1390
lemma bij_betw_UNION_chain:
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1391
  assumes CH: "\<And> i j. \<lbrakk>i \<in> I; j \<in> I\<rbrakk> \<Longrightarrow> A i \<le> A j \<or> A j \<le> A i" and
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1392
         BIJ: "\<And> i. i \<in> I \<Longrightarrow> bij_betw f (A i) (A' i)"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60307
diff changeset
  1393
  shows "bij_betw f (\<Union>i \<in> I. A i) (\<Union>i \<in> I. A' i)"
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1394
proof (unfold bij_betw_def, auto)
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1395
  have "\<And> i. i \<in> I \<Longrightarrow> inj_on f (A i)"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1396
  using BIJ bij_betw_def[of f] by auto
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60307
diff changeset
  1397
  thus "inj_on f (\<Union>i \<in> I. A i)"
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1398
  using CH inj_on_UNION_chain[of I A f] by auto
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1399
next
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1400
  fix i x
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1401
  assume *: "i \<in> I" "x \<in> A i"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1402
  hence "f x \<in> A' i" using BIJ bij_betw_def[of f] by auto
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1403
  thus "\<exists>j \<in> I. f x \<in> A' j" using * by blast
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1404
next
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1405
  fix i x'
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1406
  assume *: "i \<in> I" "x' \<in> A' i"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1407
  hence "\<exists>x \<in> A i. x' = f x" using BIJ bij_betw_def[of f] by blast
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1408
  then have "\<exists>j \<in> I. \<exists>x \<in> A j. x' = f x"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1409
    using * by blast
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1410
  then show "x' \<in> f ` (\<Union>x\<in>I. A x)" by blast
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1411
qed
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1412
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1413
(*injectivity's required.  Left-to-right inclusion holds even if A is empty*)
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1414
lemma image_INT:
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1415
   "[| inj_on f C;  ALL x:A. B x <= C;  j:A |]
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1416
    ==> f ` (INTER A B) = (INT x:A. f ` B x)"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1417
apply (simp add: inj_on_def, blast)
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1418
done
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1419
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1420
(*Compare with image_INT: no use of inj_on, and if f is surjective then
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1421
  it doesn't matter whether A is empty*)
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1422
lemma bij_image_INT: "bij f ==> f ` (INTER A B) = (INT x:A. f ` B x)"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1423
apply (simp add: bij_def)
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1424
apply (simp add: inj_on_def surj_def, blast)
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1425
done
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1426
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1427
lemma UNION_fun_upd:
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1428
  "UNION J (A(i:=B)) = (UNION (J-{i}) A \<union> (if i\<in>J then B else {}))"
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1429
by (auto split: if_splits)
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1430
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1431
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1432
subsubsection \<open>Complement\<close>
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1433
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
  1434
lemma Compl_INT [simp]: "- (\<Inter>x\<in>A. B x) = (\<Union>x\<in>A. -B x)"
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
  1435
  by (fact uminus_INF)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1436
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
  1437
lemma Compl_UN [simp]: "- (\<Union>x\<in>A. B x) = (\<Inter>x\<in>A. -B x)"
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
  1438
  by (fact uminus_SUP)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1439
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1440
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1441
subsubsection \<open>Miniscoping and maxiscoping\<close>
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1442
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1443
text \<open>\medskip Miniscoping: pushing in quantifiers and big Unions
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1444
           and Intersections.\<close>
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1445
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1446
lemma UN_simps [simp]:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1447
  "\<And>a B C. (\<Union>x\<in>C. insert a (B x)) = (if C={} then {} else insert a (\<Union>x\<in>C. B x))"
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
  1448
  "\<And>A B C. (\<Union>x\<in>C. A x \<union> B) = ((if C={} then {} else (\<Union>x\<in>C. A x) \<union> B))"
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1449
  "\<And>A B C. (\<Union>x\<in>C. A \<union> B x) = ((if C={} then {} else A \<union> (\<Union>x\<in>C. B x)))"
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
  1450
  "\<And>A B C. (\<Union>x\<in>C. A x \<inter> B) = ((\<Union>x\<in>C. A x) \<inter> B)"
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1451
  "\<And>A B C. (\<Union>x\<in>C. A \<inter> B x) = (A \<inter>(\<Union>x\<in>C. B x))"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1452
  "\<And>A B C. (\<Union>x\<in>C. A x - B) = ((\<Union>x\<in>C. A x) - B)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1453
  "\<And>A B C. (\<Union>x\<in>C. A - B x) = (A - (\<Inter>x\<in>C. B x))"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1454
  "\<And>A B. (\<Union>x\<in>\<Union>A. B x) = (\<Union>y\<in>A. \<Union>x\<in>y. B x)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1455
  "\<And>A B C. (\<Union>z\<in>UNION A B. C z) = (\<Union>x\<in>A. \<Union>z\<in>B x. C z)"
43831
e323be6b02a5 tuned notation and proofs
haftmann
parents: 43818
diff changeset
  1456
  "\<And>A B f. (\<Union>x\<in>f`A. B x) = (\<Union>a\<in>A. B (f a))"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1457
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1458
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1459
lemma INT_simps [simp]:
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
  1460
  "\<And>A B C. (\<Inter>x\<in>C. A x \<inter> B) = (if C={} then UNIV else (\<Inter>x\<in>C. A x) \<inter> B)"
43831
e323be6b02a5 tuned notation and proofs
haftmann
parents: 43818
diff changeset
  1461
  "\<And>A B C. (\<Inter>x\<in>C. A \<inter> B x) = (if C={} then UNIV else A \<inter>(\<Inter>x\<in>C. B x))"
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1462
  "\<And>A B C. (\<Inter>x\<in>C. A x - B) = (if C={} then UNIV else (\<Inter>x\<in>C. A x) - B)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1463
  "\<And>A B C. (\<Inter>x\<in>C. A - B x) = (if C={} then UNIV else A - (\<Union>x\<in>C. B x))"
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1464
  "\<And>a B C. (\<Inter>x\<in>C. insert a (B x)) = insert a (\<Inter>x\<in>C. B x)"
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1465
  "\<And>A B C. (\<Inter>x\<in>C. A x \<union> B) = ((\<Inter>x\<in>C. A x) \<union> B)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1466
  "\<And>A B C. (\<Inter>x\<in>C. A \<union> B x) = (A \<union> (\<Inter>x\<in>C. B x))"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1467
  "\<And>A B. (\<Inter>x\<in>\<Union>A. B x) = (\<Inter>y\<in>A. \<Inter>x\<in>y. B x)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1468
  "\<And>A B C. (\<Inter>z\<in>UNION A B. C z) = (\<Inter>x\<in>A. \<Inter>z\<in>B x. C z)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1469
  "\<And>A B f. (\<Inter>x\<in>f`A. B x) = (\<Inter>a\<in>A. B (f a))"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1470
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1471
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1472
lemma UN_ball_bex_simps [simp]:
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1473
  "\<And>A P. (\<forall>x\<in>\<Union>A. P x) \<longleftrightarrow> (\<forall>y\<in>A. \<forall>x\<in>y. P x)"
43967
610efb6bda1f more coherent structure in and across theories
haftmann
parents: 43944
diff changeset
  1474
  "\<And>A B P. (\<forall>x\<in>UNION A B. P x) = (\<forall>a\<in>A. \<forall>x\<in> B a. P x)"
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1475
  "\<And>A P. (\<exists>x\<in>\<Union>A. P x) \<longleftrightarrow> (\<exists>y\<in>A. \<exists>x\<in>y. P x)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1476
  "\<And>A B P. (\<exists>x\<in>UNION A B. P x) \<longleftrightarrow> (\<exists>a\<in>A. \<exists>x\<in>B a. P x)"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1477
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1478
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
  1479
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1480
text \<open>\medskip Maxiscoping: pulling out big Unions and Intersections.\<close>
13860
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1481
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1482
lemma UN_extend_simps:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1483
  "\<And>a B C. insert a (\<Union>x\<in>C. B x) = (if C={} then {a} else (\<Union>x\<in>C. insert a (B x)))"
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
  1484
  "\<And>A B C. (\<Union>x\<in>C. A x) \<union> B = (if C={} then B else (\<Union>x\<in>C. A x \<union> B))"
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1485
  "\<And>A B C. A \<union> (\<Union>x\<in>C. B x) = (if C={} then A else (\<Union>x\<in>C. A \<union> B x))"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1486
  "\<And>A B C. ((\<Union>x\<in>C. A x) \<inter> B) = (\<Union>x\<in>C. A x \<inter> B)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1487
  "\<And>A B C. (A \<inter> (\<Union>x\<in>C. B x)) = (\<Union>x\<in>C. A \<inter> B x)"
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1488
  "\<And>A B C. ((\<Union>x\<in>C. A x) - B) = (\<Union>x\<in>C. A x - B)"
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1489
  "\<And>A B C. (A - (\<Inter>x\<in>C. B x)) = (\<Union>x\<in>C. A - B x)"
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1490
  "\<And>A B. (\<Union>y\<in>A. \<Union>x\<in>y. B x) = (\<Union>x\<in>\<Union>A. B x)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1491
  "\<And>A B C. (\<Union>x\<in>A. \<Union>z\<in>B x. C z) = (\<Union>z\<in>UNION A B. C z)"
43831
e323be6b02a5 tuned notation and proofs
haftmann
parents: 43818
diff changeset
  1492
  "\<And>A B f. (\<Union>a\<in>A. B (f a)) = (\<Union>x\<in>f`A. B x)"
13860
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1493
  by auto
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1494
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1495
lemma INT_extend_simps:
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1496
  "\<And>A B C. (\<Inter>x\<in>C. A x) \<inter> B = (if C={} then B else (\<Inter>x\<in>C. A x \<inter> B))"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1497
  "\<And>A B C. A \<inter> (\<Inter>x\<in>C. B x) = (if C={} then A else (\<Inter>x\<in>C. A \<inter> B x))"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1498
  "\<And>A B C. (\<Inter>x\<in>C. A x) - B = (if C={} then UNIV - B else (\<Inter>x\<in>C. A x - B))"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1499
  "\<And>A B C. A - (\<Union>x\<in>C. B x) = (if C={} then A else (\<Inter>x\<in>C. A - B x))"
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1500
  "\<And>a B C. insert a (\<Inter>x\<in>C. B x) = (\<Inter>x\<in>C. insert a (B x))"
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1501
  "\<And>A B C. ((\<Inter>x\<in>C. A x) \<union> B) = (\<Inter>x\<in>C. A x \<union> B)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1502
  "\<And>A B C. A \<union> (\<Inter>x\<in>C. B x) = (\<Inter>x\<in>C. A \<union> B x)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1503
  "\<And>A B. (\<Inter>y\<in>A. \<Inter>x\<in>y. B x) = (\<Inter>x\<in>\<Union>A. B x)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1504
  "\<And>A B C. (\<Inter>x\<in>A. \<Inter>z\<in>B x. C z) = (\<Inter>z\<in>UNION A B. C z)"
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1505
  "\<And>A B f. (\<Inter>a\<in>A. B (f a)) = (\<Inter>x\<in>f`A. B x)"
13860
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1506
  by auto
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1507
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1508
text \<open>Finally\<close>
43872
6b917e5877d2 more consistent theorem names
haftmann
parents: 43871
diff changeset
  1509
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1510
no_notation
46691
72d81e789106 tuned syntax declarations; tuned structure
haftmann
parents: 46689
diff changeset
  1511
  less_eq (infix "\<sqsubseteq>" 50) and
72d81e789106 tuned syntax declarations; tuned structure
haftmann
parents: 46689
diff changeset
  1512
  less (infix "\<sqsubset>" 50)
32135
f645b51e8e54 set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
haftmann
parents: 32120
diff changeset
  1513
30596
140b22f22071 tuned some theorem and attribute bindings
haftmann
parents: 30531
diff changeset
  1514
lemmas mem_simps =
140b22f22071 tuned some theorem and attribute bindings
haftmann
parents: 30531
diff changeset
  1515
  insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff
140b22f22071 tuned some theorem and attribute bindings
haftmann
parents: 30531
diff changeset
  1516
  mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1517
  -- \<open>Each of these has ALREADY been added @{text "[simp]"} above.\<close>
21669
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21549
diff changeset
  1518
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1519
end
49905
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
  1520