src/HOL/Complete_Lattices.thy
author haftmann
Thu, 08 Nov 2018 09:11:52 +0100
changeset 69260 0a9688695a1b
parent 69164 74f1b0f10b2b
child 69274 ff7e6751a1a7
permissions -rw-r--r--
removed relics of ASCII syntax for indexed big operators
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
     1
(*  Title:      HOL/Complete_Lattices.thy
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
     2
    Author:     Tobias Nipkow
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
     3
    Author:     Lawrence C Paulson
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
     4
    Author:     Markus Wenzel
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
     5
    Author:     Florian Haftmann
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
     6
    Author:     Viorel Preoteasa (Complete Distributive Lattices)     
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
     7
*)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
     8
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
     9
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
    10
44860
56101fa00193 renamed theory Complete_Lattice to Complete_Lattices, in accordance with Lattices, Orderings etc.
haftmann
parents: 44845
diff changeset
    11
theory Complete_Lattices
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
    12
  imports Fun
32139
e271a64f03ff moved complete_lattice &c. into separate theory
haftmann
parents: 32135
diff changeset
    13
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
    14
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
    15
subsection \<open>Syntactic infimum and supremum operations\<close>
32879
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    16
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    17
class Inf =
68980
5717fbc55521 added spaces because otherwise nonatomic arguments look awful: BIGf x -> BIG f x
nipkow
parents: 68860
diff changeset
    18
  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
    19
begin
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    20
68796
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
    21
abbreviation (input) INFIMUM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" \<comment> \<open>legacy\<close>
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
    22
  where "INFIMUM A f \<equiv> \<Sqinter>(f ` A)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    23
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    24
end
32879
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    25
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    26
class Sup =
68980
5717fbc55521 added spaces because otherwise nonatomic arguments look awful: BIGf x -> BIG f x
nipkow
parents: 68860
diff changeset
    27
  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
    28
begin
32879
7f5ce7af45fd added syntactic Inf and Sup
haftmann
parents: 32678
diff changeset
    29
68796
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
    30
abbreviation (input) SUPREMUM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" \<comment> \<open>legacy\<close>
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
    31
  where "SUPREMUM A f \<equiv> \<Squnion>(f ` A)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
    32
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    33
end
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    34
68801
c898c2b1fd58 deprecation of ASCII syntax for indexed big operators
haftmann
parents: 68797
diff changeset
    35
syntax
62048
fefd79f6b232 retain ASCII syntax for output, when HOL/Library/Lattice_Syntax is not present (amending e96292f32c3c);
wenzelm
parents: 61955
diff changeset
    36
  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3INF _./ _)" [0, 10] 10)
68801
c898c2b1fd58 deprecation of ASCII syntax for indexed big operators
haftmann
parents: 68797
diff changeset
    37
  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3INF _\<in>_./ _)" [0, 0, 10] 10)
62048
fefd79f6b232 retain ASCII syntax for output, when HOL/Library/Lattice_Syntax is not present (amending e96292f32c3c);
wenzelm
parents: 61955
diff changeset
    38
  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3SUP _./ _)" [0, 10] 10)
68801
c898c2b1fd58 deprecation of ASCII syntax for indexed big operators
haftmann
parents: 68797
diff changeset
    39
  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3SUP _\<in>_./ _)" [0, 0, 10] 10)
62048
fefd79f6b232 retain ASCII syntax for output, when HOL/Library/Lattice_Syntax is not present (amending e96292f32c3c);
wenzelm
parents: 61955
diff changeset
    40
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
    41
syntax
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    42
  "_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
    43
  "_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
    44
  "_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
    45
  "_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
    46
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69164
diff changeset
    47
syntax (input) \<comment> \<open>legacy input syntax\<close>
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69164
diff changeset
    48
  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3INF _./ _)" [0, 10] 10)
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69164
diff changeset
    49
  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3INF _:_./ _)" [0, 0, 10] 10)
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69164
diff changeset
    50
  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3SUP _./ _)" [0, 10] 10)
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69164
diff changeset
    51
  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3SUP _:_./ _)" [0, 0, 10] 10)
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69164
diff changeset
    52
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 54147
diff changeset
    53
translations
68796
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
    54
  "\<Sqinter>x y. f"   \<rightleftharpoons> "\<Sqinter>x. \<Sqinter>y. f"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
    55
  "\<Sqinter>x. f"     \<rightleftharpoons> "\<Sqinter>CONST range (\<lambda>x. f)"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
    56
  "\<Sqinter>x\<in>A. f"   \<rightleftharpoons> "CONST Inf ((\<lambda>x. f) ` A)"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
    57
  "\<Squnion>x y. f"   \<rightleftharpoons> "\<Squnion>x. \<Squnion>y. f"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
    58
  "\<Squnion>x. f"     \<rightleftharpoons> "\<Squnion>CONST range (\<lambda>x. f)"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
    59
  "\<Squnion>x\<in>A. f"   \<rightleftharpoons> "CONST Sup ((\<lambda>x. f) `  A)"
46691
72d81e789106 tuned syntax declarations; tuned structure
haftmann
parents: 46689
diff changeset
    60
68797
haftmann
parents: 68796
diff changeset
    61
context Inf
haftmann
parents: 68796
diff changeset
    62
begin
haftmann
parents: 68796
diff changeset
    63
haftmann
parents: 68796
diff changeset
    64
lemma INF_image [simp]: " \<Sqinter>(g ` f ` A) = \<Sqinter>((g \<circ> f) ` A)"
haftmann
parents: 68796
diff changeset
    65
  by (simp add: image_comp)
haftmann
parents: 68796
diff changeset
    66
haftmann
parents: 68796
diff changeset
    67
lemma INF_identity_eq [simp]: "(\<Sqinter>x\<in>A. x) = \<Sqinter>A"
haftmann
parents: 68796
diff changeset
    68
  by simp
haftmann
parents: 68796
diff changeset
    69
haftmann
parents: 68796
diff changeset
    70
lemma INF_id_eq [simp]: "\<Sqinter>(id ` A) = \<Sqinter>A"
haftmann
parents: 68796
diff changeset
    71
  by simp
haftmann
parents: 68796
diff changeset
    72
haftmann
parents: 68796
diff changeset
    73
lemma INF_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> \<Sqinter>(C ` A) = \<Sqinter>(D ` B)"
haftmann
parents: 68796
diff changeset
    74
  by (simp add: image_def)
haftmann
parents: 68796
diff changeset
    75
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
    76
lemma INF_cong_strong [cong]:
68797
haftmann
parents: 68796
diff changeset
    77
  "A = B \<Longrightarrow> (\<And>x. x \<in> B =simp=> C x = D x) \<Longrightarrow> \<Sqinter>(C ` A) = \<Sqinter>(D ` B)"
haftmann
parents: 68796
diff changeset
    78
  unfolding simp_implies_def by (fact INF_cong)
haftmann
parents: 68796
diff changeset
    79
haftmann
parents: 68796
diff changeset
    80
end
haftmann
parents: 68796
diff changeset
    81
haftmann
parents: 68796
diff changeset
    82
context Sup
haftmann
parents: 68796
diff changeset
    83
begin
haftmann
parents: 68796
diff changeset
    84
haftmann
parents: 68796
diff changeset
    85
lemma SUP_image [simp]: "\<Squnion>(g ` f ` A) = \<Squnion>((g \<circ> f) ` A)"
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
    86
by(fact Inf.INF_image)
68797
haftmann
parents: 68796
diff changeset
    87
haftmann
parents: 68796
diff changeset
    88
lemma SUP_identity_eq [simp]: "(\<Squnion>x\<in>A. x) = \<Squnion>A"
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
    89
by(fact Inf.INF_identity_eq)
68797
haftmann
parents: 68796
diff changeset
    90
haftmann
parents: 68796
diff changeset
    91
lemma SUP_id_eq [simp]: "\<Squnion>(id ` A) = \<Squnion>A"
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
    92
by(fact Inf.INF_id_eq)
68797
haftmann
parents: 68796
diff changeset
    93
haftmann
parents: 68796
diff changeset
    94
lemma SUP_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> \<Squnion>(C ` A) = \<Squnion>(D ` B)"
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
    95
by (fact Inf.INF_cong)
68797
haftmann
parents: 68796
diff changeset
    96
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
    97
lemma SUP_cong_strong [cong]:
68797
haftmann
parents: 68796
diff changeset
    98
  "A = B \<Longrightarrow> (\<And>x. x \<in> B =simp=> C x = D x) \<Longrightarrow> \<Squnion>(C ` A) = \<Squnion>(D ` B)"
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
    99
by (fact Inf.INF_cong_strong)
68797
haftmann
parents: 68796
diff changeset
   100
haftmann
parents: 68796
diff changeset
   101
end
haftmann
parents: 68796
diff changeset
   102
62048
fefd79f6b232 retain ASCII syntax for output, when HOL/Library/Lattice_Syntax is not present (amending e96292f32c3c);
wenzelm
parents: 61955
diff changeset
   103
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   104
subsection \<open>Abstract complete lattices\<close>
32139
e271a64f03ff moved complete_lattice &c. into separate theory
haftmann
parents: 32135
diff changeset
   105
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   106
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
   107
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
   108
along with assumptions that define bottom and top
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   109
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
   110
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   111
class complete_lattice = lattice + Inf + Sup + bot + top +
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   112
  assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<le> x"
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   113
    and Inf_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> \<Sqinter>A"
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   114
    and Sup_upper: "x \<in> A \<Longrightarrow> x \<le> \<Squnion>A"
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   115
    and Sup_least: "(\<And>x. x \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> \<Squnion>A \<le> z"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   116
    and Inf_empty [simp]: "\<Sqinter>{} = \<top>"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   117
    and 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
   118
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
   119
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   120
subclass bounded_lattice
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   121
proof
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   122
  fix a
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   123
  show "\<bottom> \<le> a"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   124
    by (auto intro: Sup_least simp only: Sup_empty [symmetric])
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   125
  show "a \<le> \<top>"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   126
    by (auto intro: Inf_greatest simp only: Inf_empty [symmetric])
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   127
qed
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   128
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 64966
diff changeset
   129
lemma dual_complete_lattice: "class.complete_lattice Sup Inf sup (\<ge>) (>) inf \<top> \<bottom>"
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   130
  by (auto intro!: class.complete_lattice.intro dual_lattice)
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   131
    (unfold_locales, (fact Inf_empty Sup_empty Sup_upper Sup_least Inf_lower Inf_greatest)+)
32678
de1f7d4da21a added dual for complete lattice
haftmann
parents: 32642
diff changeset
   132
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   133
end
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   134
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   135
context complete_lattice
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   136
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
   137
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   138
lemma Sup_eqI:
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   139
  "(\<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
   140
  by (blast intro: antisym Sup_least Sup_upper)
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   141
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   142
lemma Inf_eqI:
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   143
  "(\<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
   144
  by (blast intro: antisym Inf_greatest Inf_lower)
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   145
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   146
lemma SUP_eqI:
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   147
  "(\<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
   148
  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
   149
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   150
lemma INF_eqI:
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   151
  "(\<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
   152
  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
   153
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   154
lemma INF_lower: "i \<in> A \<Longrightarrow> (\<Sqinter>i\<in>A. f i) \<le> f i"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   155
  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
   156
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   157
lemma INF_greatest: "(\<And>i. i \<in> A \<Longrightarrow> u \<le> f i) \<Longrightarrow> u \<le> (\<Sqinter>i\<in>A. f i)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   158
  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
   159
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   160
lemma SUP_upper: "i \<in> A \<Longrightarrow> f i \<le> (\<Squnion>i\<in>A. f i)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   161
  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
   162
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   163
lemma SUP_least: "(\<And>i. i \<in> A \<Longrightarrow> f i \<le> u) \<Longrightarrow> (\<Squnion>i\<in>A. f i) \<le> u"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   164
  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
   165
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   166
lemma Inf_lower2: "u \<in> A \<Longrightarrow> u \<le> v \<Longrightarrow> \<Sqinter>A \<le> v"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   167
  using Inf_lower [of u A] by auto
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   168
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   169
lemma INF_lower2: "i \<in> A \<Longrightarrow> f i \<le> u \<Longrightarrow> (\<Sqinter>i\<in>A. f i) \<le> u"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   170
  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
   171
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   172
lemma Sup_upper2: "u \<in> A \<Longrightarrow> v \<le> u \<Longrightarrow> v \<le> \<Squnion>A"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   173
  using Sup_upper [of u A] by auto
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   174
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   175
lemma SUP_upper2: "i \<in> A \<Longrightarrow> u \<le> f i \<Longrightarrow> u \<le> (\<Squnion>i\<in>A. f i)"
44103
cedaca00789f more uniform naming scheme for Inf/INF and Sup/SUP lemmas
haftmann
parents: 44085
diff changeset
   176
  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
   177
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   178
lemma le_Inf_iff: "b \<le> \<Sqinter>A \<longleftrightarrow> (\<forall>a\<in>A. b \<le> a)"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   179
  by (auto intro: Inf_greatest dest: Inf_lower)
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   180
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   181
lemma le_INF_iff: "u \<le> (\<Sqinter>i\<in>A. f i) \<longleftrightarrow> (\<forall>i\<in>A. u \<le> f i)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   182
  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
   183
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   184
lemma Sup_le_iff: "\<Squnion>A \<le> b \<longleftrightarrow> (\<forall>a\<in>A. a \<le> b)"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   185
  by (auto intro: Sup_least dest: Sup_upper)
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   186
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   187
lemma SUP_le_iff: "(\<Squnion>i\<in>A. f i) \<le> u \<longleftrightarrow> (\<forall>i\<in>A. f i \<le> u)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   188
  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
   189
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   190
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
   191
  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
   192
68797
haftmann
parents: 68796
diff changeset
   193
lemma INF_insert [simp]: "(\<Sqinter>x\<in>insert a A. f x) = f a \<sqinter> \<Sqinter>(f ` A)"
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
   194
  by (simp cong del: INF_cong_strong)
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   195
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52141
diff changeset
   196
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
   197
  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
   198
68797
haftmann
parents: 68796
diff changeset
   199
lemma SUP_insert [simp]: "(\<Squnion>x\<in>insert a A. f x) = f a \<squnion> \<Squnion>(f ` A)"
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
   200
  by (simp cong del: SUP_cong_strong)
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
   201
44067
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   202
lemma INF_empty [simp]: "(\<Sqinter>x\<in>{}. f x) = \<top>"
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
   203
  by (simp cong del: INF_cong_strong)
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   204
44067
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   205
lemma SUP_empty [simp]: "(\<Squnion>x\<in>{}. f x) = \<bottom>"
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
   206
  by (simp cong del: SUP_cong_strong)
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   207
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   208
lemma Inf_UNIV [simp]: "\<Sqinter>UNIV = \<bottom>"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   209
  by (auto intro!: antisym Inf_lower)
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   210
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   211
lemma Sup_UNIV [simp]: "\<Squnion>UNIV = \<top>"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   212
  by (auto intro!: antisym Sup_upper)
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   213
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   214
lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<le> a}"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   215
  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
   216
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   217
lemma Sup_Inf:  "\<Squnion>A = \<Sqinter>{b. \<forall>a \<in> A. a \<le> b}"
44040
32881ab55eac tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44039
diff changeset
   218
  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
   219
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   220
lemma Inf_superset_mono: "B \<subseteq> A \<Longrightarrow> \<Sqinter>A \<le> \<Sqinter>B"
43899
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   221
  by (auto intro: Inf_greatest Inf_lower)
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   222
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   223
lemma Sup_subset_mono: "A \<subseteq> B \<Longrightarrow> \<Squnion>A \<le> \<Squnion>B"
43899
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   224
  by (auto intro: Sup_least Sup_upper)
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   225
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   226
lemma Inf_mono:
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   227
  assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b"
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   228
  shows "\<Sqinter>A \<le> \<Sqinter>B"
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   229
proof (rule Inf_greatest)
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   230
  fix b assume "b \<in> B"
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   231
  with assms obtain a where "a \<in> A" and "a \<le> b" by blast
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   232
  from \<open>a \<in> A\<close> have "\<Sqinter>A \<le> a" by (rule Inf_lower)
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   233
  with \<open>a \<le> b\<close> show "\<Sqinter>A \<le> b" by auto
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   234
qed
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 37767
diff changeset
   235
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   236
lemma INF_mono: "(\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> (\<Sqinter>n\<in>A. f n) \<le> (\<Sqinter>n\<in>B. g n)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   237
  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
   238
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69164
diff changeset
   239
lemma INF_mono': "(\<And>x. f x \<le> g x) \<Longrightarrow> (\<Sqinter>x\<in>A. f x) \<le> (\<Sqinter>x\<in>A. g x)"
68860
f443ec10447d Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents: 68802
diff changeset
   240
  by (rule INF_mono) auto
f443ec10447d Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents: 68802
diff changeset
   241
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   242
lemma Sup_mono:
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   243
  assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<le> b"
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   244
  shows "\<Squnion>A \<le> \<Squnion>B"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   245
proof (rule Sup_least)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   246
  fix a assume "a \<in> A"
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   247
  with assms obtain b where "b \<in> B" and "a \<le> b" by blast
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   248
  from \<open>b \<in> B\<close> have "b \<le> \<Squnion>B" by (rule Sup_upper)
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   249
  with \<open>a \<le> b\<close> show "a \<le> \<Squnion>B" by auto
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   250
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
   251
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   252
lemma SUP_mono: "(\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<le> g m) \<Longrightarrow> (\<Squnion>n\<in>A. f n) \<le> (\<Squnion>n\<in>B. g n)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   253
  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
   254
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69164
diff changeset
   255
lemma SUP_mono': "(\<And>x. f x \<le> g x) \<Longrightarrow> (\<Squnion>x\<in>A. f x) \<le> (\<Squnion>x\<in>A. g x)"
68860
f443ec10447d Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents: 68802
diff changeset
   256
  by (rule SUP_mono) auto
f443ec10447d Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents: 68802
diff changeset
   257
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   258
lemma INF_superset_mono: "B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<le> g x) \<Longrightarrow> (\<Sqinter>x\<in>A. f x) \<le> (\<Sqinter>x\<in>B. g x)"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   259
  \<comment> \<open>The last inclusion is POSITIVE!\<close>
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   260
  by (blast intro: INF_mono dest: subsetD)
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   261
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   262
lemma SUP_subset_mono: "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x) \<Longrightarrow> (\<Squnion>x\<in>A. f x) \<le> (\<Squnion>x\<in>B. g x)"
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   263
  by (blast intro: SUP_mono dest: subsetD)
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   264
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   265
lemma Inf_less_eq:
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   266
  assumes "\<And>v. v \<in> A \<Longrightarrow> v \<le> u"
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   267
    and "A \<noteq> {}"
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   268
  shows "\<Sqinter>A \<le> u"
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   269
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   270
  from \<open>A \<noteq> {}\<close> obtain v where "v \<in> A" by blast
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   271
  moreover from \<open>v \<in> A\<close> assms(1) have "v \<le> u" by blast
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   272
  ultimately show ?thesis by (rule Inf_lower2)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   273
qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   274
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   275
lemma less_eq_Sup:
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   276
  assumes "\<And>v. v \<in> A \<Longrightarrow> u \<le> v"
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   277
    and "A \<noteq> {}"
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   278
  shows "u \<le> \<Squnion>A"
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   279
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   280
  from \<open>A \<noteq> {}\<close> obtain v where "v \<in> A" by blast
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   281
  moreover from \<open>v \<in> A\<close> assms(1) have "u \<le> v" by blast
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   282
  ultimately show ?thesis by (rule Sup_upper2)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   283
qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   284
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
   285
lemma INF_eq:
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
   286
  assumes "\<And>i. i \<in> A \<Longrightarrow> \<exists>j\<in>B. f i \<ge> g j"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   287
    and "\<And>j. j \<in> B \<Longrightarrow> \<exists>i\<in>A. g j \<ge> f i"
68797
haftmann
parents: 68796
diff changeset
   288
  shows "\<Sqinter>(f ` A) = \<Sqinter>(g ` B)"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
   289
  by (intro antisym INF_greatest) (blast intro: INF_lower2 dest: assms)+
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
   290
56212
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 56166
diff changeset
   291
lemma SUP_eq:
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   292
  assumes "\<And>i. i \<in> A \<Longrightarrow> \<exists>j\<in>B. f i \<le> g j"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   293
    and "\<And>j. j \<in> B \<Longrightarrow> \<exists>i\<in>A. g j \<le> f i"
68797
haftmann
parents: 68796
diff changeset
   294
  shows "\<Squnion>(f ` A) = \<Squnion>(g ` B)"
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   295
  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
   296
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   297
lemma less_eq_Inf_inter: "\<Sqinter>A \<squnion> \<Sqinter>B \<le> \<Sqinter>(A \<inter> B)"
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   298
  by (auto intro: Inf_greatest Inf_lower)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   299
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   300
lemma Sup_inter_less_eq: "\<Squnion>(A \<inter> B) \<le> \<Squnion>A \<sqinter> \<Squnion>B "
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   301
  by (auto intro: Sup_least Sup_upper)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   302
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   303
lemma Inf_union_distrib: "\<Sqinter>(A \<union> B) = \<Sqinter>A \<sqinter> \<Sqinter>B"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   304
  by (rule antisym) (auto intro: Inf_greatest Inf_lower le_infI1 le_infI2)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   305
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   306
lemma INF_union: "(\<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
   307
  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
   308
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   309
lemma Sup_union_distrib: "\<Squnion>(A \<union> B) = \<Squnion>A \<squnion> \<Squnion>B"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   310
  by (rule antisym) (auto intro: Sup_least Sup_upper le_supI1 le_supI2)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   311
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   312
lemma SUP_union: "(\<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
   313
  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
   314
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   315
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
   316
  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
   317
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   318
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)"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   319
  (is "?L = ?R")
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   320
proof (rule antisym)
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   321
  show "?L \<le> ?R"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   322
    by (auto intro: le_supI1 le_supI2 SUP_upper SUP_mono)
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   323
  show "?R \<le> ?L"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   324
    by (rule SUP_least) (auto intro: le_supI1 le_supI2 SUP_upper)
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   325
qed
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   326
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   327
lemma Inf_top_conv [simp]:
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   328
  "\<Sqinter>A = \<top> \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   329
  "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   330
proof -
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   331
  show "\<Sqinter>A = \<top> \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   332
  proof
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   333
    assume "\<forall>x\<in>A. x = \<top>"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   334
    then have "A = {} \<or> A = {\<top>}" by auto
44919
482f1807976e tune proofs
noschinl
parents: 44918
diff changeset
   335
    then show "\<Sqinter>A = \<top>" by auto
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   336
  next
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   337
    assume "\<Sqinter>A = \<top>"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   338
    show "\<forall>x\<in>A. x = \<top>"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   339
    proof (rule ccontr)
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   340
      assume "\<not> (\<forall>x\<in>A. x = \<top>)"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   341
      then obtain x where "x \<in> A" and "x \<noteq> \<top>" by blast
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   342
      then obtain B where "A = insert x B" by blast
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   343
      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
   344
    qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   345
  qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   346
  then show "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)" by auto
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   347
qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   348
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   349
lemma INF_top_conv [simp]:
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   350
  "(\<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
   351
  "\<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
   352
  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
   353
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   354
lemma Sup_bot_conv [simp]:
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   355
  "\<Squnion>A = \<bottom> \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   356
  "\<bottom> = \<Squnion>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)"
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   357
  using dual_complete_lattice
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   358
  by (rule complete_lattice.Inf_top_conv)+
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   359
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   360
lemma SUP_bot_conv [simp]:
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   361
  "(\<Squnion>x\<in>A. B x) = \<bottom> \<longleftrightarrow> (\<forall>x\<in>A. B x = \<bottom>)"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   362
  "\<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
   363
  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
   364
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   365
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
   366
  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
   367
43870
92129f505125 structuring duals together
haftmann
parents: 43868
diff changeset
   368
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
   369
  by (auto intro: antisym SUP_upper SUP_least)
43870
92129f505125 structuring duals together
haftmann
parents: 43868
diff changeset
   370
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   371
lemma INF_top [simp]: "(\<Sqinter>x\<in>A. \<top>) = \<top>"
44921
58eef4843641 tuned proofs
huffman
parents: 44920
diff changeset
   372
  by (cases "A = {}") simp_all
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
   373
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44860
diff changeset
   374
lemma SUP_bot [simp]: "(\<Squnion>x\<in>A. \<bottom>) = \<bottom>"
44921
58eef4843641 tuned proofs
huffman
parents: 44920
diff changeset
   375
  by (cases "A = {}") simp_all
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
   376
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   377
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
   378
  by (iprover intro: INF_lower INF_greatest order_trans antisym)
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   379
43870
92129f505125 structuring duals together
haftmann
parents: 43868
diff changeset
   380
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
   381
  by (iprover intro: SUP_upper SUP_least order_trans antisym)
43870
92129f505125 structuring duals together
haftmann
parents: 43868
diff changeset
   382
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   383
lemma INF_absorb:
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   384
  assumes "k \<in> I"
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   385
  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
   386
proof -
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   387
  from assms obtain J where "I = insert k J" by blast
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   388
  then show ?thesis by simp
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   389
qed
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   390
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   391
lemma SUP_absorb:
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   392
  assumes "k \<in> I"
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   393
  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
   394
proof -
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   395
  from assms obtain J where "I = insert k J" by blast
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   396
  then show ?thesis by simp
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   397
qed
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   398
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   399
lemma INF_inf_const1: "I \<noteq> {} \<Longrightarrow> (\<Sqinter>i\<in>I. inf x (f i)) = inf x (\<Sqinter>i\<in>I. f i)"
57448
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   400
  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
   401
     (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
   402
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   403
lemma INF_inf_const2: "I \<noteq> {} \<Longrightarrow> (\<Sqinter>i\<in>I. inf (f i) x) = inf (\<Sqinter>i\<in>I. f i) x"
57448
159e45728ceb more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents: 57197
diff changeset
   404
  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
   405
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   406
lemma INF_constant: "(\<Sqinter>y\<in>A. c) = (if A = {} then \<top> else c)"
44921
58eef4843641 tuned proofs
huffman
parents: 44920
diff changeset
   407
  by simp
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   408
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   409
lemma SUP_constant: "(\<Squnion>y\<in>A. c) = (if A = {} then \<bottom> else c)"
44921
58eef4843641 tuned proofs
huffman
parents: 44920
diff changeset
   410
  by simp
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   411
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   412
lemma less_INF_D:
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   413
  assumes "y < (\<Sqinter>i\<in>A. f i)" "i \<in> A"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   414
  shows "y < f i"
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   415
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   416
  note \<open>y < (\<Sqinter>i\<in>A. f i)\<close>
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   417
  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
   418
    by (rule INF_lower)
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   419
  finally show "y < f i" .
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   420
qed
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   421
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   422
lemma SUP_lessD:
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   423
  assumes "(\<Squnion>i\<in>A. f i) < y" "i \<in> A"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   424
  shows "f i < y"
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   425
proof -
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   426
  have "f i \<le> (\<Squnion>i\<in>A. f i)"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   427
    using \<open>i \<in> A\<close> by (rule SUP_upper)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   428
  also note \<open>(\<Squnion>i\<in>A. f i) < y\<close>
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   429
  finally show "f i < y" .
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   430
qed
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   431
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   432
lemma INF_UNIV_bool_expand: "(\<Sqinter>b. A b) = A True \<sqinter> A False"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   433
  by (simp add: UNIV_bool inf_commute)
43868
9684251c7ec1 more lemmas about Sup
haftmann
parents: 43867
diff changeset
   434
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   435
lemma SUP_UNIV_bool_expand: "(\<Squnion>b. A b) = A True \<squnion> A False"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   436
  by (simp add: UNIV_bool sup_commute)
43871
79c3231e0593 more lemmas about SUP
haftmann
parents: 43870
diff changeset
   437
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   438
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
   439
  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
   440
68797
haftmann
parents: 68796
diff changeset
   441
lemma INF_le_SUP: "A \<noteq> {} \<Longrightarrow> \<Sqinter>(f ` A) \<le> \<Squnion>(f ` A)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   442
  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
   443
68797
haftmann
parents: 68796
diff changeset
   444
lemma INF_eq_const: "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i = x) \<Longrightarrow> \<Sqinter>(f ` I) = x"
54414
72949fae4f81 add equalities for SUP and INF over constant functions
hoelzl
parents: 54259
diff changeset
   445
  by (auto intro: INF_eqI)
72949fae4f81 add equalities for SUP and INF over constant functions
hoelzl
parents: 54259
diff changeset
   446
68797
haftmann
parents: 68796
diff changeset
   447
lemma SUP_eq_const: "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i = x) \<Longrightarrow> \<Squnion>(f ` I) = x"
56248
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   448
  by (auto intro: SUP_eqI)
54414
72949fae4f81 add equalities for SUP and INF over constant functions
hoelzl
parents: 54259
diff changeset
   449
68797
haftmann
parents: 68796
diff changeset
   450
lemma INF_eq_iff: "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<le> c) \<Longrightarrow> \<Sqinter>(f ` I) = 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
   451
  using INF_eq_const [of I f c] INF_lower [of _ I f]
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
   452
  by (auto intro: antisym cong del: INF_cong_strong)
56248
67dc9549fa15 generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
haftmann
parents: 56218
diff changeset
   453
68797
haftmann
parents: 68796
diff changeset
   454
lemma SUP_eq_iff: "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> c \<le> f i) \<Longrightarrow> \<Squnion>(f ` I) = 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
   455
  using SUP_eq_const [of I f c] SUP_upper [of _ I f]
69164
74f1b0f10b2b uniform naming of strong congruence rules
nipkow
parents: 69020
diff changeset
   456
  by (auto intro: antisym cong del: SUP_cong_strong)
54414
72949fae4f81 add equalities for SUP and INF over constant functions
hoelzl
parents: 54259
diff changeset
   457
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
   458
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
   459
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   460
context complete_lattice
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   461
begin
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   462
lemma Sup_Inf_le: "Sup (Inf ` {f ` A | f . (\<forall> Y \<in> A . f Y \<in> Y)}) \<le> Inf (Sup ` A)"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   463
  by (rule SUP_least, clarify, rule INF_greatest, simp add: INF_lower2 Sup_upper)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   464
end 
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   465
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   466
class complete_distrib_lattice = complete_lattice +
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   467
  assumes Inf_Sup_le: "Inf (Sup ` A) \<le> Sup (Inf ` {f ` A | f . (\<forall> Y \<in> A . f Y \<in> Y)})"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   468
begin
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   469
  
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   470
lemma Inf_Sup: "Inf (Sup ` A) = Sup (Inf ` {f ` A | f . (\<forall> Y \<in> A . f Y \<in> Y)})"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   471
  by (rule antisym, rule Inf_Sup_le, rule Sup_Inf_le)
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   472
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   473
subclass distrib_lattice
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   474
proof
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   475
  fix a b c
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   476
  show "a \<squnion> b \<sqinter> c = (a \<squnion> b) \<sqinter> (a \<squnion> c)"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   477
  proof (rule antisym, simp_all, safe)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   478
    show "b \<sqinter> c \<le> a \<squnion> b"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   479
      by (rule le_infI1, simp)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   480
    show "b \<sqinter> c \<le> a \<squnion> c"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   481
      by (rule le_infI2, simp)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   482
    have [simp]: "a \<sqinter> c \<le> a \<squnion> b \<sqinter> c"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   483
      by (rule le_infI1, simp)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   484
    have [simp]: "b \<sqinter> a \<le> a \<squnion> b \<sqinter> c"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   485
      by (rule le_infI2, simp)
68797
haftmann
parents: 68796
diff changeset
   486
    have "\<Sqinter>(Sup ` {{a, b}, {a, c}}) =
haftmann
parents: 68796
diff changeset
   487
      \<Squnion>(Inf ` {f ` {{a, b}, {a, c}} | f. \<forall>Y\<in>{{a, b}, {a, c}}. f Y \<in> Y})"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   488
      by (rule Inf_Sup)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   489
    from this show "(a \<squnion> b) \<sqinter> (a \<squnion> c) \<le> a \<squnion> b \<sqinter> c"
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   490
      apply simp
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   491
      by (rule SUP_least, safe, simp_all)
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   492
  qed
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   493
qed
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   494
end
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
   495
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   496
context complete_lattice
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   497
begin
56074
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   498
context
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   499
  fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   500
  assumes "mono f"
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   501
begin
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   502
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   503
lemma mono_Inf: "f (\<Sqinter>A) \<le> (\<Sqinter>x\<in>A. f x)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   504
  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
   505
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   506
lemma mono_Sup: "(\<Squnion>x\<in>A. f x) \<le> f (\<Squnion>A)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   507
  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
   508
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   509
lemma mono_INF: "f (\<Sqinter>i\<in>I. A i) \<le> (\<Sqinter>x\<in>I. f (A x))"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   510
  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
   511
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   512
lemma mono_SUP: "(\<Squnion>x\<in>I. f (A x)) \<le> f (\<Squnion>i\<in>I. A i)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   513
  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
   514
56074
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   515
end
30a60277aa6e monotonicity in complete lattices
haftmann
parents: 56015
diff changeset
   516
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   517
end
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   518
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
   519
class complete_boolean_algebra = boolean_algebra + complete_distrib_lattice
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   520
begin
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   521
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   522
lemma uminus_Inf: "- (\<Sqinter>A) = \<Squnion>(uminus ` A)"
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   523
proof (rule antisym)
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   524
  show "- \<Sqinter>A \<le> \<Squnion>(uminus ` A)"
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   525
    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
   526
  show "\<Squnion>(uminus ` A) \<le> - \<Sqinter>A"
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   527
    by (rule Sup_least, rule compl_le_swap1, rule Inf_lower) auto
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   528
qed
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   529
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   530
lemma uminus_INF: "- (\<Sqinter>x\<in>A. B x) = (\<Squnion>x\<in>A. - B x)"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
   531
  by (simp add: uminus_Inf image_image)
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   532
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   533
lemma uminus_Sup: "- (\<Squnion>A) = \<Sqinter>(uminus ` A)"
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   534
proof -
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   535
  have "\<Squnion>A = - \<Sqinter>(uminus ` A)"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   536
    by (simp add: image_image uminus_INF)
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   537
  then show ?thesis by simp
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   538
qed
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   539
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   540
lemma uminus_SUP: "- (\<Squnion>x\<in>A. B x) = (\<Sqinter>x\<in>A. - B x)"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
   541
  by (simp add: uminus_Sup image_image)
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   542
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   543
end
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   544
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   545
class complete_linorder = linorder + complete_lattice
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   546
begin
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   547
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   548
lemma dual_complete_linorder:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 64966
diff changeset
   549
  "class.complete_linorder Sup Inf sup (\<ge>) (>) inf \<top> \<bottom>"
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   550
  by (rule class.complete_linorder.intro, rule dual_complete_lattice, rule dual_linorder)
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   551
51386
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   552
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
   553
  by (auto intro: antisym simp add: min_def fun_eq_iff)
51386
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   554
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   555
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
   556
  by (auto intro: antisym simp add: max_def fun_eq_iff)
51386
616f68ddcb7f generalized subclass relation;
haftmann
parents: 51341
diff changeset
   557
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   558
lemma Inf_less_iff: "\<Sqinter>S < a \<longleftrightarrow> (\<exists>x\<in>S. x < a)"
63172
d4f459eb7ed0 tuned proofs;
wenzelm
parents: 63099
diff changeset
   559
  by (simp add: not_le [symmetric] le_Inf_iff)
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   560
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   561
lemma INF_less_iff: "(\<Sqinter>i\<in>A. f i) < a \<longleftrightarrow> (\<exists>x\<in>A. f x < a)"
63172
d4f459eb7ed0 tuned proofs;
wenzelm
parents: 63099
diff changeset
   562
  by (simp add: Inf_less_iff [of "f ` A"])
44041
01d6ab227069 tuned order: pushing INF and SUP to Inf and Sup
haftmann
parents: 44040
diff changeset
   563
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   564
lemma less_Sup_iff: "a < \<Squnion>S \<longleftrightarrow> (\<exists>x\<in>S. a < x)"
63172
d4f459eb7ed0 tuned proofs;
wenzelm
parents: 63099
diff changeset
   565
  by (simp add: not_le [symmetric] Sup_le_iff)
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   566
63820
9f004fbf9d5c discontinued theory-local special syntax for lattice orderings
haftmann
parents: 63576
diff changeset
   567
lemma less_SUP_iff: "a < (\<Squnion>i\<in>A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)"
63172
d4f459eb7ed0 tuned proofs;
wenzelm
parents: 63099
diff changeset
   568
  by (simp add: less_Sup_iff [of _ "f ` A"])
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   569
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   570
lemma Sup_eq_top_iff [simp]: "\<Squnion>A = \<top> \<longleftrightarrow> (\<forall>x<\<top>. \<exists>i\<in>A. x < i)"
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   571
proof
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   572
  assume *: "\<Squnion>A = \<top>"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   573
  show "(\<forall>x<\<top>. \<exists>i\<in>A. x < i)"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   574
    unfolding * [symmetric]
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   575
  proof (intro allI impI)
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   576
    fix x
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   577
    assume "x < \<Squnion>A"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   578
    then show "\<exists>i\<in>A. x < i"
63172
d4f459eb7ed0 tuned proofs;
wenzelm
parents: 63099
diff changeset
   579
      by (simp add: less_Sup_iff)
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   580
  qed
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   581
next
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   582
  assume *: "\<forall>x<\<top>. \<exists>i\<in>A. x < i"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   583
  show "\<Squnion>A = \<top>"
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   584
  proof (rule ccontr)
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   585
    assume "\<Squnion>A \<noteq> \<top>"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   586
    with top_greatest [of "\<Squnion>A"] have "\<Squnion>A < \<top>"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   587
      unfolding le_less by auto
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   588
    with * have "\<Squnion>A < \<Squnion>A"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   589
      unfolding less_Sup_iff by auto
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   590
    then show False by auto
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   591
  qed
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   592
qed
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   593
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   594
lemma SUP_eq_top_iff [simp]: "(\<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
   595
  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
   596
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   597
lemma Inf_eq_bot_iff [simp]: "\<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
   598
  using dual_complete_linorder
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   599
  by (rule complete_linorder.Sup_eq_top_iff)
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
   600
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   601
lemma INF_eq_bot_iff [simp]: "(\<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
   602
  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
   603
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   604
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
   605
proof safe
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   606
  fix y
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   607
  assume "x \<ge> \<Sqinter>A" "y > x"
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   608
  then have "y > \<Sqinter>A" by auto
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   609
  then show "\<exists>a\<in>A. y > a"
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   610
    unfolding Inf_less_iff .
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 49905
diff changeset
   611
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
   612
68802
3974935e0252 some modernization of notation
haftmann
parents: 68801
diff changeset
   613
lemma INF_le_iff: "\<Sqinter>(f ` A) \<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
   614
  using Inf_le_iff [of "f ` A"] by simp
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   615
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   616
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
   617
proof safe
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   618
  fix y
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   619
  assume "x \<le> \<Squnion>A" "y < x"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   620
  then have "y < \<Squnion>A" by auto
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   621
  then show "\<exists>a\<in>A. y < a"
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   622
    unfolding less_Sup_iff .
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   623
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
   624
68802
3974935e0252 some modernization of notation
haftmann
parents: 68801
diff changeset
   625
lemma le_SUP_iff: "x \<le> \<Squnion>(f ` A) \<longleftrightarrow> (\<forall>y<x. \<exists>i\<in>A. y < f i)"
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
   626
  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
   627
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   628
end
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   629
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   630
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
   631
44024
de7642fcbe1e class complete_distrib_lattice
haftmann
parents: 43967
diff changeset
   632
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
   633
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
   634
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   635
definition [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
   636
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   637
definition [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
   638
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   639
instance
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   640
  by standard (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
   641
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   642
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
   643
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   644
lemma not_False_in_image_Ball [simp]: "False \<notin> P ` A \<longleftrightarrow> Ball A P"
49905
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   645
  by auto
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   646
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   647
lemma True_in_image_Bex [simp]: "True \<in> P ` A \<longleftrightarrow> Bex A P"
49905
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   648
  by auto
a81f95693c68 simp results for simplification results of Inf/Sup expressions on bool;
haftmann
parents: 46884
diff changeset
   649
68802
3974935e0252 some modernization of notation
haftmann
parents: 68801
diff changeset
   650
lemma INF_bool_eq [simp]: "(\<lambda>A f. \<Sqinter>(f ` A)) = Ball"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
   651
  by (simp add: fun_eq_iff)
32120
53a21a5e6889 attempt for more concise setup of non-etacontracting binders
haftmann
parents: 32117
diff changeset
   652
68802
3974935e0252 some modernization of notation
haftmann
parents: 68801
diff changeset
   653
lemma SUP_bool_eq [simp]: "(\<lambda>A f. \<Squnion>(f ` A)) = Bex"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
   654
  by (simp add: fun_eq_iff)
32120
53a21a5e6889 attempt for more concise setup of non-etacontracting binders
haftmann
parents: 32117
diff changeset
   655
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   656
instance bool :: complete_boolean_algebra
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
   657
  by (standard, fastforce)
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
   658
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   659
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
   660
57197
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   661
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
   662
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
   663
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   664
definition "\<Sqinter>A = (\<lambda>x. \<Sqinter>f\<in>A. f x)"
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   665
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   666
lemma Inf_apply [simp, code]: "(\<Sqinter>A) x = (\<Sqinter>f\<in>A. f x)"
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   667
  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
   668
57197
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   669
instance ..
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   670
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   671
end
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   672
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   673
instantiation "fun" :: (type, Sup) Sup
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   674
begin
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   675
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   676
definition "\<Squnion>A = (\<lambda>x. \<Squnion>f\<in>A. f x)"
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   677
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   678
lemma Sup_apply [simp, code]: "(\<Squnion>A) x = (\<Squnion>f\<in>A. f x)"
41080
294956ff285b nice syntax for lattice INFI, SUPR;
haftmann
parents: 40872
diff changeset
   679
  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
   680
57197
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   681
instance ..
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   682
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   683
end
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   684
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   685
instantiation "fun" :: (type, complete_lattice) complete_lattice
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   686
begin
4cf607675df8 Sup/Inf on functions decoupled from complete_lattice.
nipkow
parents: 56742
diff changeset
   687
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   688
instance
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   689
  by standard (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
   690
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
   691
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
   692
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   693
lemma INF_apply [simp]: "(\<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
   694
  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
   695
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   696
lemma SUP_apply [simp]: "(\<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
   697
  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
   698
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   699
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
   700
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   701
lemma Inf1_I: "(\<And>P. P \<in> A \<Longrightarrow> P a) \<Longrightarrow> (\<Sqinter>A) a"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   702
  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
   703
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   704
lemma INF1_I: "(\<And>x. x \<in> A \<Longrightarrow> B x b) \<Longrightarrow> (\<Sqinter>x\<in>A. B x) b"
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
   705
  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
   706
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   707
lemma INF2_I: "(\<And>x. x \<in> A \<Longrightarrow> B x b c) \<Longrightarrow> (\<Sqinter>x\<in>A. B x) b c"
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
   708
  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
   709
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   710
lemma Inf2_I: "(\<And>r. r \<in> A \<Longrightarrow> r a b) \<Longrightarrow> (\<Sqinter>A) a b"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   711
  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
   712
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   713
lemma Inf1_D: "(\<Sqinter>A) a \<Longrightarrow> P \<in> A \<Longrightarrow> P a"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   714
  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
   715
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   716
lemma INF1_D: "(\<Sqinter>x\<in>A. B x) b \<Longrightarrow> a \<in> A \<Longrightarrow> B a b"
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
   717
  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
   718
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   719
lemma Inf2_D: "(\<Sqinter>A) a b \<Longrightarrow> r \<in> A \<Longrightarrow> r a b"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   720
  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
   721
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   722
lemma INF2_D: "(\<Sqinter>x\<in>A. B x) b c \<Longrightarrow> a \<in> A \<Longrightarrow> B a b c"
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
   723
  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
   724
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
   725
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
   726
  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
   727
  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
   728
  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
   729
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
   730
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
   731
  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
   732
  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
   733
  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
   734
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
   735
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
   736
  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
   737
  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
   738
  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
   739
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
   740
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
   741
  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
   742
  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
   743
  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
   744
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   745
lemma Sup1_I: "P \<in> A \<Longrightarrow> P a \<Longrightarrow> (\<Squnion>A) a"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 46882
diff changeset
   746
  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
   747
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   748
lemma SUP1_I: "a \<in> A \<Longrightarrow> B a b \<Longrightarrow> (\<Squnion>x\<in>A. B x) b"
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
   749
  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
   750
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   751
lemma Sup2_I: "r \<in> A \<Longrightarrow> r a b \<Longrightarrow> (\<Squnion>A) a b"
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
   752
  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
   753
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   754
lemma SUP2_I: "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
   755
  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
   756
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
   757
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
   758
  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
   759
  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
   760
  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
   761
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
   762
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
   763
  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
   764
  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
   765
  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
   766
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
   767
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
   768
  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
   769
  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
   770
  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
   771
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
   772
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
   773
  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
   774
  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
   775
  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
   776
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
   777
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   778
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
   779
45960
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   780
instantiation "set" :: (type) complete_lattice
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   781
begin
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   782
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   783
definition "\<Sqinter>A = {x. \<Sqinter>((\<lambda>B. x \<in> B) ` A)}"
45960
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   784
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   785
definition "\<Squnion>A = {x. \<Squnion>((\<lambda>B. x \<in> B) ` A)}"
45960
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   786
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   787
instance
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   788
  by standard (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
   789
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   790
end
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
   791
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   792
subsubsection \<open>Inter\<close>
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   793
68980
5717fbc55521 added spaces because otherwise nonatomic arguments look awful: BIGf x -> BIG f x
nipkow
parents: 68860
diff changeset
   794
abbreviation Inter :: "'a set set \<Rightarrow> 'a set"  ("\<Inter> _" [900] 900)
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61799
diff changeset
   795
  where "\<Inter>S \<equiv> \<Sqinter>S"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   796
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   797
lemma Inter_eq: "\<Inter>A = {x. \<forall>B \<in> A. x \<in> B}"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   798
proof (rule set_eqI)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   799
  fix x
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   800
  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
   801
    by auto
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   802
  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
   803
    by (simp add: Inf_set_def image_def)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   804
qed
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   805
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   806
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
   807
  by (unfold Inter_eq) blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   808
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   809
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
   810
  by (simp add: Inter_eq)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   811
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   812
text \<open>
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   813
  \<^medskip> A ``destruct'' rule -- every @{term X} in @{term C}
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   814
  contains @{term A} as an element, but @{prop "A \<in> X"} can hold when
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   815
  @{prop "X \<in> C"} does not!  This rule is analogous to \<open>spec\<close>.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   816
\<close>
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   817
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   818
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
   819
  by auto
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   820
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   821
lemma InterE [elim]: "A \<in> \<Inter>C \<Longrightarrow> (X \<notin> C \<Longrightarrow> R) \<Longrightarrow> (A \<in> X \<Longrightarrow> R) \<Longrightarrow> R"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   822
  \<comment> \<open>``Classical'' elimination rule -- does not require proving
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   823
    @{prop "X \<in> C"}.\<close>
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   824
  unfolding Inter_eq by blast
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   825
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   826
lemma Inter_lower: "B \<in> A \<Longrightarrow> \<Inter>A \<subseteq> B"
43740
3316e6831801 more succinct proofs
haftmann
parents: 43739
diff changeset
   827
  by (fact Inf_lower)
3316e6831801 more succinct proofs
haftmann
parents: 43739
diff changeset
   828
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   829
lemma Inter_subset: "(\<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
   830
  by (fact Inf_less_eq)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   831
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61799
diff changeset
   832
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
   833
  by (fact Inf_greatest)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   834
44067
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   835
lemma Inter_empty: "\<Inter>{} = UNIV"
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   836
  by (fact Inf_empty) (* already simp *)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   837
44067
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   838
lemma Inter_UNIV: "\<Inter>UNIV = {}"
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   839
  by (fact Inf_UNIV) (* already simp *)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   840
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   841
lemma Inter_insert: "\<Inter>(insert a B) = a \<inter> \<Inter>B"
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   842
  by (fact Inf_insert) (* already simp *)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   843
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   844
lemma Inter_Un_subset: "\<Inter>A \<union> \<Inter>B \<subseteq> \<Inter>(A \<inter> B)"
43899
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   845
  by (fact less_eq_Inf_inter)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   846
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   847
lemma Inter_Un_distrib: "\<Inter>(A \<union> B) = \<Inter>A \<inter> \<Inter>B"
43756
15ea1a07a8cf tuned proofs
haftmann
parents: 43755
diff changeset
   848
  by (fact Inf_union_distrib)
15ea1a07a8cf tuned proofs
haftmann
parents: 43755
diff changeset
   849
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   850
lemma Inter_UNIV_conv [simp]:
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   851
  "\<Inter>A = UNIV \<longleftrightarrow> (\<forall>x\<in>A. x = UNIV)"
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   852
  "UNIV = \<Inter>A \<longleftrightarrow> (\<forall>x\<in>A. x = UNIV)"
43801
097732301fc0 more generalization towards complete lattices
haftmann
parents: 43756
diff changeset
   853
  by (fact Inf_top_conv)+
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   854
43741
fac11b64713c tuned proofs and notation
haftmann
parents: 43740
diff changeset
   855
lemma Inter_anti_mono: "B \<subseteq> A \<Longrightarrow> \<Inter>A \<subseteq> \<Inter>B"
43899
60ef6abb2f92 avoid misunderstandable names
haftmann
parents: 43898
diff changeset
   856
  by (fact Inf_superset_mono)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   857
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   858
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   859
subsubsection \<open>Intersections of families\<close>
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   860
68796
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
   861
abbreviation (input) INTER :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" \<comment> \<open>legacy\<close>
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
   862
  where "INTER \<equiv> INFIMUM"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   863
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
   864
syntax (ASCII)
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
   865
  "_INTER1"     :: "pttrns \<Rightarrow> 'b set \<Rightarrow> 'b set"           ("(3INT _./ _)" [0, 10] 10)
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
   866
  "_INTER"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> 'b set"  ("(3INT _:_./ _)" [0, 0, 10] 10)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   867
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   868
syntax (latex output)
62789
ce15dd971965 explicit property for unbreakable block;
wenzelm
parents: 62390
diff changeset
   869
  "_INTER1"     :: "pttrns \<Rightarrow> 'b set \<Rightarrow> 'b set"           ("(3\<Inter>(\<open>unbreakable\<close>\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
ce15dd971965 explicit property for unbreakable block;
wenzelm
parents: 62390
diff changeset
   870
  "_INTER"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> 'b set"  ("(3\<Inter>(\<open>unbreakable\<close>\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
   871
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
   872
syntax
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
   873
  "_INTER1"     :: "pttrns \<Rightarrow> 'b set \<Rightarrow> 'b set"           ("(3\<Inter>_./ _)" [0, 10] 10)
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
   874
  "_INTER"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> 'b set"  ("(3\<Inter>_\<in>_./ _)" [0, 0, 10] 10)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   875
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   876
translations
68796
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
   877
  "\<Inter>x y. f"  \<rightleftharpoons> "\<Inter>x. \<Inter>y. f"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
   878
  "\<Inter>x. f"    \<rightleftharpoons> "\<Inter>CONST range (\<lambda>x. f)"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
   879
  "\<Inter>x\<in>A. f"  \<rightleftharpoons> "CONST Inter ((\<lambda>x. f) ` A)"
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   880
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   881
lemma INTER_eq: "(\<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
   882
  by (auto intro!: INF_eqI)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   883
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
   884
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
   885
  using Inter_iff [of _ "B ` A"] by simp
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   886
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
   887
lemma INT_I [intro!]: "(\<And>x. x \<in> A \<Longrightarrow> b \<in> B x) \<Longrightarrow> b \<in> (\<Inter>x\<in>A. B x)"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
   888
  by auto
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   889
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
   890
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
   891
  by auto
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   892
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
   893
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"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   894
  \<comment> \<open>"Classical" elimination -- by the Excluded Middle on @{prop "a\<in>A"}.\<close>
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
   895
  by auto
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   896
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   897
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
   898
  by blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   899
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   900
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
   901
  by blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   902
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
   903
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
   904
  by (fact INF_lower)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   905
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
   906
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
   907
  by (fact INF_greatest)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   908
44067
5feac96f0e78 declare {INF,SUP}_empty [simp]
huffman
parents: 44041
diff changeset
   909
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
   910
  by (fact INF_empty)
43854
f1d23df1adde generalized some lemmas
haftmann
parents: 43853
diff changeset
   911
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
   912
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
   913
  by (fact INF_absorb)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   914
43854
f1d23df1adde generalized some lemmas
haftmann
parents: 43853
diff changeset
   915
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
   916
  by (fact le_INF_iff)
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   917
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   918
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
   919
  by (fact INF_insert)
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   920
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   921
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
   922
  by (fact INF_union)
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   923
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   924
lemma INT_insert_distrib: "u \<in> A \<Longrightarrow> (\<Inter>x\<in>A. insert a (B x)) = insert a (\<Inter>x\<in>A. B x)"
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   925
  by blast
43854
f1d23df1adde generalized some lemmas
haftmann
parents: 43853
diff changeset
   926
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   927
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
   928
  by (fact INF_constant)
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   929
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   930
lemma INTER_UNIV_conv:
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   931
  "(UNIV = (\<Inter>x\<in>A. B x)) = (\<forall>x\<in>A. B x = UNIV)"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   932
  "((\<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
   933
  by (fact INF_top_conv)+ (* already simp *)
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   934
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   935
lemma INT_bool_eq: "(\<Inter>b. A b) = A True \<inter> A False"
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
   936
  by (fact INF_UNIV_bool_expand)
43865
db18f4d0cc7d further generalization from sets to complete lattices
haftmann
parents: 43854
diff changeset
   937
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   938
lemma INT_anti_mono: "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)"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   939
  \<comment> \<open>The last inclusion is POSITIVE!\<close>
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
   940
  by (fact INF_superset_mono)
41082
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   941
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   942
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
   943
  by blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   944
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
   945
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
   946
  by blast
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   947
9ff94e7cc3b3 bot comes before top, inf before sup etc.
haftmann
parents: 41080
diff changeset
   948
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   949
subsubsection \<open>Union\<close>
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
   950
68980
5717fbc55521 added spaces because otherwise nonatomic arguments look awful: BIGf x -> BIG f x
nipkow
parents: 68860
diff changeset
   951
abbreviation Union :: "'a set set \<Rightarrow> 'a set"  ("\<Union> _" [900] 900)
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61799
diff changeset
   952
  where "\<Union>S \<equiv> \<Squnion>S"
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
   953
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   954
lemma Union_eq: "\<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
   955
proof (rule set_eqI)
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
   956
  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
   957
  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
   958
    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
   959
  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
   960
    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
   961
qed
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
   962
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   963
lemma Union_iff [simp]: "A \<in> \<Union>C \<longleftrightarrow> (\<exists>X\<in>C. A\<in>X)"
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
   964
  by (unfold Union_eq) blast
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
   965
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   966
lemma UnionI [intro]: "X \<in> C \<Longrightarrow> A \<in> X \<Longrightarrow> A \<in> \<Union>C"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   967
  \<comment> \<open>The order of the premises presupposes that @{term C} is rigid;
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
   968
    @{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
   969
  by auto
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
   970
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
   971
lemma UnionE [elim!]: "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
   972
  by auto
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
   973
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
   974
lemma Union_upper: "B \<in> A \<Longrightarrow> B \<subseteq> \<Union>A"
43901
3ab6c30d256d proof tuning
haftmann
parents: 43900
diff changeset
   975
  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
   976
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
   977
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
   978
  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
   979
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   980
lemma Union_empty: "\<Union>{} = {}"
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   981
  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
   982
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   983
lemma Union_UNIV: "\<Union>UNIV = UNIV"
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   984
  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
   985
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   986
lemma Union_insert: "\<Union>insert a B = a \<union> \<Union>B"
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   987
  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
   988
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
   989
lemma Union_Un_distrib [simp]: "\<Union>(A \<union> B) = \<Union>A \<union> \<Union>B"
43901
3ab6c30d256d proof tuning
haftmann
parents: 43900
diff changeset
   990
  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
   991
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
   992
lemma Union_Int_subset: "\<Union>(A \<inter> B) \<subseteq> \<Union>A \<inter> \<Union>B"
43901
3ab6c30d256d proof tuning
haftmann
parents: 43900
diff changeset
   993
  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
   994
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   995
lemma Union_empty_conv: "(\<Union>A = {}) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   996
  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
   997
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   998
lemma empty_Union_conv: "({} = \<Union>A) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
   999
  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
  1000
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
  1001
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
  1002
  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
  1003
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
  1004
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
  1005
  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
  1006
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1007
lemma Union_mono: "A \<subseteq> B \<Longrightarrow> \<Union>A \<subseteq> \<Union>B"
43901
3ab6c30d256d proof tuning
haftmann
parents: 43900
diff changeset
  1008
  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
  1009
63469
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63365
diff changeset
  1010
lemma Union_subsetI: "(\<And>x. x \<in> A \<Longrightarrow> \<exists>y. y \<in> B \<and> x \<subseteq> y) \<Longrightarrow> \<Union>A \<subseteq> \<Union>B"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63365
diff changeset
  1011
  by blast
32115
8f10fb3bb46e swapped bootstrap order of UNION/Union and INTER/Inter in theory Set
haftmann
parents: 32082
diff changeset
  1012
63879
15bbf6360339 simple new lemmas, mostly about sets
paulson <lp15@cam.ac.uk>
parents: 63820
diff changeset
  1013
lemma disjnt_inj_on_iff:
15bbf6360339 simple new lemmas, mostly about sets
paulson <lp15@cam.ac.uk>
parents: 63820
diff changeset
  1014
     "\<lbrakk>inj_on f (\<Union>\<A>); X \<in> \<A>; Y \<in> \<A>\<rbrakk> \<Longrightarrow> disjnt (f ` X) (f ` Y) \<longleftrightarrow> disjnt X Y"
15bbf6360339 simple new lemmas, mostly about sets
paulson <lp15@cam.ac.uk>
parents: 63820
diff changeset
  1015
  apply (auto simp: disjnt_def)
15bbf6360339 simple new lemmas, mostly about sets
paulson <lp15@cam.ac.uk>
parents: 63820
diff changeset
  1016
  using inj_on_eq_iff by fastforce
15bbf6360339 simple new lemmas, mostly about sets
paulson <lp15@cam.ac.uk>
parents: 63820
diff changeset
  1017
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1018
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1019
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
  1020
68796
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
  1021
abbreviation (input) UNION :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" \<comment> \<open>legacy\<close>
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
  1022
  where "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
  1023
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
  1024
syntax (ASCII)
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34007
diff changeset
  1025
  "_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
  1026
  "_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
  1027
3698947146b2 closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
haftmann
parents: 32064
diff changeset
  1028
syntax (latex output)
62789
ce15dd971965 explicit property for unbreakable block;
wenzelm
parents: 62390
diff changeset
  1029
  "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>(\<open>unbreakable\<close>\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
ce15dd971965 explicit property for unbreakable block;
wenzelm
parents: 62390
diff changeset
  1030
  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>(\<open>unbreakable\<close>\<^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
  1031
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
  1032
syntax
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
  1033
  "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>_./ _)" [0, 10] 10)
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
  1034
  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>_\<in>_./ _)" [0, 0, 10] 10)
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
  1035
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
  1036
translations
68796
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
  1037
  "\<Union>x y. f"   \<rightleftharpoons> "\<Union>x. \<Union>y. f"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
  1038
  "\<Union>x. f"     \<rightleftharpoons> "\<Union>CONST range (\<lambda>x. f)"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
  1039
  "\<Union>x\<in>A. f"   \<rightleftharpoons> "CONST Union ((\<lambda>x. f) ` 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
  1040
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1041
text \<open>
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
  1042
  Note the difference between ordinary syntax of indexed
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
  1043
  unions and intersections (e.g.\ \<open>\<Union>a\<^sub>1\<in>A\<^sub>1. B\<close>)
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61952
diff changeset
  1044
  and their \LaTeX\ rendition: @{term"\<Union>a\<^sub>1\<in>A\<^sub>1. B"}.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1045
\<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
  1046
67673
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
  1047
lemma disjoint_UN_iff: "disjnt A (\<Union>i\<in>I. B i) \<longleftrightarrow> (\<forall>i\<in>I. disjnt A (B i))"
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
  1048
  by (auto simp: disjnt_def)
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
  1049
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1050
lemma UNION_eq: "(\<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
  1051
  by (auto intro!: SUP_eqI)
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1052
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1053
lemma bind_UNION [code]: "Set.bind A f = UNION A f"
45960
e1b09bfb52f1 lattice type class instances for `set`; added code lemma for Set.bind
haftmann
parents: 45013
diff changeset
  1054
  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
  1055
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1056
lemma member_bind [simp]: "x \<in> Set.bind P f \<longleftrightarrow> x \<in> UNION P f "
46036
6a86cc88b02f fundamental theorems on Set.bind
haftmann
parents: 45960
diff changeset
  1057
  by (simp add: bind_UNION)
6a86cc88b02f fundamental theorems on Set.bind
haftmann
parents: 45960
diff changeset
  1058
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60307
diff changeset
  1059
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
  1060
  by blast
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60172
diff changeset
  1061
46036
6a86cc88b02f fundamental theorems on Set.bind
haftmann
parents: 45960
diff changeset
  1062
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
  1063
  using Union_iff [of _ "B ` A"] by simp
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1064
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1065
lemma UN_I [intro]: "a \<in> A \<Longrightarrow> b \<in> B a \<Longrightarrow> b \<in> (\<Union>x\<in>A. B x)"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
  1066
  \<comment> \<open>The order of the premises presupposes that @{term A} is rigid;
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1067
    @{term b} may be flexible.\<close>
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1068
  by auto
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1069
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1070
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"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
  1071
  by auto
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
  1072
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1073
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
  1074
  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
  1075
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1076
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
  1077
  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
  1078
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1079
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
  1080
  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
  1081
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1082
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
  1083
  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
  1084
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1085
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
  1086
  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
  1087
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1088
lemma UN_empty2: "(\<Union>x\<in>A. {}) = {}"
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1089
  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
  1090
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1091
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
  1092
  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
  1093
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
  1094
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
  1095
  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
  1096
44085
a65e26f1427b move legacy candiates to bottom; marked candidates for default simp rules
haftmann
parents: 44084
diff changeset
  1097
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
  1098
  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
  1099
43967
610efb6bda1f more coherent structure in and across theories
haftmann
parents: 43944
diff changeset
  1100
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
  1101
  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
  1102
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
  1103
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
  1104
  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
  1105
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
  1106
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
  1107
  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
  1108
67673
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
  1109
lemma UNION_singleton_eq_range: "(\<Union>x\<in>A. {f x}) = f ` A"
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
  1110
  by blast
c8caefb20564 lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
  1111
43944
b1b436f75070 dropped errorneous hint
haftmann
parents: 43943
diff changeset
  1112
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
  1113
  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
  1114
44920
4657b4c11138 remove some redundant [simp] declarations;
huffman
parents: 44919
diff changeset
  1115
lemma UNION_empty_conv:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1116
  "{} = (\<Union>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = {})"
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1117
  "(\<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
  1118
  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
  1119
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1120
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
  1121
  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
  1122
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
  1123
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
  1124
  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
  1125
43900
7162691e740b generalization; various notation and proof tuning
haftmann
parents: 43899
diff changeset
  1126
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
  1127
  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
  1128
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
  1129
lemma Un_eq_UN: "A \<union> B = (\<Union>b. if b then A else B)"
62390
842917225d56 more canonical names
nipkow
parents: 62343
diff changeset
  1130
  by safe (auto simp add: if_split_mem2)
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
  1131
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1132
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
  1133
  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
  1134
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
  1135
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
  1136
  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
  1137
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
  1138
lemma UN_mono:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1139
  "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
  1140
    (\<Union>x\<in>A. f x) \<subseteq> (\<Union>x\<in>B. g x)"
43940
26ca0bad226a class complete_linorder
haftmann
parents: 43901
diff changeset
  1141
  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
  1142
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1143
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
  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 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
  1147
  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
  1148
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1149
lemma vimage_eq_UN: "f -` B = (\<Union>y\<in>B. f -` {y})"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
  1150
  \<comment> \<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
  1151
  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
  1152
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1153
lemma image_UN: "f ` UNION A B = (\<Union>x\<in>A. f ` B x)"
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1154
  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
  1155
45013
05031b71a89a official status for UN_singleton
haftmann
parents: 44995
diff changeset
  1156
lemma UN_singleton [simp]: "(\<Union>x\<in>A. {x}) = A"
05031b71a89a official status for UN_singleton
haftmann
parents: 44995
diff changeset
  1157
  by blast
05031b71a89a official status for UN_singleton
haftmann
parents: 44995
diff changeset
  1158
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 64966
diff changeset
  1159
lemma inj_on_image: "inj_on f (\<Union>A) \<Longrightarrow> inj_on ((`) f) A"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62789
diff changeset
  1160
  unfolding inj_on_def by blast
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1161
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1162
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1163
subsubsection \<open>Distributive laws\<close>
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1164
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1165
lemma Int_Union: "A \<inter> \<Union>B = (\<Union>C\<in>B. A \<inter> C)"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
  1166
  by blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1167
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1168
lemma Un_Inter: "A \<union> \<Inter>B = (\<Inter>C\<in>B. A \<union> C)"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
  1169
  by blast
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1170
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1171
lemma Int_Union2: "\<Union>B \<inter> A = (\<Union>C\<in>B. C \<inter> A)"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
  1172
  by blast
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1173
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1174
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
  1175
  by (rule sym) (rule INF_inf_distrib)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1176
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1177
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
  1178
  by (rule sym) (rule SUP_sup_distrib)
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1179
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1180
lemma Int_Inter_image: "(\<Inter>x\<in>C. A x \<inter> B x) = \<Inter>(A ` C) \<inter> \<Inter>(B ` C)"  (* FIXME drop *)
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1181
  by (simp add: INT_Int_distrib)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1182
69020
4f94e262976d elimination of near duplication involving Rolle's theorem and the MVT
paulson <lp15@cam.ac.uk>
parents: 68980
diff changeset
  1183
lemma Int_Inter_eq: "A \<inter> \<Inter>\<B> = (if \<B>={} then A else (\<Inter>B\<in>\<B>. A \<inter> B))"
4f94e262976d elimination of near duplication involving Rolle's theorem and the MVT
paulson <lp15@cam.ac.uk>
parents: 68980
diff changeset
  1184
                    "\<Inter>\<B> \<inter> A = (if \<B>={} then A else (\<Inter>B\<in>\<B>. B \<inter> A))"
4f94e262976d elimination of near duplication involving Rolle's theorem and the MVT
paulson <lp15@cam.ac.uk>
parents: 68980
diff changeset
  1185
  by auto
4f94e262976d elimination of near duplication involving Rolle's theorem and the MVT
paulson <lp15@cam.ac.uk>
parents: 68980
diff changeset
  1186
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1187
lemma Un_Union_image: "(\<Union>x\<in>C. A x \<union> B x) = \<Union>(A ` C) \<union> \<Union>(B ` C)"  (* FIXME drop *)
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
  1188
  \<comment> \<open>Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5:\<close>
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
  1189
  \<comment> \<open>Union of a family of unions\<close>
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56076
diff changeset
  1190
  by (simp add: UN_Un_distrib)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1191
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1192
lemma Un_INT_distrib: "B \<union> (\<Inter>i\<in>I. A i) = (\<Inter>i\<in>I. B \<union> A i)"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
  1193
  by blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1194
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1195
lemma Int_UN_distrib: "B \<inter> (\<Union>i\<in>I. A i) = (\<Union>i\<in>I. B \<inter> A i)"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
  1196
  \<comment> \<open>Halmos, Naive Set Theory, page 35.\<close>
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
  1197
  by blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1198
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1199
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)"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
  1200
  by blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1201
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1202
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)"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
  1203
  by blast
44039
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1204
c3d0dac940fc generalized lemmas to complete lattices
haftmann
parents: 44032
diff changeset
  1205
lemma Union_disjoint: "(\<Union>C \<inter> A = {}) \<longleftrightarrow> (\<forall>B\<in>C. B \<inter> A = {})"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
  1206
  by blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1207
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1208
lemma SUP_UNION: "(\<Squnion>x\<in>(\<Union>y\<in>A. g y). f x) = (\<Squnion>y\<in>A. \<Squnion>x\<in>g y. f x :: _ :: complete_lattice)"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1209
  by (rule order_antisym) (blast intro: SUP_least SUP_upper2)+
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1210
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1211
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1212
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
  1213
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1214
lemma inj_on_Inter: "S \<noteq> {} \<Longrightarrow> (\<And>A. A \<in> S \<Longrightarrow> inj_on f A) \<Longrightarrow> inj_on f (\<Inter>S)"
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1215
  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
  1216
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1217
lemma inj_on_INTER: "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> inj_on f (A i)) \<Longrightarrow> inj_on f (\<Inter>i \<in> I. A i)"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
  1218
  unfolding inj_on_def by safe simp
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1219
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1220
lemma inj_on_UNION_chain:
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1221
  assumes chain: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> A i \<le> A j \<or> A j \<le> A i"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1222
    and inj: "\<And>i. i \<in> I \<Longrightarrow> inj_on f (A i)"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60307
diff changeset
  1223
  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
  1224
proof -
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1225
  have "x = y"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1226
    if *: "i \<in> I" "j \<in> I"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1227
    and **: "x \<in> A i" "y \<in> A j"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1228
    and ***: "f x = f y"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1229
    for i j x y
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1230
    using chain [OF *]
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1231
  proof
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1232
    assume "A i \<le> A j"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1233
    with ** have "x \<in> A j" by auto
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1234
    with inj * ** *** show ?thesis
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1235
      by (auto simp add: inj_on_def)
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1236
  next
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1237
    assume "A j \<le> A i"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1238
    with ** have "y \<in> A i" by auto
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1239
    with inj * ** *** show ?thesis
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1240
      by (auto simp add: inj_on_def)
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1241
  qed
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1242
  then show ?thesis
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1243
    by (unfold inj_on_def UNION_eq) auto
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1244
qed
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1245
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1246
lemma bij_betw_UNION_chain:
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1247
  assumes chain: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> A i \<le> A j \<or> A j \<le> A i"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1248
    and bij: "\<And>i. i \<in> I \<Longrightarrow> bij_betw f (A i) (A' i)"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60307
diff changeset
  1249
  shows "bij_betw f (\<Union>i \<in> I. A i) (\<Union>i \<in> I. A' i)"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1250
  unfolding bij_betw_def
63576
ba972a7dbeba tuned proof;
wenzelm
parents: 63575
diff changeset
  1251
proof safe
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1252
  have "\<And>i. i \<in> I \<Longrightarrow> inj_on f (A i)"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1253
    using bij bij_betw_def[of f] by auto
63576
ba972a7dbeba tuned proof;
wenzelm
parents: 63575
diff changeset
  1254
  then show "inj_on f (UNION I A)"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1255
    using chain inj_on_UNION_chain[of I A f] by auto
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1256
next
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1257
  fix i x
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1258
  assume *: "i \<in> I" "x \<in> A i"
63576
ba972a7dbeba tuned proof;
wenzelm
parents: 63575
diff changeset
  1259
  with bij have "f x \<in> A' i"
ba972a7dbeba tuned proof;
wenzelm
parents: 63575
diff changeset
  1260
    by (auto simp: bij_betw_def)
ba972a7dbeba tuned proof;
wenzelm
parents: 63575
diff changeset
  1261
  with * show "f x \<in> UNION I A'" by blast
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1262
next
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1263
  fix i x'
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1264
  assume *: "i \<in> I" "x' \<in> A' i"
63576
ba972a7dbeba tuned proof;
wenzelm
parents: 63575
diff changeset
  1265
  with bij have "\<exists>x \<in> A i. x' = f x"
ba972a7dbeba tuned proof;
wenzelm
parents: 63575
diff changeset
  1266
    unfolding bij_betw_def by blast
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1267
  with * have "\<exists>j \<in> I. \<exists>x \<in> A j. x' = f x"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1268
    by blast
63576
ba972a7dbeba tuned proof;
wenzelm
parents: 63575
diff changeset
  1269
  then show "x' \<in> f ` UNION I A"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1270
    by blast
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1271
qed
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1272
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1273
(*injectivity's required.  Left-to-right inclusion holds even if A is empty*)
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1274
lemma image_INT: "inj_on f C \<Longrightarrow> \<forall>x\<in>A. B x \<subseteq> C \<Longrightarrow> j \<in> A \<Longrightarrow> f ` (INTER A B) = (INT x:A. f ` B x)"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1275
  by (auto simp add: inj_on_def) blast
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1276
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1277
lemma bij_image_INT: "bij f \<Longrightarrow> f ` (INTER A B) = (INT x:A. f ` B x)"
64966
d53d7ca3303e added inj_def (redundant, analogous to surj_def, bij_def);
wenzelm
parents: 63879
diff changeset
  1278
  by (auto simp: bij_def inj_def surj_def) blast
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1279
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1280
lemma UNION_fun_upd: "UNION J (A(i := B)) = UNION (J - {i}) A \<union> (if i \<in> J then B else {})"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62048
diff changeset
  1281
  by (auto simp add: set_eq_iff)
63365
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1282
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1283
lemma bij_betw_Pow:
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1284
  assumes "bij_betw f A B"
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1285
  shows "bij_betw (image f) (Pow A) (Pow B)"
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1286
proof -
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1287
  from assms have "inj_on f A"
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1288
    by (rule bij_betw_imp_inj_on)
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1289
  then have "inj_on f (\<Union>Pow A)"
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1290
    by simp
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1291
  then have "inj_on (image f) (Pow A)"
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1292
    by (rule inj_on_image)
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1293
  then have "bij_betw (image f) (Pow A) (image f ` Pow A)"
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1294
    by (rule inj_on_imp_bij_betw)
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1295
  moreover from assms have "f ` A = B"
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1296
    by (rule bij_betw_imp_surj_on)
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1297
  then have "image f ` Pow A = Pow B"
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1298
    by (rule image_Pow_surj)
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1299
  ultimately show ?thesis by simp
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1300
qed
5340fb6633d0 more theorems
haftmann
parents: 63172
diff changeset
  1301
56015
57e2cfba9c6e bootstrap fundamental Fun theory immediately after Set theory, without dependency on complete lattices
haftmann
parents: 54414
diff changeset
  1302
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1303
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
  1304
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
  1305
lemma Compl_INT [simp]: "- (\<Inter>x\<in>A. B x) = (\<Union>x\<in>A. -B x)"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
  1306
  by blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1307
43873
8a2f339641c1 more on complement
haftmann
parents: 43872
diff changeset
  1308
lemma Compl_UN [simp]: "- (\<Union>x\<in>A. B x) = (\<Inter>x\<in>A. -B x)"
67829
2a6ef5ba4822 Changes to complete distributive lattices due to Viorel Preoteasa
Manuel Eberl <eberlm@in.tum.de>
parents: 67673
diff changeset
  1309
  by blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1310
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1311
subsubsection \<open>Miniscoping and maxiscoping\<close>
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1312
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1313
text \<open>\<^medskip> Miniscoping: pushing in quantifiers and big Unions and Intersections.\<close>
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1314
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1315
lemma UN_simps [simp]:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1316
  "\<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
  1317
  "\<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
  1318
  "\<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
  1319
  "\<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
  1320
  "\<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
  1321
  "\<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
  1322
  "\<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
  1323
  "\<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
  1324
  "\<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
  1325
  "\<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
  1326
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1327
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1328
lemma INT_simps [simp]:
44032
cb768f4ceaf9 solving duality problem for complete_distrib_lattice; tuned
haftmann
parents: 44029
diff changeset
  1329
  "\<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
  1330
  "\<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
  1331
  "\<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
  1332
  "\<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
  1333
  "\<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
  1334
  "\<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
  1335
  "\<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
  1336
  "\<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
  1337
  "\<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
  1338
  "\<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
  1339
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1340
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1341
lemma UN_ball_bex_simps [simp]:
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1342
  "\<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
  1343
  "\<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
  1344
  "\<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
  1345
  "\<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
  1346
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1347
43943
e6928fc2332a moved some lemmas
haftmann
parents: 43940
diff changeset
  1348
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63469
diff changeset
  1349
text \<open>\<^medskip> Maxiscoping: pulling out big Unions and Intersections.\<close>
13860
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1350
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1351
lemma UN_extend_simps:
43817
d53350bc65a4 tuned notation
haftmann
parents: 43814
diff changeset
  1352
  "\<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
  1353
  "\<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
  1354
  "\<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
  1355
  "\<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
  1356
  "\<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
  1357
  "\<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
  1358
  "\<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
  1359
  "\<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
  1360
  "\<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
  1361
  "\<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
  1362
  by auto
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1363
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1364
lemma INT_extend_simps:
43852
7411fbf0a325 tuned notation
haftmann
parents: 43831
diff changeset
  1365
  "\<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
  1366
  "\<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
  1367
  "\<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
  1368
  "\<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
  1369
  "\<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
  1370
  "\<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
  1371
  "\<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
  1372
  "\<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
  1373
  "\<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
  1374
  "\<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
  1375
  by auto
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1376
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60585
diff changeset
  1377
text \<open>Finally\<close>
43872
6b917e5877d2 more consistent theorem names
haftmann
parents: 43871
diff changeset
  1378
30596
140b22f22071 tuned some theorem and attribute bindings
haftmann
parents: 30531
diff changeset
  1379
lemmas mem_simps =
140b22f22071 tuned some theorem and attribute bindings
haftmann
parents: 30531
diff changeset
  1380
  insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff
140b22f22071 tuned some theorem and attribute bindings
haftmann
parents: 30531
diff changeset
  1381
  mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
  1382
  \<comment> \<open>Each of these has ALREADY been added \<open>[simp]\<close> above.\<close>
21669
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21549
diff changeset
  1383
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1384
end