| author | wenzelm | 
| Sun, 30 Jan 2011 18:36:35 +0100 | |
| changeset 41649 | 57181bb1dfe0 | 
| parent 41107 | 8795cd75965e | 
| child 42163 | 392fd6c4669c | 
| permissions | -rw-r--r-- | 
| 32139 | 1 | (* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel *) | 
| 923 | 2 | |
| 11979 | 3 | header {* Set theory for higher-order logic *}
 | 
| 4 | ||
| 15131 | 5 | theory Set | 
| 30304 
d8e4cd2ac2a1
set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
 haftmann parents: 
29901diff
changeset | 6 | imports Lattices | 
| 15131 | 7 | begin | 
| 11979 | 8 | |
| 32081 | 9 | subsection {* Sets as predicates *}
 | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 10 | |
| 37677 | 11 | types 'a set = "'a \<Rightarrow> bool" | 
| 3820 | 12 | |
| 37677 | 13 | definition Collect :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set" where -- "comprehension"
 | 
| 14 | "Collect P = P" | |
| 30304 
d8e4cd2ac2a1
set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
 haftmann parents: 
29901diff
changeset | 15 | |
| 37677 | 16 | definition member :: "'a \<Rightarrow> 'a set \<Rightarrow> bool" where -- "membership" | 
| 17 | mem_def: "member x A = A x" | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 18 | |
| 21210 | 19 | notation | 
| 37677 | 20 |   member  ("op :") and
 | 
| 21 |   member  ("(_/ : _)" [50, 51] 50)
 | |
| 11979 | 22 | |
| 37677 | 23 | abbreviation not_member where | 
| 24 | "not_member x A \<equiv> ~ (x : A)" -- "non-membership" | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 25 | |
| 21210 | 26 | notation | 
| 37677 | 27 |   not_member  ("op ~:") and
 | 
| 28 |   not_member  ("(_/ ~: _)" [50, 51] 50)
 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 29 | |
| 21210 | 30 | notation (xsymbols) | 
| 37677 | 31 |   member      ("op \<in>") and
 | 
| 32 |   member      ("(_/ \<in> _)" [50, 51] 50) and
 | |
| 33 |   not_member  ("op \<notin>") and
 | |
| 34 |   not_member  ("(_/ \<notin> _)" [50, 51] 50)
 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 35 | |
| 21210 | 36 | notation (HTML output) | 
| 37677 | 37 |   member      ("op \<in>") and
 | 
| 38 |   member      ("(_/ \<in> _)" [50, 51] 50) and
 | |
| 39 |   not_member  ("op \<notin>") and
 | |
| 40 |   not_member  ("(_/ \<notin> _)" [50, 51] 50)
 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 41 | |
| 41107 | 42 | |
| 43 | ||
| 32081 | 44 | text {* Set comprehensions *}
 | 
| 45 | ||
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 46 | syntax | 
| 35115 | 47 |   "_Coll" :: "pttrn => bool => 'a set"    ("(1{_./ _})")
 | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 48 | translations | 
| 35115 | 49 |   "{x. P}" == "CONST Collect (%x. P)"
 | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 50 | |
| 32081 | 51 | syntax | 
| 35115 | 52 |   "_Collect" :: "idt => 'a set => bool => 'a set"    ("(1{_ :/ _./ _})")
 | 
| 32081 | 53 | syntax (xsymbols) | 
| 35115 | 54 |   "_Collect" :: "idt => 'a set => bool => 'a set"    ("(1{_ \<in>/ _./ _})")
 | 
| 32081 | 55 | translations | 
| 35115 | 56 |   "{x:A. P}" => "{x. x:A & P}"
 | 
| 32081 | 57 | |
| 41107 | 58 | lemma mem_Collect_eq [iff]: "a \<in> {x. P x} = P a"
 | 
| 32081 | 59 | by (simp add: Collect_def mem_def) | 
| 60 | ||
| 41107 | 61 | lemma Collect_mem_eq [simp]: "{x. x \<in> A} = A"
 | 
| 32081 | 62 | by (simp add: Collect_def mem_def) | 
| 63 | ||
| 41107 | 64 | lemma CollectI: "P a \<Longrightarrow> a \<in> {x. P x}"
 | 
| 32081 | 65 | by simp | 
| 66 | ||
| 41107 | 67 | lemma CollectD: "a \<in> {x. P x} \<Longrightarrow> P a"
 | 
| 32081 | 68 | by simp | 
| 69 | ||
| 41107 | 70 | lemma Collect_cong: "(\<And>x. P x = Q x) ==> {x. P x} = {x. Q x}"
 | 
| 32081 | 71 | by simp | 
| 72 | ||
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 73 | text {*
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 74 | Simproc for pulling @{text "x=t"} in @{text "{x. \<dots> & x=t & \<dots>}"}
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 75 | to the front (and similarly for @{text "t=x"}):
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 76 | *} | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 77 | |
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 78 | setup {*
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 79 | let | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 80 |   val Coll_perm_tac = rtac @{thm Collect_cong} 1 THEN rtac @{thm iffI} 1 THEN
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 81 |     ALLGOALS(EVERY'[REPEAT_DETERM o (etac @{thm conjE}),
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 82 |                     DEPTH_SOLVE_1 o (ares_tac [@{thm conjI}])])
 | 
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
38649diff
changeset | 83 |   val defColl_regroup = Simplifier.simproc_global @{theory}
 | 
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 84 |     "defined Collect" ["{x. P x & Q x}"]
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 85 | (Quantifier1.rearrange_Coll Coll_perm_tac) | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 86 | in | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 87 | Simplifier.map_simpset (fn ss => ss addsimprocs [defColl_regroup]) | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 88 | end | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 89 | *} | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 90 | |
| 32081 | 91 | lemmas CollectE = CollectD [elim_format] | 
| 92 | ||
| 41107 | 93 | lemma set_eqI: | 
| 94 | assumes "\<And>x. x \<in> A \<longleftrightarrow> x \<in> B" | |
| 95 | shows "A = B" | |
| 96 | proof - | |
| 97 |   from assms have "{x. x \<in> A} = {x. x \<in> B}" by simp
 | |
| 98 | then show ?thesis by simp | |
| 99 | qed | |
| 100 | ||
| 101 | lemma set_eq_iff [no_atp]: | |
| 102 | "A = B \<longleftrightarrow> (\<forall>x. x \<in> A \<longleftrightarrow> x \<in> B)" | |
| 103 | by (auto intro:set_eqI) | |
| 104 | ||
| 32081 | 105 | text {* Set enumerations *}
 | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 106 | |
| 32264 
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
 haftmann parents: 
32139diff
changeset | 107 | abbreviation empty :: "'a set" ("{}") where
 | 
| 
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
 haftmann parents: 
32139diff
changeset | 108 |   "{} \<equiv> bot"
 | 
| 31456 | 109 | |
| 110 | definition insert :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where | |
| 32081 | 111 |   insert_compr: "insert a B = {x. x = a \<or> x \<in> B}"
 | 
| 31456 | 112 | |
| 113 | syntax | |
| 35115 | 114 |   "_Finset" :: "args => 'a set"    ("{(_)}")
 | 
| 31456 | 115 | translations | 
| 35115 | 116 |   "{x, xs}" == "CONST insert x {xs}"
 | 
| 117 |   "{x}" == "CONST insert x {}"
 | |
| 31456 | 118 | |
| 32081 | 119 | |
| 120 | subsection {* Subsets and bounded quantifiers *}
 | |
| 121 | ||
| 122 | abbreviation | |
| 123 | subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where | |
| 124 | "subset \<equiv> less" | |
| 125 | ||
| 126 | abbreviation | |
| 127 | subset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where | |
| 128 | "subset_eq \<equiv> less_eq" | |
| 129 | ||
| 130 | notation (output) | |
| 131 |   subset  ("op <") and
 | |
| 132 |   subset  ("(_/ < _)" [50, 51] 50) and
 | |
| 133 |   subset_eq  ("op <=") and
 | |
| 134 |   subset_eq  ("(_/ <= _)" [50, 51] 50)
 | |
| 135 | ||
| 136 | notation (xsymbols) | |
| 137 |   subset  ("op \<subset>") and
 | |
| 138 |   subset  ("(_/ \<subset> _)" [50, 51] 50) and
 | |
| 139 |   subset_eq  ("op \<subseteq>") and
 | |
| 140 |   subset_eq  ("(_/ \<subseteq> _)" [50, 51] 50)
 | |
| 141 | ||
| 142 | notation (HTML output) | |
| 143 |   subset  ("op \<subset>") and
 | |
| 144 |   subset  ("(_/ \<subset> _)" [50, 51] 50) and
 | |
| 145 |   subset_eq  ("op \<subseteq>") and
 | |
| 146 |   subset_eq  ("(_/ \<subseteq> _)" [50, 51] 50)
 | |
| 147 | ||
| 148 | abbreviation (input) | |
| 149 | supset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where | |
| 150 | "supset \<equiv> greater" | |
| 151 | ||
| 152 | abbreviation (input) | |
| 153 | supset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where | |
| 154 | "supset_eq \<equiv> greater_eq" | |
| 155 | ||
| 156 | notation (xsymbols) | |
| 157 |   supset  ("op \<supset>") and
 | |
| 158 |   supset  ("(_/ \<supset> _)" [50, 51] 50) and
 | |
| 159 |   supset_eq  ("op \<supseteq>") and
 | |
| 160 |   supset_eq  ("(_/ \<supseteq> _)" [50, 51] 50)
 | |
| 161 | ||
| 37387 
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
 haftmann parents: 
36009diff
changeset | 162 | definition Ball :: "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
 haftmann parents: 
36009diff
changeset | 163 | "Ball A P \<longleftrightarrow> (\<forall>x. x \<in> A \<longrightarrow> P x)" -- "bounded universal quantifiers" | 
| 32077 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 164 | |
| 37387 
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
 haftmann parents: 
36009diff
changeset | 165 | definition Bex :: "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
 haftmann parents: 
36009diff
changeset | 166 | "Bex A P \<longleftrightarrow> (\<exists>x. x \<in> A \<and> P x)" -- "bounded existential quantifiers" | 
| 32077 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 167 | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 168 | syntax | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 169 |   "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3ALL _:_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 170 |   "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3EX _:_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 171 |   "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3EX! _:_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 172 |   "_Bleast"     :: "id => 'a set => bool => 'a"           ("(3LEAST _:_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 173 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 174 | syntax (HOL) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 175 |   "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3! _:_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 176 |   "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3? _:_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 177 |   "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3?! _:_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 178 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 179 | syntax (xsymbols) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 180 |   "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3\<forall>_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 181 |   "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3\<exists>_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 182 |   "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3\<exists>!_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 183 |   "_Bleast"     :: "id => 'a set => bool => 'a"           ("(3LEAST_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 184 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 185 | syntax (HTML output) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 186 |   "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3\<forall>_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 187 |   "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3\<exists>_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 188 |   "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3\<exists>!_\<in>_./ _)" [0, 0, 10] 10)
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 189 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 190 | translations | 
| 35115 | 191 | "ALL x:A. P" == "CONST Ball A (%x. P)" | 
| 192 | "EX x:A. P" == "CONST Bex A (%x. P)" | |
| 193 | "EX! x:A. P" => "EX! x. x:A & P" | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 194 | "LEAST x:A. P" => "LEAST x. x:A & P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 195 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 196 | syntax (output) | 
| 14804 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 197 |   "_setlessAll" :: "[idt, 'a, bool] => bool"  ("(3ALL _<_./ _)"  [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 198 |   "_setlessEx"  :: "[idt, 'a, bool] => bool"  ("(3EX _<_./ _)"  [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 199 |   "_setleAll"   :: "[idt, 'a, bool] => bool"  ("(3ALL _<=_./ _)" [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 200 |   "_setleEx"    :: "[idt, 'a, bool] => bool"  ("(3EX _<=_./ _)" [0, 0, 10] 10)
 | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19870diff
changeset | 201 |   "_setleEx1"   :: "[idt, 'a, bool] => bool"  ("(3EX! _<=_./ _)" [0, 0, 10] 10)
 | 
| 14804 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 202 | |
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 203 | syntax (xsymbols) | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 204 |   "_setlessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<subset>_./ _)"  [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 205 |   "_setlessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<subset>_./ _)"  [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 206 |   "_setleAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<subseteq>_./ _)" [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 207 |   "_setleEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<subseteq>_./ _)" [0, 0, 10] 10)
 | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19870diff
changeset | 208 |   "_setleEx1"   :: "[idt, 'a, bool] => bool"   ("(3\<exists>!_\<subseteq>_./ _)" [0, 0, 10] 10)
 | 
| 14804 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 209 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19637diff
changeset | 210 | syntax (HOL output) | 
| 14804 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 211 |   "_setlessAll" :: "[idt, 'a, bool] => bool"   ("(3! _<_./ _)"  [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 212 |   "_setlessEx"  :: "[idt, 'a, bool] => bool"   ("(3? _<_./ _)"  [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 213 |   "_setleAll"   :: "[idt, 'a, bool] => bool"   ("(3! _<=_./ _)" [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 214 |   "_setleEx"    :: "[idt, 'a, bool] => bool"   ("(3? _<=_./ _)" [0, 0, 10] 10)
 | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19870diff
changeset | 215 |   "_setleEx1"   :: "[idt, 'a, bool] => bool"   ("(3?! _<=_./ _)" [0, 0, 10] 10)
 | 
| 14804 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 216 | |
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 217 | syntax (HTML output) | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 218 |   "_setlessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<subset>_./ _)"  [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 219 |   "_setlessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<subset>_./ _)"  [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 220 |   "_setleAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<subseteq>_./ _)" [0, 0, 10] 10)
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 221 |   "_setleEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<subseteq>_./ _)" [0, 0, 10] 10)
 | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19870diff
changeset | 222 |   "_setleEx1"   :: "[idt, 'a, bool] => bool"   ("(3\<exists>!_\<subseteq>_./ _)" [0, 0, 10] 10)
 | 
| 14804 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 223 | |
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 224 | translations | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 225 | "\<forall>A\<subset>B. P" => "ALL A. A \<subset> B --> P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 226 | "\<exists>A\<subset>B. P" => "EX A. A \<subset> B & P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 227 | "\<forall>A\<subseteq>B. P" => "ALL A. A \<subseteq> B --> P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 228 | "\<exists>A\<subseteq>B. P" => "EX A. A \<subseteq> B & P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 229 | "\<exists>!A\<subseteq>B. P" => "EX! A. A \<subseteq> B & P" | 
| 14804 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 230 | |
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 231 | print_translation {*
 | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 232 | let | 
| 35115 | 233 |   val Type (set_type, _) = @{typ "'a set"};   (* FIXME 'a => bool (!?!) *)
 | 
| 234 |   val All_binder = Syntax.binder_name @{const_syntax All};
 | |
| 235 |   val Ex_binder = Syntax.binder_name @{const_syntax Ex};
 | |
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38715diff
changeset | 236 |   val impl = @{const_syntax HOL.implies};
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 237 |   val conj = @{const_syntax HOL.conj};
 | 
| 35115 | 238 |   val sbset = @{const_syntax subset};
 | 
| 239 |   val sbset_eq = @{const_syntax subset_eq};
 | |
| 21819 | 240 | |
| 241 | val trans = | |
| 35115 | 242 |    [((All_binder, impl, sbset), @{syntax_const "_setlessAll"}),
 | 
| 243 |     ((All_binder, impl, sbset_eq), @{syntax_const "_setleAll"}),
 | |
| 244 |     ((Ex_binder, conj, sbset), @{syntax_const "_setlessEx"}),
 | |
| 245 |     ((Ex_binder, conj, sbset_eq), @{syntax_const "_setleEx"})];
 | |
| 21819 | 246 | |
| 247 | fun mk v v' c n P = | |
| 248 | if v = v' andalso not (Term.exists_subterm (fn Free (x, _) => x = v | _ => false) n) | |
| 249 | then Syntax.const c $ Syntax.mark_bound v' $ n $ P else raise Match; | |
| 250 | ||
| 251 | fun tr' q = (q, | |
| 35115 | 252 |         fn [Const (@{syntax_const "_bound"}, _) $ Free (v, Type (T, _)),
 | 
| 253 | Const (c, _) $ | |
| 254 |               (Const (d, _) $ (Const (@{syntax_const "_bound"}, _) $ Free (v', _)) $ n) $ P] =>
 | |
| 255 | if T = set_type then | |
| 256 | (case AList.lookup (op =) trans (q, c, d) of | |
| 257 | NONE => raise Match | |
| 258 | | SOME l => mk v v' l n P) | |
| 259 | else raise Match | |
| 260 | | _ => raise Match); | |
| 14804 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 261 | in | 
| 21819 | 262 | [tr' All_binder, tr' Ex_binder] | 
| 14804 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 263 | end | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 264 | *} | 
| 
8de39d3e8eb6
Corrected printer bug for bounded quantifiers Q x<=y. P
 nipkow parents: 
14752diff
changeset | 265 | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 266 | |
| 11979 | 267 | text {*
 | 
| 268 |   \medskip Translate between @{text "{e | x1...xn. P}"} and @{text
 | |
| 269 |   "{u. EX x1..xn. u = e & P}"}; @{text "{y. EX x1..xn. y = e & P}"} is
 | |
| 270 |   only translated if @{text "[0..n] subset bvs(e)"}.
 | |
| 271 | *} | |
| 272 | ||
| 35115 | 273 | syntax | 
| 274 |   "_Setcompr" :: "'a => idts => bool => 'a set"    ("(1{_ |/_./ _})")
 | |
| 275 | ||
| 11979 | 276 | parse_translation {*
 | 
| 277 | let | |
| 35115 | 278 |     val ex_tr = snd (mk_binder_tr ("EX ", @{const_syntax Ex}));
 | 
| 3947 | 279 | |
| 35115 | 280 |     fun nvars (Const (@{syntax_const "_idts"}, _) $ _ $ idts) = nvars idts + 1
 | 
| 11979 | 281 | | nvars _ = 1; | 
| 282 | ||
| 283 | fun setcompr_tr [e, idts, b] = | |
| 284 | let | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 285 |         val eq = Syntax.const @{const_syntax HOL.eq} $ Bound (nvars idts) $ e;
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 286 |         val P = Syntax.const @{const_syntax HOL.conj} $ eq $ b;
 | 
| 11979 | 287 | val exP = ex_tr [idts, P]; | 
| 35115 | 288 |       in Syntax.const @{const_syntax Collect} $ Term.absdummy (dummyT, exP) end;
 | 
| 11979 | 289 | |
| 35115 | 290 |   in [(@{syntax_const "_Setcompr"}, setcompr_tr)] end;
 | 
| 11979 | 291 | *} | 
| 923 | 292 | |
| 35115 | 293 | print_translation {*
 | 
| 294 |  [Syntax.preserve_binder_abs2_tr' @{const_syntax Ball} @{syntax_const "_Ball"},
 | |
| 295 |   Syntax.preserve_binder_abs2_tr' @{const_syntax Bex} @{syntax_const "_Bex"}]
 | |
| 296 | *} -- {* to avoid eta-contraction of body *}
 | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 297 | |
| 13763 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13653diff
changeset | 298 | print_translation {*
 | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13653diff
changeset | 299 | let | 
| 35115 | 300 |   val ex_tr' = snd (mk_binder_tr' (@{const_syntax Ex}, "DUMMY"));
 | 
| 13763 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13653diff
changeset | 301 | |
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13653diff
changeset | 302 | fun setcompr_tr' [Abs (abs as (_, _, P))] = | 
| 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13653diff
changeset | 303 | let | 
| 35115 | 304 |       fun check (Const (@{const_syntax Ex}, _) $ Abs (_, _, P), n) = check (P, n + 1)
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 305 |         | check (Const (@{const_syntax HOL.conj}, _) $
 | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 306 |               (Const (@{const_syntax HOL.eq}, _) $ Bound m $ e) $ P, n) =
 | 
| 13763 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13653diff
changeset | 307 | n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso | 
| 33038 | 308 | subset (op =) (0 upto (n - 1), add_loose_bnos (e, 0, [])) | 
| 35115 | 309 | | check _ = false; | 
| 923 | 310 | |
| 11979 | 311 | fun tr' (_ $ abs) = | 
| 312 | let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr' [abs] | |
| 35115 | 313 |           in Syntax.const @{syntax_const "_Setcompr"} $ e $ idts $ Q end;
 | 
| 314 | in | |
| 315 | if check (P, 0) then tr' P | |
| 316 | else | |
| 317 | let | |
| 318 | val (x as _ $ Free(xN, _), t) = atomic_abs_tr' abs; | |
| 319 |           val M = Syntax.const @{syntax_const "_Coll"} $ x $ t;
 | |
| 320 | in | |
| 321 | case t of | |
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 322 |             Const (@{const_syntax HOL.conj}, _) $
 | 
| 37677 | 323 |               (Const (@{const_syntax Set.member}, _) $
 | 
| 35115 | 324 |                 (Const (@{syntax_const "_bound"}, _) $ Free (yN, _)) $ A) $ P =>
 | 
| 325 |             if xN = yN then Syntax.const @{syntax_const "_Collect"} $ x $ A $ P else M
 | |
| 326 | | _ => M | |
| 327 | end | |
| 13763 
f94b569cd610
added print translations tha avoid eta contraction for important binders.
 nipkow parents: 
13653diff
changeset | 328 | end; | 
| 35115 | 329 |   in [(@{const_syntax Collect}, setcompr_tr')] end;
 | 
| 11979 | 330 | *} | 
| 331 | ||
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 332 | setup {*
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 333 | let | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 334 |   val unfold_bex_tac = unfold_tac @{thms "Bex_def"};
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 335 | fun prove_bex_tac ss = unfold_bex_tac ss THEN Quantifier1.prove_one_point_ex_tac; | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 336 | val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac; | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 337 |   val unfold_ball_tac = unfold_tac @{thms "Ball_def"};
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 338 | fun prove_ball_tac ss = unfold_ball_tac ss THEN Quantifier1.prove_one_point_all_tac; | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 339 | val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac; | 
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
38649diff
changeset | 340 |   val defBEX_regroup = Simplifier.simproc_global @{theory}
 | 
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 341 | "defined BEX" ["EX x:A. P x & Q x"] rearrange_bex; | 
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
38649diff
changeset | 342 |   val defBALL_regroup = Simplifier.simproc_global @{theory}
 | 
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 343 | "defined BALL" ["ALL x:A. P x --> Q x"] rearrange_ball; | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 344 | in | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 345 | Simplifier.map_simpset (fn ss => ss addsimprocs [defBALL_regroup, defBEX_regroup]) | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 346 | end | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 347 | *} | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 348 | |
| 11979 | 349 | lemma ballI [intro!]: "(!!x. x:A ==> P x) ==> ALL x:A. P x" | 
| 350 | by (simp add: Ball_def) | |
| 351 | ||
| 352 | lemmas strip = impI allI ballI | |
| 353 | ||
| 354 | lemma bspec [dest?]: "ALL x:A. P x ==> x:A ==> P x" | |
| 355 | by (simp add: Ball_def) | |
| 356 | ||
| 357 | text {*
 | |
| 358 | Gives better instantiation for bound: | |
| 359 | *} | |
| 360 | ||
| 26339 | 361 | declaration {* fn _ =>
 | 
| 362 |   Classical.map_cs (fn cs => cs addbefore ("bspec", datac @{thm bspec} 1))
 | |
| 11979 | 363 | *} | 
| 364 | ||
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 365 | ML {*
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 366 | structure Simpdata = | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 367 | struct | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 368 | |
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 369 | open Simpdata; | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 370 | |
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 371 | val mksimps_pairs = [(@{const_name Ball}, @{thms bspec})] @ mksimps_pairs;
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 372 | |
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 373 | end; | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 374 | |
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 375 | open Simpdata; | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 376 | *} | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 377 | |
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 378 | declaration {* fn _ =>
 | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 379 | Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 380 | *} | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 381 | |
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 382 | lemma ballE [elim]: "ALL x:A. P x ==> (P x ==> Q) ==> (x ~: A ==> Q) ==> Q" | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 383 | by (unfold Ball_def) blast | 
| 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 384 | |
| 11979 | 385 | lemma bexI [intro]: "P x ==> x:A ==> EX x:A. P x" | 
| 386 |   -- {* Normally the best argument order: @{prop "P x"} constrains the
 | |
| 387 |     choice of @{prop "x:A"}. *}
 | |
| 388 | by (unfold Bex_def) blast | |
| 389 | ||
| 13113 | 390 | lemma rev_bexI [intro?]: "x:A ==> P x ==> EX x:A. P x" | 
| 11979 | 391 |   -- {* The best argument order when there is only one @{prop "x:A"}. *}
 | 
| 392 | by (unfold Bex_def) blast | |
| 393 | ||
| 394 | lemma bexCI: "(ALL x:A. ~P x ==> P a) ==> a:A ==> EX x:A. P x" | |
| 395 | by (unfold Bex_def) blast | |
| 396 | ||
| 397 | lemma bexE [elim!]: "EX x:A. P x ==> (!!x. x:A ==> P x ==> Q) ==> Q" | |
| 398 | by (unfold Bex_def) blast | |
| 399 | ||
| 400 | lemma ball_triv [simp]: "(ALL x:A. P) = ((EX x. x:A) --> P)" | |
| 401 |   -- {* Trival rewrite rule. *}
 | |
| 402 | by (simp add: Ball_def) | |
| 403 | ||
| 404 | lemma bex_triv [simp]: "(EX x:A. P) = ((EX x. x:A) & P)" | |
| 405 |   -- {* Dual form for existentials. *}
 | |
| 406 | by (simp add: Bex_def) | |
| 407 | ||
| 408 | lemma bex_triv_one_point1 [simp]: "(EX x:A. x = a) = (a:A)" | |
| 409 | by blast | |
| 410 | ||
| 411 | lemma bex_triv_one_point2 [simp]: "(EX x:A. a = x) = (a:A)" | |
| 412 | by blast | |
| 413 | ||
| 414 | lemma bex_one_point1 [simp]: "(EX x:A. x = a & P x) = (a:A & P a)" | |
| 415 | by blast | |
| 416 | ||
| 417 | lemma bex_one_point2 [simp]: "(EX x:A. a = x & P x) = (a:A & P a)" | |
| 418 | by blast | |
| 419 | ||
| 420 | lemma ball_one_point1 [simp]: "(ALL x:A. x = a --> P x) = (a:A --> P a)" | |
| 421 | by blast | |
| 422 | ||
| 423 | lemma ball_one_point2 [simp]: "(ALL x:A. a = x --> P x) = (a:A --> P a)" | |
| 424 | by blast | |
| 425 | ||
| 426 | ||
| 32081 | 427 | text {* Congruence rules *}
 | 
| 11979 | 428 | |
| 16636 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 429 | lemma ball_cong: | 
| 11979 | 430 | "A = B ==> (!!x. x:B ==> P x = Q x) ==> | 
| 431 | (ALL x:A. P x) = (ALL x:B. Q x)" | |
| 432 | by (simp add: Ball_def) | |
| 433 | ||
| 16636 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 434 | lemma strong_ball_cong [cong]: | 
| 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 435 | "A = B ==> (!!x. x:B =simp=> P x = Q x) ==> | 
| 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 436 | (ALL x:A. P x) = (ALL x:B. Q x)" | 
| 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 437 | by (simp add: simp_implies_def Ball_def) | 
| 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 438 | |
| 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 439 | lemma bex_cong: | 
| 11979 | 440 | "A = B ==> (!!x. x:B ==> P x = Q x) ==> | 
| 441 | (EX x:A. P x) = (EX x:B. Q x)" | |
| 442 | by (simp add: Bex_def cong: conj_cong) | |
| 1273 | 443 | |
| 16636 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 444 | lemma strong_bex_cong [cong]: | 
| 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 445 | "A = B ==> (!!x. x:B =simp=> P x = Q x) ==> | 
| 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 446 | (EX x:A. P x) = (EX x:B. Q x)" | 
| 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 447 | by (simp add: simp_implies_def Bex_def cong: conj_cong) | 
| 
1ed737a98198
Added strong_ball_cong and strong_bex_cong (these are now the standard
 berghofe parents: 
15950diff
changeset | 448 | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 449 | |
| 32081 | 450 | subsection {* Basic operations *}
 | 
| 451 | ||
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 452 | subsubsection {* Subsets *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 453 | |
| 33022 
c95102496490
 Removal of the unused atpset concept, the atp attribute and some related code.
 paulson parents: 
32888diff
changeset | 454 | lemma subsetI [intro!]: "(\<And>x. x \<in> A \<Longrightarrow> x \<in> B) \<Longrightarrow> A \<subseteq> B" | 
| 32888 | 455 | unfolding mem_def by (rule le_funI, rule le_boolI) | 
| 30352 | 456 | |
| 11979 | 457 | text {*
 | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 458 |   \medskip Map the type @{text "'a set => anything"} to just @{typ
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 459 |   'a}; for overloading constants whose first argument has type @{typ
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 460 | "'a set"}. | 
| 11979 | 461 | *} | 
| 462 | ||
| 30596 | 463 | lemma subsetD [elim, intro?]: "A \<subseteq> B ==> c \<in> A ==> c \<in> B" | 
| 32888 | 464 | unfolding mem_def by (erule le_funE, erule le_boolE) | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 465 |   -- {* Rule in Modus Ponens style. *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 466 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 467 | lemma rev_subsetD [no_atp,intro?]: "c \<in> A ==> A \<subseteq> B ==> c \<in> B" | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 468 |   -- {* The same, with reversed premises for use with @{text erule} --
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 469 |       cf @{text rev_mp}. *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 470 | by (rule subsetD) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 471 | |
| 11979 | 472 | text {*
 | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 473 |   \medskip Converts @{prop "A \<subseteq> B"} to @{prop "x \<in> A ==> x \<in> B"}.
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 474 | *} | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 475 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 476 | lemma subsetCE [no_atp,elim]: "A \<subseteq> B ==> (c \<notin> A ==> P) ==> (c \<in> B ==> P) ==> P" | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 477 |   -- {* Classical elimination rule. *}
 | 
| 32888 | 478 | unfolding mem_def by (blast dest: le_funE le_boolE) | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 479 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 480 | lemma subset_eq [no_atp]: "A \<le> B = (\<forall>x\<in>A. x \<in> B)" by blast | 
| 2388 | 481 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 482 | lemma contra_subsetD [no_atp]: "A \<subseteq> B ==> c \<notin> B ==> c \<notin> A" | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 483 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 484 | |
| 33022 
c95102496490
 Removal of the unused atpset concept, the atp attribute and some related code.
 paulson parents: 
32888diff
changeset | 485 | lemma subset_refl [simp]: "A \<subseteq> A" | 
| 32081 | 486 | by (fact order_refl) | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 487 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 488 | lemma subset_trans: "A \<subseteq> B ==> B \<subseteq> C ==> A \<subseteq> C" | 
| 32081 | 489 | by (fact order_trans) | 
| 490 | ||
| 491 | lemma set_rev_mp: "x:A ==> A \<subseteq> B ==> x:B" | |
| 492 | by (rule subsetD) | |
| 493 | ||
| 494 | lemma set_mp: "A \<subseteq> B ==> x:A ==> x:B" | |
| 495 | by (rule subsetD) | |
| 496 | ||
| 33044 | 497 | lemma eq_mem_trans: "a=b ==> b \<in> A ==> a \<in> A" | 
| 498 | by simp | |
| 499 | ||
| 32081 | 500 | lemmas basic_trans_rules [trans] = | 
| 33044 | 501 | order_trans_rules set_rev_mp set_mp eq_mem_trans | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 502 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 503 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 504 | subsubsection {* Equality *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 505 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 506 | lemma subset_antisym [intro!]: "A \<subseteq> B ==> B \<subseteq> A ==> A = B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 507 |   -- {* Anti-symmetry of the subset relation. *}
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39213diff
changeset | 508 | by (iprover intro: set_eqI subsetD) | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 509 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 510 | text {*
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 511 | \medskip Equality rules from ZF set theory -- are they appropriate | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 512 | here? | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 513 | *} | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 514 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 515 | lemma equalityD1: "A = B ==> A \<subseteq> B" | 
| 34209 | 516 | by simp | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 517 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 518 | lemma equalityD2: "A = B ==> B \<subseteq> A" | 
| 34209 | 519 | by simp | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 520 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 521 | text {*
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 522 |   \medskip Be careful when adding this to the claset as @{text
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 523 |   subset_empty} is in the simpset: @{prop "A = {}"} goes to @{prop "{}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 524 |   \<subseteq> A"} and @{prop "A \<subseteq> {}"} and then back to @{prop "A = {}"}!
 | 
| 30352 | 525 | *} | 
| 526 | ||
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 527 | lemma equalityE: "A = B ==> (A \<subseteq> B ==> B \<subseteq> A ==> P) ==> P" | 
| 34209 | 528 | by simp | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 529 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 530 | lemma equalityCE [elim]: | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 531 | "A = B ==> (c \<in> A ==> c \<in> B ==> P) ==> (c \<notin> A ==> c \<notin> B ==> P) ==> P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 532 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 533 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 534 | lemma eqset_imp_iff: "A = B ==> (x : A) = (x : B)" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 535 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 536 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 537 | lemma eqelem_imp_iff: "x = y ==> (x : A) = (y : A)" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 538 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 539 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 540 | |
| 41082 | 541 | subsubsection {* The empty set *}
 | 
| 542 | ||
| 543 | lemma empty_def: | |
| 544 |   "{} = {x. False}"
 | |
| 545 | by (simp add: bot_fun_def bot_bool_def Collect_def) | |
| 546 | ||
| 547 | lemma empty_iff [simp]: "(c : {}) = False"
 | |
| 548 | by (simp add: empty_def) | |
| 549 | ||
| 550 | lemma emptyE [elim!]: "a : {} ==> P"
 | |
| 551 | by simp | |
| 552 | ||
| 553 | lemma empty_subsetI [iff]: "{} \<subseteq> A"
 | |
| 554 |     -- {* One effect is to delete the ASSUMPTION @{prop "{} <= A"} *}
 | |
| 555 | by blast | |
| 556 | ||
| 557 | lemma equals0I: "(!!y. y \<in> A ==> False) ==> A = {}"
 | |
| 558 | by blast | |
| 559 | ||
| 560 | lemma equals0D: "A = {} ==> a \<notin> A"
 | |
| 561 |     -- {* Use for reasoning about disjointness: @{text "A Int B = {}"} *}
 | |
| 562 | by blast | |
| 563 | ||
| 564 | lemma ball_empty [simp]: "Ball {} P = True"
 | |
| 565 | by (simp add: Ball_def) | |
| 566 | ||
| 567 | lemma bex_empty [simp]: "Bex {} P = False"
 | |
| 568 | by (simp add: Bex_def) | |
| 569 | ||
| 570 | ||
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 571 | subsubsection {* The universal set -- UNIV *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 572 | |
| 32264 
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
 haftmann parents: 
32139diff
changeset | 573 | abbreviation UNIV :: "'a set" where | 
| 
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
 haftmann parents: 
32139diff
changeset | 574 | "UNIV \<equiv> top" | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 575 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 576 | lemma UNIV_def: | 
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 577 |   "UNIV = {x. True}"
 | 
| 41076 
a7fba340058c
primitive definitions of bot/top/inf/sup for bool and fun are named with canonical suffix `_def` rather than `_eq`;
 haftmann parents: 
40872diff
changeset | 578 | by (simp add: top_fun_def top_bool_def Collect_def) | 
| 32081 | 579 | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 580 | lemma UNIV_I [simp]: "x : UNIV" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 581 | by (simp add: UNIV_def) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 582 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 583 | declare UNIV_I [intro]  -- {* unsafe makes it less likely to cause problems *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 584 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 585 | lemma UNIV_witness [intro?]: "EX x. x : UNIV" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 586 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 587 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 588 | lemma subset_UNIV [simp]: "A \<subseteq> UNIV" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 589 | by (rule subsetI) (rule UNIV_I) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 590 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 591 | text {*
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 592 |   \medskip Eta-contracting these two rules (to remove @{text P})
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 593 | causes them to be ignored because of their interaction with | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 594 | congruence rules. | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 595 | *} | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 596 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 597 | lemma ball_UNIV [simp]: "Ball UNIV P = All P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 598 | by (simp add: Ball_def) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 599 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 600 | lemma bex_UNIV [simp]: "Bex UNIV P = Ex P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 601 | by (simp add: Bex_def) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 602 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 603 | lemma UNIV_eq_I: "(\<And>x. x \<in> A) \<Longrightarrow> UNIV = A" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 604 | by auto | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 605 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 606 | lemma UNIV_not_empty [iff]: "UNIV ~= {}"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 607 | by (blast elim: equalityE) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 608 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 609 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 610 | subsubsection {* The Powerset operator -- Pow *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 611 | |
| 32077 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 612 | definition Pow :: "'a set => 'a set set" where | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 613 |   Pow_def: "Pow A = {B. B \<le> A}"
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 614 | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 615 | lemma Pow_iff [iff]: "(A \<in> Pow B) = (A \<subseteq> B)" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 616 | by (simp add: Pow_def) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 617 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 618 | lemma PowI: "A \<subseteq> B ==> A \<in> Pow B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 619 | by (simp add: Pow_def) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 620 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 621 | lemma PowD: "A \<in> Pow B ==> A \<subseteq> B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 622 | by (simp add: Pow_def) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 623 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 624 | lemma Pow_bottom: "{} \<in> Pow B"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 625 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 626 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 627 | lemma Pow_top: "A \<in> Pow A" | 
| 34209 | 628 | by simp | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 629 | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 630 | lemma Pow_not_empty: "Pow A \<noteq> {}"
 | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 631 | using Pow_top by blast | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 632 | |
| 41076 
a7fba340058c
primitive definitions of bot/top/inf/sup for bool and fun are named with canonical suffix `_def` rather than `_eq`;
 haftmann parents: 
40872diff
changeset | 633 | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 634 | subsubsection {* Set complement *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 635 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 636 | lemma Compl_iff [simp]: "(c \<in> -A) = (c \<notin> A)" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 637 | by (simp add: mem_def fun_Compl_def bool_Compl_def) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 638 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 639 | lemma ComplI [intro!]: "(c \<in> A ==> False) ==> c \<in> -A" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 640 | by (unfold mem_def fun_Compl_def bool_Compl_def) blast | 
| 923 | 641 | |
| 11979 | 642 | text {*
 | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 643 | \medskip This form, with negated conclusion, works well with the | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 644 | Classical prover. Negated assumptions behave like formulae on the | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 645 | right side of the notional turnstile ... *} | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 646 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 647 | lemma ComplD [dest!]: "c : -A ==> c~:A" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 648 | by (simp add: mem_def fun_Compl_def bool_Compl_def) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 649 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 650 | lemmas ComplE = ComplD [elim_format] | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 651 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 652 | lemma Compl_eq: "- A = {x. ~ x : A}" by blast
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 653 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 654 | |
| 41082 | 655 | subsubsection {* Binary intersection *}
 | 
| 656 | ||
| 657 | abbreviation inter :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Int" 70) where | |
| 658 | "op Int \<equiv> inf" | |
| 659 | ||
| 660 | notation (xsymbols) | |
| 661 | inter (infixl "\<inter>" 70) | |
| 662 | ||
| 663 | notation (HTML output) | |
| 664 | inter (infixl "\<inter>" 70) | |
| 665 | ||
| 666 | lemma Int_def: | |
| 667 |   "A \<inter> B = {x. x \<in> A \<and> x \<in> B}"
 | |
| 668 | by (simp add: inf_fun_def inf_bool_def Collect_def mem_def) | |
| 669 | ||
| 670 | lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)" | |
| 671 | by (unfold Int_def) blast | |
| 672 | ||
| 673 | lemma IntI [intro!]: "c:A ==> c:B ==> c : A Int B" | |
| 674 | by simp | |
| 675 | ||
| 676 | lemma IntD1: "c : A Int B ==> c:A" | |
| 677 | by simp | |
| 678 | ||
| 679 | lemma IntD2: "c : A Int B ==> c:B" | |
| 680 | by simp | |
| 681 | ||
| 682 | lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P" | |
| 683 | by simp | |
| 684 | ||
| 685 | lemma mono_Int: "mono f \<Longrightarrow> f (A \<inter> B) \<subseteq> f A \<inter> f B" | |
| 686 | by (fact mono_inf) | |
| 687 | ||
| 688 | ||
| 689 | subsubsection {* Binary union *}
 | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 690 | |
| 32683 
7c1fe854ca6a
inter and union are mere abbreviations for inf and sup
 haftmann parents: 
32456diff
changeset | 691 | abbreviation union :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Un" 65) where | 
| 41076 
a7fba340058c
primitive definitions of bot/top/inf/sup for bool and fun are named with canonical suffix `_def` rather than `_eq`;
 haftmann parents: 
40872diff
changeset | 692 | "union \<equiv> sup" | 
| 32081 | 693 | |
| 694 | notation (xsymbols) | |
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 695 | union (infixl "\<union>" 65) | 
| 32081 | 696 | |
| 697 | notation (HTML output) | |
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 698 | union (infixl "\<union>" 65) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 699 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 700 | lemma Un_def: | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 701 |   "A \<union> B = {x. x \<in> A \<or> x \<in> B}"
 | 
| 41076 
a7fba340058c
primitive definitions of bot/top/inf/sup for bool and fun are named with canonical suffix `_def` rather than `_eq`;
 haftmann parents: 
40872diff
changeset | 702 | by (simp add: sup_fun_def sup_bool_def Collect_def mem_def) | 
| 32081 | 703 | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 704 | lemma Un_iff [simp]: "(c : A Un B) = (c:A | c:B)" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 705 | by (unfold Un_def) blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 706 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 707 | lemma UnI1 [elim?]: "c:A ==> c : A Un B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 708 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 709 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 710 | lemma UnI2 [elim?]: "c:B ==> c : A Un B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 711 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 712 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 713 | text {*
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 714 |   \medskip Classical introduction rule: no commitment to @{prop A} vs
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 715 |   @{prop B}.
 | 
| 11979 | 716 | *} | 
| 717 | ||
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 718 | lemma UnCI [intro!]: "(c~:B ==> c:A) ==> c : A Un B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 719 | by auto | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 720 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 721 | lemma UnE [elim!]: "c : A Un B ==> (c:A ==> P) ==> (c:B ==> P) ==> P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 722 | by (unfold Un_def) blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 723 | |
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 724 | lemma insert_def: "insert a B = {x. x = a} \<union> B"
 | 
| 32081 | 725 | by (simp add: Collect_def mem_def insert_compr Un_def) | 
| 726 | ||
| 727 | lemma mono_Un: "mono f \<Longrightarrow> f A \<union> f B \<subseteq> f (A \<union> B)" | |
| 32683 
7c1fe854ca6a
inter and union are mere abbreviations for inf and sup
 haftmann parents: 
32456diff
changeset | 728 | by (fact mono_sup) | 
| 32081 | 729 | |
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 730 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 731 | subsubsection {* Set difference *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 732 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 733 | lemma Diff_iff [simp]: "(c : A - B) = (c:A & c~:B)" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 734 | by (simp add: mem_def fun_diff_def bool_diff_def) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 735 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 736 | lemma DiffI [intro!]: "c : A ==> c ~: B ==> c : A - B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 737 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 738 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 739 | lemma DiffD1: "c : A - B ==> c : A" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 740 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 741 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 742 | lemma DiffD2: "c : A - B ==> c : B ==> P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 743 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 744 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 745 | lemma DiffE [elim!]: "c : A - B ==> (c:A ==> c~:B ==> P) ==> P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 746 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 747 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 748 | lemma set_diff_eq: "A - B = {x. x : A & ~ x : B}" by blast
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 749 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 750 | lemma Compl_eq_Diff_UNIV: "-A = (UNIV - A)" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 751 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 752 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 753 | |
| 31456 | 754 | subsubsection {* Augmenting a set -- @{const insert} *}
 | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 755 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 756 | lemma insert_iff [simp]: "(a : insert b A) = (a = b | a:A)" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 757 | by (unfold insert_def) blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 758 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 759 | lemma insertI1: "a : insert a B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 760 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 761 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 762 | lemma insertI2: "a : B ==> a : insert b B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 763 | by simp | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 764 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 765 | lemma insertE [elim!]: "a : insert b A ==> (a = b ==> P) ==> (a:A ==> P) ==> P" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 766 | by (unfold insert_def) blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 767 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 768 | lemma insertCI [intro!]: "(a~:B ==> a = b) ==> a: insert b B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 769 |   -- {* Classical introduction rule. *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 770 | by auto | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 771 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 772 | lemma subset_insert_iff: "(A \<subseteq> insert x B) = (if x:A then A - {x} \<subseteq> B else A \<subseteq> B)"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 773 | by auto | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 774 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 775 | lemma set_insert: | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 776 | assumes "x \<in> A" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 777 | obtains B where "A = insert x B" and "x \<notin> B" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 778 | proof | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 779 |   from assms show "A = insert x (A - {x})" by blast
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 780 | next | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 781 |   show "x \<notin> A - {x}" by blast
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 782 | qed | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 783 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 784 | lemma insert_ident: "x ~: A ==> x ~: B ==> (insert x A = insert x B) = (A = B)" | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 785 | by auto | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 786 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 787 | subsubsection {* Singletons, using insert *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 788 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 789 | lemma singletonI [intro!,no_atp]: "a : {a}"
 | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 790 |     -- {* Redundant? But unlike @{text insertCI}, it proves the subgoal immediately! *}
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 791 | by (rule insertI1) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 792 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 793 | lemma singletonD [dest!,no_atp]: "b : {a} ==> b = a"
 | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 794 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 795 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 796 | lemmas singletonE = singletonD [elim_format] | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 797 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 798 | lemma singleton_iff: "(b : {a}) = (b = a)"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 799 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 800 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 801 | lemma singleton_inject [dest!]: "{a} = {b} ==> a = b"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 802 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 803 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 804 | lemma singleton_insert_inj_eq [iff,no_atp]: | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 805 |      "({b} = insert a A) = (a = b & A \<subseteq> {b})"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 806 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 807 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 808 | lemma singleton_insert_inj_eq' [iff,no_atp]: | 
| 30531 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 809 |      "(insert a A = {b}) = (a = b & A \<subseteq> {b})"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 810 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 811 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 812 | lemma subset_singletonD: "A \<subseteq> {x} ==> A = {} | A = {x}"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 813 | by fast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 814 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 815 | lemma singleton_conv [simp]: "{x. x = a} = {a}"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 816 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 817 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 818 | lemma singleton_conv2 [simp]: "{x. a = x} = {a}"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 819 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 820 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 821 | lemma diff_single_insert: "A - {x} \<subseteq> B ==> x \<in> A ==> A \<subseteq> insert x B"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 822 | by blast | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 823 | |
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 824 | lemma doubleton_eq_iff: "({a,b} = {c,d}) = (a=c & b=d | a=d & b=c)"
 | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 825 | by (blast elim: equalityE) | 
| 
ab3d61baf66a
reverted to old version of Set.thy -- strange effects have to be traced first
 haftmann parents: 
30352diff
changeset | 826 | |
| 11979 | 827 | |
| 32077 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 828 | subsubsection {* Image of a set under a function *}
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 829 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 830 | text {*
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 831 |   Frequently @{term b} does not have the syntactic form of @{term "f x"}.
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 832 | *} | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 833 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 834 | definition image :: "('a => 'b) => 'a set => 'b set" (infixr "`" 90) where
 | 
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 835 |   image_def [no_atp]: "f ` A = {y. EX x:A. y = f(x)}"
 | 
| 32077 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 836 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 837 | abbreviation | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 838 |   range :: "('a => 'b) => 'b set" where -- "of function"
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 839 | "range f == f ` UNIV" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 840 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 841 | lemma image_eqI [simp, intro]: "b = f x ==> x:A ==> b : f`A" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 842 | by (unfold image_def) blast | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 843 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 844 | lemma imageI: "x : A ==> f x : f ` A" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 845 | by (rule image_eqI) (rule refl) | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 846 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 847 | lemma rev_image_eqI: "x:A ==> b = f x ==> b : f`A" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 848 |   -- {* This version's more effective when we already have the
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 849 |     required @{term x}. *}
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 850 | by (unfold image_def) blast | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 851 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 852 | lemma imageE [elim!]: | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 853 | "b : (%x. f x)`A ==> (!!x. b = f x ==> x:A ==> P) ==> P" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 854 |   -- {* The eta-expansion gives variable-name preservation. *}
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 855 | by (unfold image_def) blast | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 856 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 857 | lemma image_Un: "f`(A Un B) = f`A Un f`B" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 858 | by blast | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 859 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 860 | lemma image_iff: "(z : f`A) = (EX x:A. z = f x)" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 861 | by blast | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 862 | |
| 38648 
52ea97d95e4b
"no_atp" a few facts that often lead to unsound proofs
 blanchet parents: 
37767diff
changeset | 863 | lemma image_subset_iff [no_atp]: "(f`A \<subseteq> B) = (\<forall>x\<in>A. f x \<in> B)" | 
| 32077 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 864 |   -- {* This rewrite rule would confuse users if made default. *}
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 865 | by blast | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 866 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 867 | lemma subset_image_iff: "(B \<subseteq> f`A) = (EX AA. AA \<subseteq> A & B = f`AA)" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 868 | apply safe | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 869 | prefer 2 apply fast | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 870 |   apply (rule_tac x = "{a. a : A & f a : B}" in exI, fast)
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 871 | done | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 872 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 873 | lemma image_subsetI: "(!!x. x \<in> A ==> f x \<in> B) ==> f`A \<subseteq> B" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 874 |   -- {* Replaces the three steps @{text subsetI}, @{text imageE},
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 875 |     @{text hypsubst}, but breaks too many existing proofs. *}
 | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 876 | by blast | 
| 11979 | 877 | |
| 878 | text {*
 | |
| 32077 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 879 | \medskip Range of a function -- just a translation for image! | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 880 | *} | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 881 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 882 | lemma range_eqI: "b = f x ==> b \<in> range f" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 883 | by simp | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 884 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 885 | lemma rangeI: "f x \<in> range f" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 886 | by simp | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 887 | |
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 888 | lemma rangeE [elim?]: "b \<in> range (\<lambda>x. f x) ==> (!!x. b = f x ==> P) ==> P" | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 889 | by blast | 
| 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 890 | |
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 891 | subsubsection {* Some rules with @{text "if"} *}
 | 
| 32081 | 892 | |
| 893 | text{* Elimination of @{text"{x. \<dots> & x=t & \<dots>}"}. *}
 | |
| 894 | ||
| 895 | lemma Collect_conv_if: "{x. x=a & P x} = (if P a then {a} else {})"
 | |
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 896 | by auto | 
| 32081 | 897 | |
| 898 | lemma Collect_conv_if2: "{x. a=x & P x} = (if P a then {a} else {})"
 | |
| 32117 
0762b9ad83df
Set.thy: prefer = over == where possible; tuned ML setup; dropped (moved) ML legacy
 haftmann parents: 
32115diff
changeset | 899 | by auto | 
| 32081 | 900 | |
| 901 | text {*
 | |
| 902 |   Rewrite rules for boolean case-splitting: faster than @{text
 | |
| 903 | "split_if [split]"}. | |
| 904 | *} | |
| 905 | ||
| 906 | lemma split_if_eq1: "((if Q then x else y) = b) = ((Q --> x = b) & (~ Q --> y = b))" | |
| 907 | by (rule split_if) | |
| 908 | ||
| 909 | lemma split_if_eq2: "(a = (if Q then x else y)) = ((Q --> a = x) & (~ Q --> a = y))" | |
| 910 | by (rule split_if) | |
| 911 | ||
| 912 | text {*
 | |
| 913 |   Split ifs on either side of the membership relation.  Not for @{text
 | |
| 914 | "[simp]"} -- can cause goals to blow up! | |
| 915 | *} | |
| 916 | ||
| 917 | lemma split_if_mem1: "((if Q then x else y) : b) = ((Q --> x : b) & (~ Q --> y : b))" | |
| 918 | by (rule split_if) | |
| 919 | ||
| 920 | lemma split_if_mem2: "(a : (if Q then x else y)) = ((Q --> a : x) & (~ Q --> a : y))" | |
| 921 | by (rule split_if [where P="%S. a : S"]) | |
| 922 | ||
| 923 | lemmas split_ifs = if_bool_eq_conj split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2 | |
| 924 | ||
| 925 | (*Would like to add these, but the existing code only searches for the | |
| 37677 | 926 | outer-level constant, which in this case is just Set.member; we instead need | 
| 32081 | 927 | to use term-nets to associate patterns with rules. Also, if a rule fails to | 
| 928 | apply, then the formula should be kept. | |
| 34974 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: 
34209diff
changeset | 929 |   [("uminus", Compl_iff RS iffD1), ("minus", [Diff_iff RS iffD1]),
 | 
| 32081 | 930 |    ("Int", [IntD1,IntD2]),
 | 
| 931 |    ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])]
 | |
| 932 | *) | |
| 933 | ||
| 934 | ||
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 935 | subsection {* Further operations and lemmas *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 936 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 937 | subsubsection {* The ``proper subset'' relation *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 938 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 939 | lemma psubsetI [intro!,no_atp]: "A \<subseteq> B ==> A \<noteq> B ==> A \<subset> B" | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 940 | by (unfold less_le) blast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 941 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 942 | lemma psubsetE [elim!,no_atp]: | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 943 | "[|A \<subset> B; [|A \<subseteq> B; ~ (B\<subseteq>A)|] ==> R|] ==> R" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 944 | by (unfold less_le) blast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 945 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 946 | lemma psubset_insert_iff: | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 947 |   "(A \<subset> insert x B) = (if x \<in> B then A \<subset> B else if x \<in> A then A - {x} \<subset> B else A \<subseteq> B)"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 948 | by (auto simp add: less_le subset_insert_iff) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 949 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 950 | lemma psubset_eq: "(A \<subset> B) = (A \<subseteq> B & A \<noteq> B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 951 | by (simp only: less_le) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 952 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 953 | lemma psubset_imp_subset: "A \<subset> B ==> A \<subseteq> B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 954 | by (simp add: psubset_eq) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 955 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 956 | lemma psubset_trans: "[| A \<subset> B; B \<subset> C |] ==> A \<subset> C" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 957 | apply (unfold less_le) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 958 | apply (auto dest: subset_antisym) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 959 | done | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 960 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 961 | lemma psubsetD: "[| A \<subset> B; c \<in> A |] ==> c \<in> B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 962 | apply (unfold less_le) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 963 | apply (auto dest: subsetD) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 964 | done | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 965 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 966 | lemma psubset_subset_trans: "A \<subset> B ==> B \<subseteq> C ==> A \<subset> C" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 967 | by (auto simp add: psubset_eq) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 968 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 969 | lemma subset_psubset_trans: "A \<subseteq> B ==> B \<subset> C ==> A \<subset> C" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 970 | by (auto simp add: psubset_eq) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 971 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 972 | lemma psubset_imp_ex_mem: "A \<subset> B ==> \<exists>b. b \<in> (B - A)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 973 | by (unfold less_le) blast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 974 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 975 | lemma atomize_ball: | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 976 | "(!!x. x \<in> A ==> P x) == Trueprop (\<forall>x\<in>A. P x)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 977 | by (simp only: Ball_def atomize_all atomize_imp) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 978 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 979 | lemmas [symmetric, rulify] = atomize_ball | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 980 | and [symmetric, defn] = atomize_ball | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 981 | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 982 | lemma image_Pow_mono: | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 983 | assumes "f ` A \<le> B" | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 984 | shows "(image f) ` (Pow A) \<le> Pow B" | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 985 | using assms by blast | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 986 | |
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 987 | lemma image_Pow_surj: | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 988 | assumes "f ` A = B" | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 989 | shows "(image f) ` (Pow A) = Pow B" | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 990 | using assms unfolding Pow_def proof(auto) | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 991 | fix Y assume *: "Y \<le> f ` A" | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 992 |   obtain X where X_def: "X = {x \<in> A. f x \<in> Y}" by blast
 | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 993 | have "f ` X = Y \<and> X \<le> A" unfolding X_def using * by auto | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 994 |   thus "Y \<in> (image f) ` {X. X \<le> A}" by blast
 | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 995 | qed | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
39910diff
changeset | 996 | |
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 997 | subsubsection {* Derived rules involving subsets. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 998 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 999 | text {* @{text insert}. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
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: 
32120diff
changeset | 1001 | lemma subset_insertI: "B \<subseteq> insert a B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1002 | by (rule subsetI) (erule insertI2) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
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: 
32120diff
changeset | 1004 | lemma subset_insertI2: "A \<subseteq> B \<Longrightarrow> A \<subseteq> insert b B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
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: 
32120diff
changeset | 1006 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1007 | lemma subset_insert: "x \<notin> A ==> (A \<subseteq> insert x B) = (A \<subseteq> B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1008 | 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: 
32120diff
changeset | 1009 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1010 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1011 | text {* \medskip Finite Union -- the least upper bound of two sets. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1012 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1013 | lemma Un_upper1: "A \<subseteq> A \<union> B" | 
| 36009 | 1014 | by (fact sup_ge1) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1015 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1016 | lemma Un_upper2: "B \<subseteq> A \<union> B" | 
| 36009 | 1017 | by (fact sup_ge2) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1018 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1019 | lemma Un_least: "A \<subseteq> C ==> B \<subseteq> C ==> A \<union> B \<subseteq> C" | 
| 36009 | 1020 | 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: 
32120diff
changeset | 1021 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1022 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1023 | text {* \medskip Finite Intersection -- the greatest lower bound of two sets. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1024 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1025 | lemma Int_lower1: "A \<inter> B \<subseteq> A" | 
| 36009 | 1026 | by (fact inf_le1) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1027 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1028 | lemma Int_lower2: "A \<inter> B \<subseteq> B" | 
| 36009 | 1029 | by (fact inf_le2) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1030 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1031 | lemma Int_greatest: "C \<subseteq> A ==> C \<subseteq> B ==> C \<subseteq> A \<inter> B" | 
| 36009 | 1032 | by (fact inf_greatest) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1033 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1034 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1035 | text {* \medskip Set difference. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1036 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1037 | lemma Diff_subset: "A - B \<subseteq> A" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1038 | 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: 
32120diff
changeset | 1039 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1040 | lemma Diff_subset_conv: "(A - B \<subseteq> C) = (A \<subseteq> B \<union> C)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1041 | 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: 
32120diff
changeset | 1042 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1043 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1044 | subsubsection {* Equalities involving union, intersection, inclusion, etc. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1045 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1046 | text {* @{text "{}"}. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1047 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1048 | lemma Collect_const [simp]: "{s. P} = (if P then UNIV else {})"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1049 |   -- {* supersedes @{text "Collect_False_empty"} *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1050 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1051 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1052 | lemma subset_empty [simp]: "(A \<subseteq> {}) = (A = {})"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1053 | 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: 
32120diff
changeset | 1054 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1055 | lemma not_psubset_empty [iff]: "\<not> (A < {})"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1056 | by (unfold less_le) blast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1057 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1058 | lemma Collect_empty_eq [simp]: "(Collect P = {}) = (\<forall>x. \<not> P x)"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1059 | 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: 
32120diff
changeset | 1060 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1061 | lemma empty_Collect_eq [simp]: "({} = Collect P) = (\<forall>x. \<not> P x)"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1062 | 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: 
32120diff
changeset | 1063 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1064 | lemma Collect_neg_eq: "{x. \<not> P x} = - {x. P x}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1065 | 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: 
32120diff
changeset | 1066 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1067 | lemma Collect_disj_eq: "{x. P x | Q x} = {x. P x} \<union> {x. Q x}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1068 | 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: 
32120diff
changeset | 1069 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1070 | lemma Collect_imp_eq: "{x. P x --> Q x} = -{x. P x} \<union> {x. Q x}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1071 | 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: 
32120diff
changeset | 1072 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1073 | lemma Collect_conj_eq: "{x. P x & Q x} = {x. P x} \<inter> {x. Q x}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1074 | 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: 
32120diff
changeset | 1075 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1076 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1077 | text {* \medskip @{text insert}. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1078 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1079 | lemma insert_is_Un: "insert a A = {a} Un A"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1080 |   -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a {}"} *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1081 | 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: 
32120diff
changeset | 1082 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1083 | lemma insert_not_empty [simp]: "insert a A \<noteq> {}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1084 | 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: 
32120diff
changeset | 1085 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1086 | lemmas empty_not_insert = insert_not_empty [symmetric, standard] | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1087 | declare empty_not_insert [simp] | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1088 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1089 | lemma insert_absorb: "a \<in> A ==> insert a 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: 
32120diff
changeset | 1090 |   -- {* @{text "[simp]"} causes recursive calls when there are nested inserts *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1091 |   -- {* with \emph{quadratic} running time *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1092 | 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: 
32120diff
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: 
32120diff
changeset | 1094 | lemma insert_absorb2 [simp]: "insert x (insert x A) = insert x A" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1095 | 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: 
32120diff
changeset | 1096 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1097 | lemma insert_commute: "insert x (insert y A) = insert y (insert x A)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1098 | 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: 
32120diff
changeset | 1099 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1100 | lemma insert_subset [simp]: "(insert x A \<subseteq> B) = (x \<in> B & A \<subseteq> B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
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: 
32120diff
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: 
32120diff
changeset | 1103 | lemma mk_disjoint_insert: "a \<in> A ==> \<exists>B. A = insert a B & a \<notin> B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1104 |   -- {* use new @{text B} rather than @{text "A - {a}"} to avoid infinite unfolding *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1105 |   apply (rule_tac x = "A - {a}" in exI, blast)
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1106 | done | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1107 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1108 | lemma insert_Collect: "insert a (Collect P) = {u. u \<noteq> a --> P u}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1109 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1110 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1111 | lemma insert_inter_insert[simp]: "insert a A \<inter> insert a B = insert a (A \<inter> B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1112 | 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: 
32120diff
changeset | 1113 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 1114 | lemma insert_disjoint [simp,no_atp]: | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1115 |  "(insert a A \<inter> B = {}) = (a \<notin> B \<and> A \<inter> B = {})"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1116 |  "({} = insert a A \<inter> B) = (a \<notin> B \<and> {} = A \<inter> B)"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1117 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1118 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 1119 | lemma disjoint_insert [simp,no_atp]: | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1120 |  "(B \<inter> insert a A = {}) = (a \<notin> B \<and> B \<inter> A = {})"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1121 |  "({} = A \<inter> insert b B) = (b \<notin> A \<and> {} = A \<inter> B)"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1122 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1123 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1124 | text {* \medskip @{text image}. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1125 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1126 | lemma image_empty [simp]: "f`{} = {}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
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: 
32120diff
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: 
32120diff
changeset | 1129 | lemma image_insert [simp]: "f ` insert a B = insert (f a) (f`B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1130 | 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: 
32120diff
changeset | 1131 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1132 | lemma image_constant: "x \<in> A ==> (\<lambda>x. c) ` A = {c}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1133 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
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: 
32120diff
changeset | 1135 | lemma image_constant_conv: "(%x. c) ` A = (if A = {} then {} else {c})"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1136 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
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: 
32120diff
changeset | 1138 | lemma image_image: "f ` (g ` A) = (\<lambda>x. f (g x)) ` A" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1139 | 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: 
32120diff
changeset | 1140 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1141 | lemma insert_image [simp]: "x \<in> A ==> insert (f x) (f`A) = f`A" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1142 | 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: 
32120diff
changeset | 1143 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1144 | lemma image_is_empty [iff]: "(f`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: 
32120diff
changeset | 1145 | 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: 
32120diff
changeset | 1146 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1147 | lemma empty_is_image[iff]: "({} = f ` 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: 
32120diff
changeset | 1148 | 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: 
32120diff
changeset | 1149 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1150 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 1151 | lemma image_Collect [no_atp]: "f ` {x. P x} = {f x | x. P 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: 
32120diff
changeset | 1152 |   -- {* NOT suitable as a default simprule: the RHS isn't simpler than the LHS,
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1153 | with its implicit quantifier and conjunction. Also image enjoys better | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1154 | equational properties than does the RHS. *} | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1155 | 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: 
32120diff
changeset | 1156 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1157 | lemma if_image_distrib [simp]: | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1158 | "(\<lambda>x. if P x then f x else g x) ` S | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1159 |     = (f ` (S \<inter> {x. P x})) \<union> (g ` (S \<inter> {x. \<not> P x}))"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1160 | by (auto simp add: image_def) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1161 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1162 | lemma image_cong: "M = N ==> (!!x. x \<in> N ==> f x = g x) ==> f`M = g`N" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1163 | by (simp add: image_def) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1164 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1165 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1166 | text {* \medskip @{text range}. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1167 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 1168 | lemma full_SetCompr_eq [no_atp]: "{u. \<exists>x. u = f x} = range f"
 | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1169 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1170 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1171 | lemma range_composition: "range (\<lambda>x. f (g x)) = f`range g" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1172 | by (subst image_image, simp) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1173 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1174 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1175 | text {* \medskip @{text Int} *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1176 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1177 | lemma Int_absorb [simp]: "A \<inter> A = A" | 
| 36009 | 1178 | by (fact inf_idem) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1179 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1180 | lemma Int_left_absorb: "A \<inter> (A \<inter> B) = A \<inter> B" | 
| 36009 | 1181 | by (fact inf_left_idem) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1182 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1183 | lemma Int_commute: "A \<inter> B = B \<inter> A" | 
| 36009 | 1184 | by (fact inf_commute) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1185 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1186 | lemma Int_left_commute: "A \<inter> (B \<inter> C) = B \<inter> (A \<inter> C)" | 
| 36009 | 1187 | by (fact inf_left_commute) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1188 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1189 | lemma Int_assoc: "(A \<inter> B) \<inter> C = A \<inter> (B \<inter> C)" | 
| 36009 | 1190 | by (fact inf_assoc) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1191 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1192 | lemmas Int_ac = Int_assoc Int_left_absorb Int_commute Int_left_commute | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1193 |   -- {* Intersection is an AC-operator *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1194 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1195 | lemma Int_absorb1: "B \<subseteq> A ==> A \<inter> B = B" | 
| 36009 | 1196 | by (fact inf_absorb2) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1197 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1198 | lemma Int_absorb2: "A \<subseteq> B ==> A \<inter> B = A" | 
| 36009 | 1199 | by (fact inf_absorb1) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1200 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1201 | lemma Int_empty_left [simp]: "{} \<inter> B = {}"
 | 
| 36009 | 1202 | by (fact inf_bot_left) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1203 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1204 | lemma Int_empty_right [simp]: "A \<inter> {} = {}"
 | 
| 36009 | 1205 | by (fact inf_bot_right) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1206 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1207 | lemma disjoint_eq_subset_Compl: "(A \<inter> B = {}) = (A \<subseteq> -B)"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1208 | 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: 
32120diff
changeset | 1209 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1210 | lemma disjoint_iff_not_equal: "(A \<inter> B = {}) = (\<forall>x\<in>A. \<forall>y\<in>B. x \<noteq> y)"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1211 | 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: 
32120diff
changeset | 1212 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1213 | lemma Int_UNIV_left [simp]: "UNIV \<inter> B = B" | 
| 36009 | 1214 | by (fact inf_top_left) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1215 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1216 | lemma Int_UNIV_right [simp]: "A \<inter> UNIV = A" | 
| 36009 | 1217 | by (fact inf_top_right) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1218 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1219 | lemma Int_Un_distrib: "A \<inter> (B \<union> C) = (A \<inter> B) \<union> (A \<inter> C)" | 
| 36009 | 1220 | by (fact inf_sup_distrib1) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1221 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1222 | lemma Int_Un_distrib2: "(B \<union> C) \<inter> A = (B \<inter> A) \<union> (C \<inter> A)" | 
| 36009 | 1223 | by (fact inf_sup_distrib2) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1224 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 1225 | lemma Int_UNIV [simp,no_atp]: "(A \<inter> B = UNIV) = (A = UNIV & B = UNIV)" | 
| 36009 | 1226 | by (fact inf_eq_top_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: 
32120diff
changeset | 1227 | |
| 38648 
52ea97d95e4b
"no_atp" a few facts that often lead to unsound proofs
 blanchet parents: 
37767diff
changeset | 1228 | lemma Int_subset_iff [no_atp, simp]: "(C \<subseteq> A \<inter> B) = (C \<subseteq> A & C \<subseteq> B)" | 
| 36009 | 1229 | by (fact le_inf_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: 
32120diff
changeset | 1230 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1231 | lemma Int_Collect: "(x \<in> A \<inter> {x. P x}) = (x \<in> A & P x)"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1232 | 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: 
32120diff
changeset | 1233 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1234 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1235 | text {* \medskip @{text Un}. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1236 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1237 | lemma Un_absorb [simp]: "A \<union> A = A" | 
| 36009 | 1238 | by (fact sup_idem) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1239 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1240 | lemma Un_left_absorb: "A \<union> (A \<union> B) = A \<union> B" | 
| 36009 | 1241 | by (fact sup_left_idem) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1242 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1243 | lemma Un_commute: "A \<union> B = B \<union> A" | 
| 36009 | 1244 | by (fact sup_commute) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1245 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1246 | lemma Un_left_commute: "A \<union> (B \<union> C) = B \<union> (A \<union> C)" | 
| 36009 | 1247 | by (fact sup_left_commute) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1248 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1249 | lemma Un_assoc: "(A \<union> B) \<union> C = A \<union> (B \<union> C)" | 
| 36009 | 1250 | by (fact sup_assoc) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1251 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1252 | lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1253 |   -- {* Union is an AC-operator *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1254 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1255 | lemma Un_absorb1: "A \<subseteq> B ==> A \<union> B = B" | 
| 36009 | 1256 | by (fact sup_absorb2) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1257 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1258 | lemma Un_absorb2: "B \<subseteq> A ==> A \<union> B = A" | 
| 36009 | 1259 | by (fact sup_absorb1) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1260 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1261 | lemma Un_empty_left [simp]: "{} \<union> B = B"
 | 
| 36009 | 1262 | by (fact sup_bot_left) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1263 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1264 | lemma Un_empty_right [simp]: "A \<union> {} = A"
 | 
| 36009 | 1265 | by (fact sup_bot_right) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1266 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1267 | lemma Un_UNIV_left [simp]: "UNIV \<union> B = UNIV" | 
| 36009 | 1268 | by (fact sup_top_left) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1269 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1270 | lemma Un_UNIV_right [simp]: "A \<union> UNIV = UNIV" | 
| 36009 | 1271 | by (fact sup_top_right) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1272 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1273 | lemma Un_insert_left [simp]: "(insert a B) \<union> C = insert a (B \<union> C)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1274 | 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: 
32120diff
changeset | 1275 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1276 | lemma Un_insert_right [simp]: "A \<union> (insert a B) = insert a (A \<union> B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1277 | 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: 
32120diff
changeset | 1278 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1279 | lemma Int_insert_left: | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1280 | "(insert a B) Int C = (if a \<in> C then insert a (B \<inter> C) else B \<inter> C)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1281 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1282 | |
| 32456 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1283 | lemma Int_insert_left_if0[simp]: | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1284 | "a \<notin> C \<Longrightarrow> (insert a B) Int C = B \<inter> C" | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1285 | by auto | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1286 | |
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1287 | lemma Int_insert_left_if1[simp]: | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1288 | "a \<in> C \<Longrightarrow> (insert a B) Int C = insert a (B Int C)" | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1289 | by auto | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1290 | |
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1291 | lemma Int_insert_right: | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1292 | "A \<inter> (insert a B) = (if a \<in> A then insert a (A \<inter> B) else A \<inter> B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1293 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1294 | |
| 32456 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1295 | lemma Int_insert_right_if0[simp]: | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1296 | "a \<notin> A \<Longrightarrow> A Int (insert a B) = A Int B" | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1297 | by auto | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1298 | |
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1299 | lemma Int_insert_right_if1[simp]: | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1300 | "a \<in> A \<Longrightarrow> A Int (insert a B) = insert a (A Int B)" | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1301 | by auto | 
| 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 nipkow parents: 
32264diff
changeset | 1302 | |
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1303 | lemma Un_Int_distrib: "A \<union> (B \<inter> C) = (A \<union> B) \<inter> (A \<union> C)" | 
| 36009 | 1304 | by (fact sup_inf_distrib1) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1305 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1306 | lemma Un_Int_distrib2: "(B \<inter> C) \<union> A = (B \<union> A) \<inter> (C \<union> A)" | 
| 36009 | 1307 | by (fact sup_inf_distrib2) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1308 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1309 | lemma Un_Int_crazy: | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1310 | "(A \<inter> B) \<union> (B \<inter> C) \<union> (C \<inter> A) = (A \<union> B) \<inter> (B \<union> C) \<inter> (C \<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: 
32120diff
changeset | 1311 | 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: 
32120diff
changeset | 1312 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1313 | lemma subset_Un_eq: "(A \<subseteq> B) = (A \<union> B = B)" | 
| 36009 | 1314 | by (fact le_iff_sup) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1315 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1316 | lemma Un_empty [iff]: "(A \<union> B = {}) = (A = {} & B = {})"
 | 
| 36009 | 1317 | by (fact sup_eq_bot_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: 
32120diff
changeset | 1318 | |
| 38648 
52ea97d95e4b
"no_atp" a few facts that often lead to unsound proofs
 blanchet parents: 
37767diff
changeset | 1319 | lemma Un_subset_iff [no_atp, simp]: "(A \<union> B \<subseteq> C) = (A \<subseteq> C & B \<subseteq> C)" | 
| 36009 | 1320 | by (fact le_sup_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: 
32120diff
changeset | 1321 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1322 | lemma Un_Diff_Int: "(A - B) \<union> (A \<inter> B) = A" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1323 | 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: 
32120diff
changeset | 1324 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1325 | lemma Diff_Int2: "A \<inter> C - B \<inter> C = A \<inter> C - B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1326 | 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: 
32120diff
changeset | 1327 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1328 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1329 | text {* \medskip Set complement *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1330 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1331 | lemma Compl_disjoint [simp]: "A \<inter> -A = {}"
 | 
| 36009 | 1332 | by (fact inf_compl_bot) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1333 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1334 | lemma Compl_disjoint2 [simp]: "-A \<inter> A = {}"
 | 
| 36009 | 1335 | by (fact compl_inf_bot) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1336 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1337 | lemma Compl_partition: "A \<union> -A = UNIV" | 
| 36009 | 1338 | by (fact sup_compl_top) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1339 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1340 | lemma Compl_partition2: "-A \<union> A = UNIV" | 
| 36009 | 1341 | by (fact compl_sup_top) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1342 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1343 | lemma double_complement [simp]: "- (-A) = (A::'a set)" | 
| 36009 | 1344 | by (fact double_compl) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1345 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1346 | lemma Compl_Un [simp]: "-(A \<union> B) = (-A) \<inter> (-B)" | 
| 36009 | 1347 | by (fact compl_sup) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1348 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1349 | lemma Compl_Int [simp]: "-(A \<inter> B) = (-A) \<union> (-B)" | 
| 36009 | 1350 | by (fact compl_inf) | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1351 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1352 | lemma subset_Compl_self_eq: "(A \<subseteq> -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: 
32120diff
changeset | 1353 | 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: 
32120diff
changeset | 1354 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1355 | lemma Un_Int_assoc_eq: "((A \<inter> B) \<union> C = A \<inter> (B \<union> C)) = (C \<subseteq> A)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1356 |   -- {* Halmos, Naive Set Theory, page 16. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1357 | 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: 
32120diff
changeset | 1358 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1359 | lemma Compl_UNIV_eq [simp]: "-UNIV = {}"
 | 
| 36009 | 1360 | by (fact compl_top_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: 
32120diff
changeset | 1361 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1362 | lemma Compl_empty_eq [simp]: "-{} = UNIV"
 | 
| 36009 | 1363 | by (fact compl_bot_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: 
32120diff
changeset | 1364 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1365 | lemma Compl_subset_Compl_iff [iff]: "(-A \<subseteq> -B) = (B \<subseteq> A)" | 
| 36009 | 1366 | by (fact compl_le_compl_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: 
32120diff
changeset | 1367 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1368 | lemma Compl_eq_Compl_iff [iff]: "(-A = -B) = (A = (B::'a set))" | 
| 36009 | 1369 | by (fact compl_eq_compl_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: 
32120diff
changeset | 1370 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1371 | text {* \medskip Bounded quantifiers.
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1372 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1373 | The following are not added to the default simpset because | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1374 |   (a) they duplicate the body and (b) there are no similar rules for @{text Int}. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1375 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1376 | lemma ball_Un: "(\<forall>x \<in> A \<union> B. P x) = ((\<forall>x\<in>A. P x) & (\<forall>x\<in>B. P x))" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1377 | 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: 
32120diff
changeset | 1378 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1379 | lemma bex_Un: "(\<exists>x \<in> A \<union> B. P x) = ((\<exists>x\<in>A. P x) | (\<exists>x\<in>B. P x))" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1380 | 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: 
32120diff
changeset | 1381 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1382 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1383 | text {* \medskip Set difference. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1384 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1385 | lemma Diff_eq: "A - B = A \<inter> (-B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1386 | 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: 
32120diff
changeset | 1387 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 1388 | lemma Diff_eq_empty_iff [simp,no_atp]: "(A - B = {}) = (A \<subseteq> B)"
 | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1389 | 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: 
32120diff
changeset | 1390 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1391 | lemma Diff_cancel [simp]: "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: 
32120diff
changeset | 1392 | 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: 
32120diff
changeset | 1393 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1394 | lemma Diff_idemp [simp]: "(A - B) - B = A - (B::'a set)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1395 | 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: 
32120diff
changeset | 1396 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1397 | lemma Diff_triv: "A \<inter> B = {} ==> A - B = A"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1398 | by (blast elim: equalityE) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1399 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1400 | lemma empty_Diff [simp]: "{} - A = {}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1401 | 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: 
32120diff
changeset | 1402 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1403 | lemma Diff_empty [simp]: "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: 
32120diff
changeset | 1404 | 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: 
32120diff
changeset | 1405 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1406 | lemma Diff_UNIV [simp]: "A - UNIV = {}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1407 | 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: 
32120diff
changeset | 1408 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 1409 | lemma Diff_insert0 [simp,no_atp]: "x \<notin> A ==> A - insert x B = A - B" | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1410 | 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: 
32120diff
changeset | 1411 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1412 | lemma Diff_insert: "A - insert a B = A - B - {a}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1413 |   -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"} *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1414 | 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: 
32120diff
changeset | 1415 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1416 | lemma Diff_insert2: "A - insert a B = A - {a} - B"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1417 |   -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"} *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1418 | 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: 
32120diff
changeset | 1419 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1420 | lemma insert_Diff_if: "insert x A - B = (if x \<in> B then A - B else insert x (A - B))" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1421 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1422 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1423 | lemma insert_Diff1 [simp]: "x \<in> B ==> insert x A - B = A - B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1424 | 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: 
32120diff
changeset | 1425 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1426 | lemma insert_Diff_single[simp]: "insert a (A - {a}) = insert 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: 
32120diff
changeset | 1427 | 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: 
32120diff
changeset | 1428 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1429 | lemma insert_Diff: "a \<in> A ==> insert a (A - {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: 
32120diff
changeset | 1430 | 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: 
32120diff
changeset | 1431 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1432 | lemma Diff_insert_absorb: "x \<notin> A ==> (insert x A) - {x} = A"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1433 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1434 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1435 | lemma Diff_disjoint [simp]: "A \<inter> (B - A) = {}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1436 | 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: 
32120diff
changeset | 1437 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1438 | lemma Diff_partition: "A \<subseteq> B ==> A \<union> (B - A) = B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1439 | 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: 
32120diff
changeset | 1440 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1441 | lemma double_diff: "A \<subseteq> B ==> B \<subseteq> C ==> B - (C - 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: 
32120diff
changeset | 1442 | 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: 
32120diff
changeset | 1443 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1444 | lemma Un_Diff_cancel [simp]: "A \<union> (B - A) = A \<union> B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1445 | 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: 
32120diff
changeset | 1446 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1447 | lemma Un_Diff_cancel2 [simp]: "(B - A) \<union> A = B \<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: 
32120diff
changeset | 1448 | 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: 
32120diff
changeset | 1449 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1450 | lemma Diff_Un: "A - (B \<union> C) = (A - B) \<inter> (A - C)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1451 | 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: 
32120diff
changeset | 1452 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1453 | lemma Diff_Int: "A - (B \<inter> C) = (A - B) \<union> (A - C)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1454 | 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: 
32120diff
changeset | 1455 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1456 | lemma Un_Diff: "(A \<union> B) - C = (A - C) \<union> (B - C)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1457 | 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: 
32120diff
changeset | 1458 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1459 | lemma Int_Diff: "(A \<inter> B) - C = A \<inter> (B - C)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1460 | 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: 
32120diff
changeset | 1461 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1462 | lemma Diff_Int_distrib: "C \<inter> (A - B) = (C \<inter> A) - (C \<inter> B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1463 | 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: 
32120diff
changeset | 1464 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1465 | lemma Diff_Int_distrib2: "(A - B) \<inter> C = (A \<inter> C) - (B \<inter> C)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1466 | 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: 
32120diff
changeset | 1467 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1468 | lemma Diff_Compl [simp]: "A - (- B) = A \<inter> B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1469 | by auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1470 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1471 | lemma Compl_Diff_eq [simp]: "- (A - B) = -A \<union> B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1472 | 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: 
32120diff
changeset | 1473 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1474 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1475 | text {* \medskip Quantification over type @{typ bool}. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1476 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1477 | lemma bool_induct: "P True \<Longrightarrow> P False \<Longrightarrow> P x" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1478 | by (cases x) auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1479 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1480 | lemma all_bool_eq: "(\<forall>b. P b) \<longleftrightarrow> P True \<and> P False" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1481 | by (auto intro: bool_induct) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1482 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1483 | lemma bool_contrapos: "P x \<Longrightarrow> \<not> P False \<Longrightarrow> P True" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1484 | by (cases x) auto | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1485 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1486 | lemma ex_bool_eq: "(\<exists>b. P b) \<longleftrightarrow> P True \<or> P False" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1487 | by (auto intro: bool_contrapos) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1488 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1489 | text {* \medskip @{text Pow} *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1490 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1491 | lemma Pow_empty [simp]: "Pow {} = {{}}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1492 | by (auto simp add: Pow_def) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1493 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1494 | lemma Pow_insert: "Pow (insert a A) = Pow A \<union> (insert a ` Pow A)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1495 |   by (blast intro: image_eqI [where ?x = "u - {a}", standard])
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1496 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1497 | lemma Pow_Compl: "Pow (- A) = {-B | B. A \<in> Pow B}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1498 | by (blast intro: exI [where ?x = "- u", standard]) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1499 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1500 | lemma Pow_UNIV [simp]: "Pow UNIV = UNIV" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1501 | 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: 
32120diff
changeset | 1502 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1503 | lemma Un_Pow_subset: "Pow A \<union> Pow B \<subseteq> Pow (A \<union> B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1504 | 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: 
32120diff
changeset | 1505 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1506 | lemma Pow_Int_eq [simp]: "Pow (A \<inter> B) = Pow A \<inter> Pow B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1507 | 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: 
32120diff
changeset | 1508 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1509 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1510 | text {* \medskip Miscellany. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1511 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1512 | lemma set_eq_subset: "(A = B) = (A \<subseteq> B & B \<subseteq> A)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1513 | 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: 
32120diff
changeset | 1514 | |
| 38648 
52ea97d95e4b
"no_atp" a few facts that often lead to unsound proofs
 blanchet parents: 
37767diff
changeset | 1515 | lemma subset_iff [no_atp]: "(A \<subseteq> B) = (\<forall>t. t \<in> A --> t \<in> B)" | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1516 | 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: 
32120diff
changeset | 1517 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1518 | lemma subset_iff_psubset_eq: "(A \<subseteq> B) = ((A \<subset> B) | (A = B))" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1519 | by (unfold less_le) blast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1520 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1521 | lemma all_not_in_conv [simp]: "(\<forall>x. x \<notin> 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: 
32120diff
changeset | 1522 | 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: 
32120diff
changeset | 1523 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1524 | lemma ex_in_conv: "(\<exists>x. x \<in> A) = (A \<noteq> {})"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1525 | 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: 
32120diff
changeset | 1526 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1527 | lemma distinct_lemma: "f x \<noteq> f y ==> x \<noteq> y" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1528 | by iprover | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1529 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1530 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1531 | subsubsection {* Monotonicity of various operations *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1532 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1533 | lemma image_mono: "A \<subseteq> B ==> f`A \<subseteq> f`B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1534 | 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: 
32120diff
changeset | 1535 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1536 | lemma Pow_mono: "A \<subseteq> B ==> Pow A \<subseteq> Pow B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1537 | 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: 
32120diff
changeset | 1538 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1539 | lemma insert_mono: "C \<subseteq> D ==> insert a C \<subseteq> insert a D" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1540 | 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: 
32120diff
changeset | 1541 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1542 | lemma Un_mono: "A \<subseteq> C ==> B \<subseteq> D ==> A \<union> B \<subseteq> C \<union> D" | 
| 36009 | 1543 | by (fact sup_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: 
32120diff
changeset | 1544 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1545 | lemma Int_mono: "A \<subseteq> C ==> B \<subseteq> D ==> A \<inter> B \<subseteq> C \<inter> D" | 
| 36009 | 1546 | by (fact inf_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: 
32120diff
changeset | 1547 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1548 | lemma Diff_mono: "A \<subseteq> C ==> D \<subseteq> B ==> A - B \<subseteq> C - D" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1549 | 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: 
32120diff
changeset | 1550 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1551 | lemma Compl_anti_mono: "A \<subseteq> B ==> -B \<subseteq> -A" | 
| 36009 | 1552 | by (fact compl_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: 
32120diff
changeset | 1553 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1554 | text {* \medskip Monotonicity of implications. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1555 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1556 | lemma in_mono: "A \<subseteq> B ==> x \<in> A --> x \<in> B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1557 | apply (rule impI) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1558 | apply (erule subsetD, assumption) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1559 | done | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1560 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1561 | lemma conj_mono: "P1 --> Q1 ==> P2 --> Q2 ==> (P1 & P2) --> (Q1 & Q2)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1562 | by iprover | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1563 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1564 | lemma disj_mono: "P1 --> Q1 ==> P2 --> Q2 ==> (P1 | P2) --> (Q1 | Q2)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1565 | by iprover | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1566 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1567 | lemma imp_mono: "Q1 --> P1 ==> P2 --> Q2 ==> (P1 --> P2) --> (Q1 --> Q2)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1568 | by iprover | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1569 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1570 | lemma imp_refl: "P --> P" .. | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1571 | |
| 33935 | 1572 | lemma not_mono: "Q --> P ==> ~ P --> ~ Q" | 
| 1573 | by iprover | |
| 1574 | ||
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1575 | lemma ex_mono: "(!!x. P x --> Q x) ==> (EX x. P x) --> (EX x. Q x)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1576 | by iprover | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1577 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1578 | lemma all_mono: "(!!x. P x --> Q x) ==> (ALL x. P x) --> (ALL x. Q x)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1579 | by iprover | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1580 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1581 | lemma Collect_mono: "(!!x. P x --> Q x) ==> Collect P \<subseteq> Collect Q" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1582 | 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: 
32120diff
changeset | 1583 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1584 | lemma Int_Collect_mono: | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1585 | "A \<subseteq> B ==> (!!x. x \<in> A ==> P x --> Q x) ==> A \<inter> Collect P \<subseteq> B \<inter> Collect Q" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1586 | 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: 
32120diff
changeset | 1587 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1588 | lemmas basic_monos = | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1589 | subset_refl imp_refl disj_mono conj_mono | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1590 | ex_mono Collect_mono in_mono | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1591 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1592 | lemma eq_to_mono: "a = b ==> c = d ==> b --> d ==> a --> c" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1593 | by iprover | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1594 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1595 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1596 | subsubsection {* Inverse image of a function *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1597 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35115diff
changeset | 1598 | definition vimage :: "('a => 'b) => 'b set => 'a set" (infixr "-`" 90) where
 | 
| 37767 | 1599 |   "f -` B == {x. f x : B}"
 | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1600 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1601 | lemma vimage_eq [simp]: "(a : f -` B) = (f a : B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1602 | by (unfold vimage_def) blast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1603 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1604 | lemma vimage_singleton_eq: "(a : f -` {b}) = (f a = b)"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1605 | by simp | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1606 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1607 | lemma vimageI [intro]: "f a = b ==> b:B ==> a : f -` B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1608 | by (unfold vimage_def) blast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1609 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1610 | lemma vimageI2: "f a : A ==> a : f -` A" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1611 | by (unfold vimage_def) fast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1612 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1613 | lemma vimageE [elim!]: "a: f -` B ==> (!!x. f a = x ==> x:B ==> P) ==> P" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1614 | by (unfold vimage_def) blast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1615 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1616 | lemma vimageD: "a : f -` A ==> f 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: 
32120diff
changeset | 1617 | by (unfold vimage_def) fast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1618 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1619 | lemma vimage_empty [simp]: "f -` {} = {}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1620 | 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: 
32120diff
changeset | 1621 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1622 | lemma vimage_Compl: "f -` (-A) = -(f -` A)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1623 | 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: 
32120diff
changeset | 1624 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1625 | lemma vimage_Un [simp]: "f -` (A Un B) = (f -` A) Un (f -` B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1626 | 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: 
32120diff
changeset | 1627 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1628 | lemma vimage_Int [simp]: "f -` (A Int B) = (f -` A) Int (f -` B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1629 | by fast | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1630 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1631 | lemma vimage_Collect_eq [simp]: "f -` Collect P = {y. P (f y)}"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1632 | 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: 
32120diff
changeset | 1633 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1634 | lemma vimage_Collect: "(!!x. P (f x) = Q x) ==> f -` (Collect P) = Collect Q" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1635 | 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: 
32120diff
changeset | 1636 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1637 | lemma vimage_insert: "f-`(insert a B) = (f-`{a}) Un (f-`B)"
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1638 |   -- {* NOT suitable for rewriting because of the recurrence of @{term "{a}"}. *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1639 | 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: 
32120diff
changeset | 1640 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1641 | lemma vimage_Diff: "f -` (A - B) = (f -` A) - (f -` B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1642 | 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: 
32120diff
changeset | 1643 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1644 | lemma vimage_UNIV [simp]: "f -` UNIV = UNIV" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1645 | 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: 
32120diff
changeset | 1646 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1647 | lemma vimage_mono: "A \<subseteq> B ==> f -` A \<subseteq> f -` B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1648 |   -- {* monotonicity *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1649 | 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: 
32120diff
changeset | 1650 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35576diff
changeset | 1651 | lemma vimage_image_eq [no_atp]: "f -` (f ` A) = {y. EX x:A. f x = f 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: 
32120diff
changeset | 1652 | by (blast intro: sym) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1653 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1654 | lemma image_vimage_subset: "f ` (f -` 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: 
32120diff
changeset | 1655 | 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: 
32120diff
changeset | 1656 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1657 | lemma image_vimage_eq [simp]: "f ` (f -` A) = A Int range f" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1658 | 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: 
32120diff
changeset | 1659 | |
| 33533 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: 
33045diff
changeset | 1660 | lemma vimage_const [simp]: "((\<lambda>x. c) -` A) = (if c \<in> A then UNIV else {})"
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: 
33045diff
changeset | 1661 | by auto | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: 
33045diff
changeset | 1662 | |
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: 
33045diff
changeset | 1663 | lemma vimage_if [simp]: "((\<lambda>x. if x \<in> B then c else d) -` A) = | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: 
33045diff
changeset | 1664 | (if c \<in> A then (if d \<in> A then UNIV else B) | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: 
33045diff
changeset | 1665 |     else if d \<in> A then -B else {})"  
 | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: 
33045diff
changeset | 1666 | by (auto simp add: vimage_def) | 
| 
40b44cb20c8c
New theory Probability/Borel.thy, and some associated lemmas
 paulson parents: 
33045diff
changeset | 1667 | |
| 35576 | 1668 | lemma vimage_inter_cong: | 
| 1669 | "(\<And> w. w \<in> S \<Longrightarrow> f w = g w) \<Longrightarrow> f -` y \<inter> S = g -` y \<inter> S" | |
| 1670 | by auto | |
| 1671 | ||
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1672 | lemma image_Int_subset: "f`(A Int B) <= f`A Int f`B" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1673 | 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: 
32120diff
changeset | 1674 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1675 | lemma image_diff_subset: "f`A - f`B <= f`(A - B)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1676 | 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: 
32120diff
changeset | 1677 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1678 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1679 | subsubsection {* Getting the Contents of a Singleton Set *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1680 | |
| 39910 | 1681 | definition the_elem :: "'a set \<Rightarrow> 'a" where | 
| 1682 |   "the_elem X = (THE x. X = {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: 
32120diff
changeset | 1683 | |
| 39910 | 1684 | lemma the_elem_eq [simp]: "the_elem {x} = x"
 | 
| 1685 | by (simp add: the_elem_def) | |
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1686 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1687 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1688 | subsubsection {* Least value operator *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1689 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1690 | lemma Least_mono: | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1691 | "mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1692 | ==> (LEAST y. y : f ` S) = f (LEAST x. x : S)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1693 |     -- {* Courtesy of Stephan Merz *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1694 | apply clarify | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1695 | apply (erule_tac P = "%x. x : S" in LeastI2_order, fast) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1696 | apply (rule LeastI2_order) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1697 | apply (auto elim: monoD intro!: order_antisym) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1698 | done | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1699 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1700 | subsection {* Misc *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1701 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1702 | text {* Rudimentary code generation *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1703 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1704 | lemma insert_code [code]: "insert y A x \<longleftrightarrow> y = x \<or> A x" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1705 | by (auto simp add: insert_compr Collect_def mem_def) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1706 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1707 | lemma vimage_code [code]: "(f -` A) x = A (f x)" | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1708 | by (simp add: vimage_def Collect_def mem_def) | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1709 | |
| 37677 | 1710 | hide_const (open) member | 
| 32135 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1711 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1712 | text {* Misc theorem and ML bindings *}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1713 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1714 | lemmas equalityI = subset_antisym | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1715 | |
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1716 | ML {*
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1717 | val Ball_def = @{thm Ball_def}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1718 | val Bex_def = @{thm Bex_def}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1719 | val CollectD = @{thm CollectD}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1720 | val CollectE = @{thm CollectE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1721 | val CollectI = @{thm CollectI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1722 | val Collect_conj_eq = @{thm Collect_conj_eq}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1723 | val Collect_mem_eq = @{thm Collect_mem_eq}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1724 | val IntD1 = @{thm IntD1}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1725 | val IntD2 = @{thm IntD2}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1726 | val IntE = @{thm IntE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1727 | val IntI = @{thm IntI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1728 | val Int_Collect = @{thm Int_Collect}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1729 | val UNIV_I = @{thm UNIV_I}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1730 | val UNIV_witness = @{thm UNIV_witness}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1731 | val UnE = @{thm UnE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1732 | val UnI1 = @{thm UnI1}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1733 | val UnI2 = @{thm UnI2}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1734 | val ballE = @{thm ballE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1735 | val ballI = @{thm ballI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1736 | val bexCI = @{thm bexCI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1737 | val bexE = @{thm bexE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1738 | val bexI = @{thm bexI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1739 | val bex_triv = @{thm bex_triv}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1740 | val bspec = @{thm bspec}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1741 | val contra_subsetD = @{thm contra_subsetD}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1742 | val distinct_lemma = @{thm distinct_lemma}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1743 | val eq_to_mono = @{thm eq_to_mono}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1744 | val equalityCE = @{thm equalityCE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1745 | val equalityD1 = @{thm equalityD1}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1746 | val equalityD2 = @{thm equalityD2}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1747 | val equalityE = @{thm equalityE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1748 | val equalityI = @{thm equalityI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1749 | val imageE = @{thm imageE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1750 | val imageI = @{thm imageI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1751 | val image_Un = @{thm image_Un}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1752 | val image_insert = @{thm image_insert}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1753 | val insert_commute = @{thm insert_commute}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1754 | val insert_iff = @{thm insert_iff}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1755 | val mem_Collect_eq = @{thm mem_Collect_eq}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1756 | val rangeE = @{thm rangeE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1757 | val rangeI = @{thm rangeI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1758 | val range_eqI = @{thm range_eqI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1759 | val subsetCE = @{thm subsetCE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1760 | val subsetD = @{thm subsetD}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1761 | val subsetI = @{thm subsetI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1762 | val subset_refl = @{thm subset_refl}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1763 | val subset_trans = @{thm subset_trans}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1764 | val vimageD = @{thm vimageD}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1765 | val vimageE = @{thm vimageE}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1766 | val vimageI = @{thm vimageI}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1767 | val vimageI2 = @{thm vimageI2}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1768 | val vimage_Collect = @{thm vimage_Collect}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1769 | val vimage_Int = @{thm vimage_Int}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1770 | val vimage_Un = @{thm vimage_Un}
 | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1771 | *} | 
| 
f645b51e8e54
set intersection and union now named inter and union; closer connection between set and lattice operations; factored out complete lattice
 haftmann parents: 
32120diff
changeset | 1772 | |
| 32077 
3698947146b2
closer relation of sets and complete lattices; corresponding consts, defs and syntax at similar places in theory text
 haftmann parents: 
32064diff
changeset | 1773 | end |