src/HOL/Set.thy
author haftmann
Thu, 05 Mar 2009 08:23:11 +0100
changeset 30304 d8e4cd2ac2a1
parent 29901 f4b3f8fbf599
child 30352 047f183c43b0
permissions -rw-r--r--
set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(*  Title:      HOL/Set.thy
12257
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
     2
    Author:     Tobias Nipkow, Lawrence C Paulson and Markus Wenzel
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
     5
header {* Set theory for higher-order logic *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
     6
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15120
diff changeset
     7
theory Set
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
     8
imports Lattices
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15120
diff changeset
     9
begin
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    10
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    11
text {* A set in HOL is simply a predicate. *}
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
2261
d926157c0a6a added "op :", "op ~:" syntax;
wenzelm
parents: 2006
diff changeset
    13
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    14
subsection {* Basic syntax *}
2261
d926157c0a6a added "op :", "op ~:" syntax;
wenzelm
parents: 2006
diff changeset
    15
3947
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
    16
global
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
    17
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
    18
types 'a set = "'a => bool"
3820
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
    19
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
consts
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    21
  Collect       :: "('a => bool) => 'a set"              -- "comprehension"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    22
  "op :"        :: "'a => 'a set => bool"                -- "membership"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    23
  insert        :: "'a => 'a set => 'a set"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    24
  Ball          :: "'a set => ('a => bool) => bool"      -- "bounded universal quantifiers"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    25
  Bex           :: "'a set => ('a => bool) => bool"      -- "bounded existential quantifiers"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
    26
  Bex1          :: "'a set => ('a => bool) => bool"      -- "bounded unique existential quantifiers"
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    27
  Pow           :: "'a set => 'a set set"                -- "powerset"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    28
  image         :: "('a => 'b) => 'a set => 'b set"      (infixr "`" 90)
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    29
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    30
local
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
    31
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 20380
diff changeset
    32
notation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
    33
  "op :"  ("op :") and
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
    34
  "op :"  ("(_/ : _)" [50, 51] 50)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    35
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
    36
abbreviation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
    37
  "not_mem x A == ~ (x : A)" -- "non-membership"
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
    38
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 20380
diff changeset
    39
notation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
    40
  not_mem  ("op ~:") and
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
    41
  not_mem  ("(_/ ~: _)" [50, 51] 50)
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
    42
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 20380
diff changeset
    43
notation (xsymbols)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
    44
  "op :"  ("op \<in>") and
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
    45
  "op :"  ("(_/ \<in> _)" [50, 51] 50) and
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
    46
  not_mem  ("op \<notin>") and
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    47
  not_mem  ("(_/ \<notin> _)" [50, 51] 50)
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
    48
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 20380
diff changeset
    49
notation (HTML output)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
    50
  "op :"  ("op \<in>") and
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
    51
  "op :"  ("(_/ \<in> _)" [50, 51] 50) and
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
    52
  not_mem  ("op \<notin>") and
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
    53
  not_mem  ("(_/ \<notin> _)" [50, 51] 50)
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
    54
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
syntax
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    56
  "@Coll"       :: "pttrn => bool => 'a set"              ("(1{_./ _})")
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    57
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    58
translations
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    59
  "{x. P}"      == "Collect (%x. P)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    60
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    61
definition empty :: "'a set" ("{}") where
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    62
  "empty \<equiv> {x. False}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    63
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    64
definition UNIV :: "'a set" where
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    65
  "UNIV \<equiv> {x. True}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    66
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    67
syntax
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    68
  "@Finset"     :: "args => 'a set"                       ("{(_)}")
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    69
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    70
translations
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    71
  "{x, xs}"     == "insert x {xs}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    72
  "{x}"         == "insert x {}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    73
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    74
definition Int :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Int" 70) where
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    75
  "A Int B \<equiv> {x. x \<in> A \<and> x \<in> B}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    76
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    77
definition Un :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Un" 65) where
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    78
  "A Un B \<equiv> {x. x \<in> A \<or> x \<in> B}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    79
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    80
notation (xsymbols)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    81
  "Int"  (infixl "\<inter>" 70) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    82
  "Un"  (infixl "\<union>" 65)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    83
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    84
notation (HTML output)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    85
  "Int"  (infixl "\<inter>" 70) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    86
  "Un"  (infixl "\<union>" 65)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    87
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
    88
syntax
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    89
  "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3ALL _:_./ _)" [0, 0, 10] 10)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    90
  "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3EX _:_./ _)" [0, 0, 10] 10)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
    91
  "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3EX! _:_./ _)" [0, 0, 10] 10)
22478
110f7f6f8a5d fixed typo
haftmann
parents: 22455
diff changeset
    92
  "_Bleast"     :: "id => 'a set => bool => 'a"           ("(3LEAST _:_./ _)" [0, 0, 10] 10)
18674
98d380757893 *** empty log message ***
nipkow
parents: 18447
diff changeset
    93
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 5931
diff changeset
    94
syntax (HOL)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    95
  "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3! _:_./ _)" [0, 0, 10] 10)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
    96
  "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3? _:_./ _)" [0, 0, 10] 10)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
    97
  "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3?! _:_./ _)" [0, 0, 10] 10)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
12114
a8e860c86252 eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents: 12023
diff changeset
    99
syntax (xsymbols)
14381
1189a8212a12 Modified UN and INT xsymbol syntax: made index subscript
nipkow
parents: 14335
diff changeset
   100
  "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3\<forall>_\<in>_./ _)" [0, 0, 10] 10)
1189a8212a12 Modified UN and INT xsymbol syntax: made index subscript
nipkow
parents: 14335
diff changeset
   101
  "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3\<exists>_\<in>_./ _)" [0, 0, 10] 10)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
   102
  "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3\<exists>!_\<in>_./ _)" [0, 0, 10] 10)
18674
98d380757893 *** empty log message ***
nipkow
parents: 18447
diff changeset
   103
  "_Bleast"     :: "id => 'a set => bool => 'a"           ("(3LEAST_\<in>_./ _)" [0, 0, 10] 10)
14381
1189a8212a12 Modified UN and INT xsymbol syntax: made index subscript
nipkow
parents: 14335
diff changeset
   104
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14551
diff changeset
   105
syntax (HTML output)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14551
diff changeset
   106
  "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3\<forall>_\<in>_./ _)" [0, 0, 10] 10)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14551
diff changeset
   107
  "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3\<exists>_\<in>_./ _)" [0, 0, 10] 10)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
   108
  "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3\<exists>!_\<in>_./ _)" [0, 0, 10] 10)
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14551
diff changeset
   109
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   110
translations
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   111
  "ALL x:A. P"  == "Ball A (%x. P)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   112
  "EX x:A. P"   == "Bex A (%x. P)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   113
  "EX! x:A. P"  == "Bex1 A (%x. P)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   114
  "LEAST x:A. P" => "LEAST x. x:A & P"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   115
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   116
definition INTER :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   117
  "INTER A B \<equiv> {y. \<forall>x\<in>A. y \<in> B x}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   118
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   119
definition UNION :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   120
  "UNION A B \<equiv> {y. \<exists>x\<in>A. y \<in> B x}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   121
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   122
definition Inter :: "'a set set \<Rightarrow> 'a set" where
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   123
  "Inter S \<equiv> INTER S (\<lambda>x. x)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   124
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   125
definition Union :: "'a set set \<Rightarrow> 'a set" where
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   126
  "Union S \<equiv> UNION S (\<lambda>x. x)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   127
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   128
notation (xsymbols)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   129
  Inter  ("\<Inter>_" [90] 90) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   130
  Union  ("\<Union>_" [90] 90)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   131
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   132
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   133
subsection {* Additional concrete syntax *}
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   134
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   135
syntax
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   136
  "@SetCompr"   :: "'a => idts => bool => 'a set"         ("(1{_ |/_./ _})")
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   137
  "@Collect"    :: "idt => 'a set => bool => 'a set"      ("(1{_ :/ _./ _})")
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   138
  "@INTER1"     :: "pttrns => 'b set => 'b set"           ("(3INT _./ _)" [0, 10] 10)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   139
  "@UNION1"     :: "pttrns => 'b set => 'b set"           ("(3UN _./ _)" [0, 10] 10)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   140
  "@INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3INT _:_./ _)" [0, 10] 10)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   141
  "@UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3UN _:_./ _)" [0, 10] 10)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   142
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 15102
diff changeset
   143
syntax (xsymbols)
15535
nipkow
parents: 15524
diff changeset
   144
  "@Collect"    :: "idt => 'a set => bool => 'a set"      ("(1{_ \<in>/ _./ _})")
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   145
  "@INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>_./ _)" [0, 10] 10)
22439
b709739c69e6 syntax: proper body priorty for derived binders;
wenzelm
parents: 22377
diff changeset
   146
  "@UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>_./ _)" [0, 10] 10)
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   147
  "@INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>_\<in>_./ _)" [0, 10] 10)
22439
b709739c69e6 syntax: proper body priorty for derived binders;
wenzelm
parents: 22377
diff changeset
   148
  "@UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>_\<in>_./ _)" [0, 10] 10)
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
   149
15120
f0359f75682e undid UN/INT syntax
nipkow
parents: 15102
diff changeset
   150
syntax (latex output)
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   151
  "@INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
22439
b709739c69e6 syntax: proper body priorty for derived binders;
wenzelm
parents: 22377
diff changeset
   152
  "@UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   153
  "@INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 10] 10)
22439
b709739c69e6 syntax: proper body priorty for derived binders;
wenzelm
parents: 22377
diff changeset
   154
  "@UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 10] 10)
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   155
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   156
translations
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   157
  "{x:A. P}"    => "{x. x:A & P}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   158
  "INT x y. B"  == "INT x. INT y. B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   159
  "INT x. B"    == "CONST INTER CONST UNIV (%x. B)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   160
  "INT x. B"    == "INT x:CONST UNIV. B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   161
  "INT x:A. B"  == "CONST INTER A (%x. B)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   162
  "UN x y. B"   == "UN x. UN y. B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   163
  "UN x. B"     == "CONST UNION CONST UNIV (%x. B)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   164
  "UN x. B"     == "UN x:CONST UNIV. B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   165
  "UN x:A. B"   == "CONST UNION A (%x. B)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   166
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   167
text {*
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
   168
  Note the difference between ordinary xsymbol syntax of indexed
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
   169
  unions and intersections (e.g.\ @{text"\<Union>a\<^isub>1\<in>A\<^isub>1. B"})
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
   170
  and their \LaTeX\ rendition: @{term"\<Union>a\<^isub>1\<in>A\<^isub>1. B"}. The
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
   171
  former does not make the index expression a subscript of the
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
   172
  union/intersection symbol because this leads to problems with nested
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   173
  subscripts in Proof General.
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   174
*}
2261
d926157c0a6a added "op :", "op ~:" syntax;
wenzelm
parents: 2006
diff changeset
   175
21333
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   176
abbreviation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   177
  subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
21819
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   178
  "subset \<equiv> less"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   179
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   180
abbreviation
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   181
  subset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
21819
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   182
  "subset_eq \<equiv> less_eq"
21333
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   183
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   184
notation (output)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   185
  subset  ("op <") and
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   186
  subset  ("(_/ < _)" [50, 51] 50) and
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   187
  subset_eq  ("op <=") and
21333
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   188
  subset_eq  ("(_/ <= _)" [50, 51] 50)
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   189
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   190
notation (xsymbols)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   191
  subset  ("op \<subset>") and
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   192
  subset  ("(_/ \<subset> _)" [50, 51] 50) and
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   193
  subset_eq  ("op \<subseteq>") and
21333
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   194
  subset_eq  ("(_/ \<subseteq> _)" [50, 51] 50)
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   195
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   196
notation (HTML output)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   197
  subset  ("op \<subset>") and
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   198
  subset  ("(_/ \<subset> _)" [50, 51] 50) and
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   199
  subset_eq  ("op \<subseteq>") and
21333
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   200
  subset_eq  ("(_/ \<subseteq> _)" [50, 51] 50)
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   201
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   202
abbreviation (input)
21819
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   203
  supset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   204
  "supset \<equiv> greater"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   205
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21384
diff changeset
   206
abbreviation (input)
21819
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   207
  supset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   208
  "supset_eq \<equiv> greater_eq"
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   209
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   210
notation (xsymbols)
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   211
  supset  ("op \<supset>") and
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   212
  supset  ("(_/ \<supset> _)" [50, 51] 50) and
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   213
  supset_eq  ("op \<supseteq>") and
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   214
  supset_eq  ("(_/ \<supseteq> _)" [50, 51] 50)
21333
eb291029d6cd dropped LOrder dependency
haftmann
parents: 21316
diff changeset
   215
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   216
abbreviation
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   217
  range :: "('a => 'b) => 'b set" where -- "of function"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   218
  "range f == f ` UNIV"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   219
14804
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   220
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   221
subsubsection "Bounded quantifiers"
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   222
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
   223
syntax (output)
14804
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   224
  "_setlessAll" :: "[idt, 'a, bool] => bool"  ("(3ALL _<_./ _)"  [0, 0, 10] 10)
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   225
  "_setlessEx"  :: "[idt, 'a, bool] => bool"  ("(3EX _<_./ _)"  [0, 0, 10] 10)
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   226
  "_setleAll"   :: "[idt, 'a, bool] => bool"  ("(3ALL _<=_./ _)" [0, 0, 10] 10)
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   227
  "_setleEx"    :: "[idt, 'a, bool] => bool"  ("(3EX _<=_./ _)" [0, 0, 10] 10)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
   228
  "_setleEx1"   :: "[idt, 'a, bool] => bool"  ("(3EX! _<=_./ _)" [0, 0, 10] 10)
14804
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   229
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   230
syntax (xsymbols)
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   231
  "_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: 14752
diff changeset
   232
  "_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: 14752
diff changeset
   233
  "_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: 14752
diff changeset
   234
  "_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: 19870
diff changeset
   235
  "_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: 14752
diff changeset
   236
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19637
diff changeset
   237
syntax (HOL output)
14804
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   238
  "_setlessAll" :: "[idt, 'a, bool] => bool"   ("(3! _<_./ _)"  [0, 0, 10] 10)
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   239
  "_setlessEx"  :: "[idt, 'a, bool] => bool"   ("(3? _<_./ _)"  [0, 0, 10] 10)
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   240
  "_setleAll"   :: "[idt, 'a, bool] => bool"   ("(3! _<=_./ _)" [0, 0, 10] 10)
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   241
  "_setleEx"    :: "[idt, 'a, bool] => bool"   ("(3? _<=_./ _)" [0, 0, 10] 10)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
   242
  "_setleEx1"   :: "[idt, 'a, bool] => bool"   ("(3?! _<=_./ _)" [0, 0, 10] 10)
14804
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   243
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   244
syntax (HTML output)
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   245
  "_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: 14752
diff changeset
   246
  "_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: 14752
diff changeset
   247
  "_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: 14752
diff changeset
   248
  "_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: 19870
diff changeset
   249
  "_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: 14752
diff changeset
   250
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   251
translations
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   252
 "\<forall>A\<subset>B. P"   =>  "ALL A. A \<subset> B --> P"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
   253
 "\<exists>A\<subset>B. P"   =>  "EX A. A \<subset> B & P"
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
   254
 "\<forall>A\<subseteq>B. P"   =>  "ALL A. A \<subseteq> B --> P"
14804
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   255
 "\<exists>A\<subseteq>B. P"   =>  "EX A. A \<subseteq> B & P"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
   256
 "\<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: 14752
diff changeset
   257
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   258
print_translation {*
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   259
let
22377
61610b1beedf tuned ML setup;
wenzelm
parents: 22172
diff changeset
   260
  val Type (set_type, _) = @{typ "'a set"};
61610b1beedf tuned ML setup;
wenzelm
parents: 22172
diff changeset
   261
  val All_binder = Syntax.binder_name @{const_syntax "All"};
61610b1beedf tuned ML setup;
wenzelm
parents: 22172
diff changeset
   262
  val Ex_binder = Syntax.binder_name @{const_syntax "Ex"};
61610b1beedf tuned ML setup;
wenzelm
parents: 22172
diff changeset
   263
  val impl = @{const_syntax "op -->"};
61610b1beedf tuned ML setup;
wenzelm
parents: 22172
diff changeset
   264
  val conj = @{const_syntax "op &"};
61610b1beedf tuned ML setup;
wenzelm
parents: 22172
diff changeset
   265
  val sbset = @{const_syntax "subset"};
61610b1beedf tuned ML setup;
wenzelm
parents: 22172
diff changeset
   266
  val sbset_eq = @{const_syntax "subset_eq"};
21819
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   267
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   268
  val trans =
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   269
   [((All_binder, impl, sbset), "_setlessAll"),
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   270
    ((All_binder, impl, sbset_eq), "_setleAll"),
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   271
    ((Ex_binder, conj, sbset), "_setlessEx"),
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   272
    ((Ex_binder, conj, sbset_eq), "_setleEx")];
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   273
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   274
  fun mk v v' c n P =
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   275
    if v = v' andalso not (Term.exists_subterm (fn Free (x, _) => x = v | _ => false) n)
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   276
    then Syntax.const c $ Syntax.mark_bound v' $ n $ P else raise Match;
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   277
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   278
  fun tr' q = (q,
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   279
    fn [Const ("_bound", _) $ Free (v, Type (T, _)), Const (c, _) $ (Const (d, _) $ (Const ("_bound", _) $ Free (v', _)) $ n) $ P] =>
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   280
         if T = (set_type) then case AList.lookup (op =) trans (q, c, d)
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   281
          of NONE => raise Match
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   282
           | SOME l => mk v v' l n P
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   283
         else raise Match
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   284
     | _ => raise Match);
14804
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   285
in
21819
8eb82ffcdd15 fixed syntax for bounded quantification
haftmann
parents: 21669
diff changeset
   286
  [tr' All_binder, tr' Ex_binder]
14804
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   287
end
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   288
*}
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   289
8de39d3e8eb6 Corrected printer bug for bounded quantifiers Q x<=y. P
nipkow
parents: 14752
diff changeset
   290
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   291
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   292
  \medskip Translate between @{text "{e | x1...xn. P}"} and @{text
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   293
  "{u. EX x1..xn. u = e & P}"}; @{text "{y. EX x1..xn. y = e & P}"} is
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   294
  only translated if @{text "[0..n] subset bvs(e)"}.
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   295
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   296
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   297
parse_translation {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   298
  let
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   299
    val ex_tr = snd (mk_binder_tr ("EX ", "Ex"));
3947
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
   300
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   301
    fun nvars (Const ("_idts", _) $ _ $ idts) = nvars idts + 1
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   302
      | nvars _ = 1;
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   303
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   304
    fun setcompr_tr [e, idts, b] =
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   305
      let
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   306
        val eq = Syntax.const "op =" $ Bound (nvars idts) $ e;
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   307
        val P = Syntax.const "op &" $ eq $ b;
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   308
        val exP = ex_tr [idts, P];
17784
5cbb52f2c478 Term.absdummy;
wenzelm
parents: 17715
diff changeset
   309
      in Syntax.const "Collect" $ Term.absdummy (dummyT, exP) end;
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   310
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   311
  in [("@SetCompr", setcompr_tr)] end;
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   312
*}
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   313
13763
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   314
(* To avoid eta-contraction of body: *)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   315
print_translation {*
13763
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   316
let
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   317
  fun btr' syn [A, Abs abs] =
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   318
    let val (x, t) = atomic_abs_tr' abs
13763
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   319
    in Syntax.const syn $ x $ A $ t end
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   320
in
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   321
[(@{const_syntax Ball}, btr' "_Ball"), (@{const_syntax Bex}, btr' "_Bex"),
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
   322
 (@{const_syntax UNION}, btr' "@UNION"),(@{const_syntax INTER}, btr' "@INTER")]
13763
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   323
end
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   324
*}
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   325
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   326
print_translation {*
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   327
let
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   328
  val ex_tr' = snd (mk_binder_tr' ("Ex", "DUMMY"));
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   329
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   330
  fun setcompr_tr' [Abs (abs as (_, _, P))] =
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   331
    let
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   332
      fun check (Const ("Ex", _) $ Abs (_, _, P), n) = check (P, n + 1)
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   333
        | check (Const ("op &", _) $ (Const ("op =", _) $ Bound m $ e) $ P, n) =
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   334
            n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   335
            ((0 upto (n - 1)) subset add_loose_bnos (e, 0, []))
13764
3e180bf68496 removed some problems with print translations
nipkow
parents: 13763
diff changeset
   336
        | check _ = false
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   337
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   338
        fun tr' (_ $ abs) =
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   339
          let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr' [abs]
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   340
          in Syntax.const "@SetCompr" $ e $ idts $ Q end;
13763
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   341
    in if check (P, 0) then tr' P
15535
nipkow
parents: 15524
diff changeset
   342
       else let val (x as _ $ Free(xN,_), t) = atomic_abs_tr' abs
nipkow
parents: 15524
diff changeset
   343
                val M = Syntax.const "@Coll" $ x $ t
nipkow
parents: 15524
diff changeset
   344
            in case t of
nipkow
parents: 15524
diff changeset
   345
                 Const("op &",_)
nipkow
parents: 15524
diff changeset
   346
                   $ (Const("op :",_) $ (Const("_bound",_) $ Free(yN,_)) $ A)
nipkow
parents: 15524
diff changeset
   347
                   $ P =>
nipkow
parents: 15524
diff changeset
   348
                   if xN=yN then Syntax.const "@Collect" $ x $ A $ P else M
nipkow
parents: 15524
diff changeset
   349
               | _ => M
nipkow
parents: 15524
diff changeset
   350
            end
13763
f94b569cd610 added print translations tha avoid eta contraction for important binders.
nipkow
parents: 13653
diff changeset
   351
    end;
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   352
  in [("Collect", setcompr_tr')] end;
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   353
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   354
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   355
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   356
subsection {* Rules and definitions *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   357
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   358
text {* Isomorphisms between predicates and sets. *}
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   359
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   360
defs
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27824
diff changeset
   361
  mem_def [code]: "x : S == S x"
4e74209f113e `code func` now just `code`
haftmann
parents: 27824
diff changeset
   362
  Collect_def [code]: "Collect P == P"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   363
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   364
defs
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   365
  Ball_def:     "Ball A P       == ALL x. x:A --> P(x)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   366
  Bex_def:      "Bex A P        == EX x. x:A & P(x)"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19870
diff changeset
   367
  Bex1_def:     "Bex1 A P       == EX! x. x:A & P(x)"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   368
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   369
instantiation "fun" :: (type, minus) minus
25510
38c15efe603b adjustions to due to instance target
haftmann
parents: 25502
diff changeset
   370
begin
38c15efe603b adjustions to due to instance target
haftmann
parents: 25502
diff changeset
   371
38c15efe603b adjustions to due to instance target
haftmann
parents: 25502
diff changeset
   372
definition
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   373
  fun_diff_def: "A - B = (%x. A x - B x)"
25762
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 25510
diff changeset
   374
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 25510
diff changeset
   375
instance ..
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 25510
diff changeset
   376
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 25510
diff changeset
   377
end
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 25510
diff changeset
   378
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   379
instantiation bool :: minus
25762
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 25510
diff changeset
   380
begin
25510
38c15efe603b adjustions to due to instance target
haftmann
parents: 25502
diff changeset
   381
38c15efe603b adjustions to due to instance target
haftmann
parents: 25502
diff changeset
   382
definition
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   383
  bool_diff_def: "A - B = (A & ~ B)"
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   384
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   385
instance ..
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   386
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   387
end
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   388
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   389
instantiation "fun" :: (type, uminus) uminus
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   390
begin
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   391
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   392
definition
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   393
  fun_Compl_def: "- A = (%x. - A x)"
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   394
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   395
instance ..
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   396
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   397
end
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   398
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   399
instantiation bool :: uminus
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   400
begin
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   401
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   402
definition
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   403
  bool_Compl_def: "- A = (~ A)"
25510
38c15efe603b adjustions to due to instance target
haftmann
parents: 25502
diff changeset
   404
38c15efe603b adjustions to due to instance target
haftmann
parents: 25502
diff changeset
   405
instance ..
38c15efe603b adjustions to due to instance target
haftmann
parents: 25502
diff changeset
   406
38c15efe603b adjustions to due to instance target
haftmann
parents: 25502
diff changeset
   407
end
22744
5cbe966d67a2 Isar definitions are now added explicitly to code theorem table
haftmann
parents: 22478
diff changeset
   408
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   409
defs
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   410
  Pow_def:      "Pow A          == {B. B <= A}"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   411
  insert_def:   "insert a B     == {x. x=a} Un B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   412
  image_def:    "f`A            == {y. EX x:A. y = f(x)}"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   413
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   414
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   415
subsection {* Lemmas and proof tool setup *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   416
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   417
subsubsection {* Relating predicates and sets *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   418
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   419
lemma mem_Collect_eq [iff]: "(a : {x. P(x)}) = P(a)"
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   420
  by (simp add: Collect_def mem_def)
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   421
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   422
lemma Collect_mem_eq [simp]: "{x. x:A} = A"
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   423
  by (simp add: Collect_def mem_def)
17085
5b57f995a179 more simprules now have names
paulson
parents: 17084
diff changeset
   424
12257
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
   425
lemma CollectI: "P(a) ==> a : {x. P(x)}"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   426
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   427
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   428
lemma CollectD: "a : {x. P(x)} ==> P(a)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   429
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   430
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   431
lemma Collect_cong: "(!!x. P x = Q x) ==> {x. P(x)} = {x. Q(x)}"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   432
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   433
12257
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
   434
lemmas CollectE = CollectD [elim_format]
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   435
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   436
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   437
subsubsection {* Bounded quantifiers *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   438
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   439
lemma ballI [intro!]: "(!!x. x:A ==> P x) ==> ALL x:A. P x"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   440
  by (simp add: Ball_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   441
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   442
lemmas strip = impI allI ballI
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   443
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   444
lemma bspec [dest?]: "ALL x:A. P x ==> x:A ==> P x"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   445
  by (simp add: Ball_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   446
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   447
lemma ballE [elim]: "ALL x:A. P x ==> (P x ==> Q) ==> (x ~: A ==> Q) ==> Q"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   448
  by (unfold Ball_def) blast
22139
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
   449
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
   450
ML {* bind_thm ("rev_ballE", permute_prems 1 1 @{thm ballE}) *}
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   451
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   452
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   453
  \medskip This tactic takes assumptions @{prop "ALL x:A. P x"} and
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   454
  @{prop "a:A"}; creates assumption @{prop "P a"}.
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   455
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   456
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   457
ML {*
22139
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
   458
  fun ball_tac i = etac @{thm ballE} i THEN contr_tac (i + 1)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   459
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   460
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   461
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   462
  Gives better instantiation for bound:
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   463
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   464
26339
7825c83c9eff eliminated change_claset/simpset;
wenzelm
parents: 26150
diff changeset
   465
declaration {* fn _ =>
7825c83c9eff eliminated change_claset/simpset;
wenzelm
parents: 26150
diff changeset
   466
  Classical.map_cs (fn cs => cs addbefore ("bspec", datac @{thm bspec} 1))
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   467
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   468
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   469
lemma bexI [intro]: "P x ==> x:A ==> EX x:A. P x"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   470
  -- {* Normally the best argument order: @{prop "P x"} constrains the
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   471
    choice of @{prop "x:A"}. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   472
  by (unfold Bex_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   473
13113
5eb9be7b72a5 rev_bexI [intro?];
wenzelm
parents: 13103
diff changeset
   474
lemma rev_bexI [intro?]: "x:A ==> P x ==> EX x:A. P x"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   475
  -- {* The best argument order when there is only one @{prop "x:A"}. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   476
  by (unfold Bex_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   477
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   478
lemma bexCI: "(ALL x:A. ~P x ==> P a) ==> a:A ==> EX x:A. P x"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   479
  by (unfold Bex_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   480
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   481
lemma bexE [elim!]: "EX x:A. P x ==> (!!x. x:A ==> P x ==> Q) ==> Q"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   482
  by (unfold Bex_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   483
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   484
lemma ball_triv [simp]: "(ALL x:A. P) = ((EX x. x:A) --> P)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   485
  -- {* Trival rewrite rule. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   486
  by (simp add: Ball_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   487
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   488
lemma bex_triv [simp]: "(EX x:A. P) = ((EX x. x:A) & P)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   489
  -- {* Dual form for existentials. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   490
  by (simp add: Bex_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   491
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   492
lemma bex_triv_one_point1 [simp]: "(EX x:A. x = a) = (a:A)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   493
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   494
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   495
lemma bex_triv_one_point2 [simp]: "(EX x:A. a = x) = (a:A)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   496
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   497
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   498
lemma bex_one_point1 [simp]: "(EX x:A. x = a & P x) = (a:A & P a)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   499
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   500
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   501
lemma bex_one_point2 [simp]: "(EX x:A. a = x & P x) = (a:A & P a)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   502
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   503
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   504
lemma ball_one_point1 [simp]: "(ALL x:A. x = a --> P x) = (a:A --> P a)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   505
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   506
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   507
lemma ball_one_point2 [simp]: "(ALL x:A. a = x --> P x) = (a:A --> P a)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   508
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   509
26480
544cef16045b replaced 'ML_setup' by 'ML';
wenzelm
parents: 26339
diff changeset
   510
ML {*
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13421
diff changeset
   511
  local
22139
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
   512
    val unfold_bex_tac = unfold_tac @{thms "Bex_def"};
18328
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 18315
diff changeset
   513
    fun prove_bex_tac ss = unfold_bex_tac ss THEN Quantifier1.prove_one_point_ex_tac;
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   514
    val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac;
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   515
22139
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
   516
    val unfold_ball_tac = unfold_tac @{thms "Ball_def"};
18328
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 18315
diff changeset
   517
    fun prove_ball_tac ss = unfold_ball_tac ss THEN Quantifier1.prove_one_point_all_tac;
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   518
    val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac;
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   519
  in
18328
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 18315
diff changeset
   520
    val defBEX_regroup = Simplifier.simproc (the_context ())
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13421
diff changeset
   521
      "defined BEX" ["EX x:A. P x & Q x"] rearrange_bex;
18328
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 18315
diff changeset
   522
    val defBALL_regroup = Simplifier.simproc (the_context ())
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13421
diff changeset
   523
      "defined BALL" ["ALL x:A. P x --> Q x"] rearrange_ball;
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   524
  end;
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13421
diff changeset
   525
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13421
diff changeset
   526
  Addsimprocs [defBALL_regroup, defBEX_regroup];
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   527
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   528
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   529
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   530
subsubsection {* Congruence rules *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   531
16636
1ed737a98198 Added strong_ball_cong and strong_bex_cong (these are now the standard
berghofe
parents: 15950
diff changeset
   532
lemma ball_cong:
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   533
  "A = B ==> (!!x. x:B ==> P x = Q x) ==>
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   534
    (ALL x:A. P x) = (ALL x:B. Q x)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   535
  by (simp add: Ball_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   536
16636
1ed737a98198 Added strong_ball_cong and strong_bex_cong (these are now the standard
berghofe
parents: 15950
diff changeset
   537
lemma strong_ball_cong [cong]:
1ed737a98198 Added strong_ball_cong and strong_bex_cong (these are now the standard
berghofe
parents: 15950
diff changeset
   538
  "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: 15950
diff changeset
   539
    (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: 15950
diff changeset
   540
  by (simp add: simp_implies_def Ball_def)
1ed737a98198 Added strong_ball_cong and strong_bex_cong (these are now the standard
berghofe
parents: 15950
diff changeset
   541
1ed737a98198 Added strong_ball_cong and strong_bex_cong (these are now the standard
berghofe
parents: 15950
diff changeset
   542
lemma bex_cong:
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   543
  "A = B ==> (!!x. x:B ==> P x = Q x) ==>
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   544
    (EX x:A. P x) = (EX x:B. Q x)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   545
  by (simp add: Bex_def cong: conj_cong)
1273
6960ec882bca added 8bit pragmas
regensbu
parents: 1068
diff changeset
   546
16636
1ed737a98198 Added strong_ball_cong and strong_bex_cong (these are now the standard
berghofe
parents: 15950
diff changeset
   547
lemma strong_bex_cong [cong]:
1ed737a98198 Added strong_ball_cong and strong_bex_cong (these are now the standard
berghofe
parents: 15950
diff changeset
   548
  "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: 15950
diff changeset
   549
    (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: 15950
diff changeset
   550
  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: 15950
diff changeset
   551
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 5931
diff changeset
   552
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   553
subsubsection {* Subsets *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   554
19295
c5d236fe9668 subsetI is often necessary
paulson
parents: 19277
diff changeset
   555
lemma subsetI [atp,intro!]: "(!!x. x:A ==> x:B) ==> A \<subseteq> B"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   556
  by (auto simp add: mem_def intro: predicate1I)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   557
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   558
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   559
  \medskip Map the type @{text "'a set => anything"} to just @{typ
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   560
  'a}; for overloading constants whose first argument has type @{typ
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   561
  "'a set"}.
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   562
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   563
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   564
lemma subsetD [elim]: "A \<subseteq> B ==> c \<in> A ==> c \<in> B"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   565
  -- {* Rule in Modus Ponens style. *}
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   566
  by (unfold mem_def) blast
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   567
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   568
declare subsetD [intro?] -- FIXME
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   569
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   570
lemma rev_subsetD: "c \<in> A ==> A \<subseteq> B ==> c \<in> B"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   571
  -- {* The same, with reversed premises for use with @{text erule} --
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   572
      cf @{text rev_mp}. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   573
  by (rule subsetD)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   574
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   575
declare rev_subsetD [intro?] -- FIXME
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   576
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   577
text {*
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   578
  \medskip Converts @{prop "A \<subseteq> B"} to @{prop "x \<in> A ==> x \<in> B"}.
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   579
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   580
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   581
ML {*
22139
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
   582
  fun impOfSubs th = th RSN (2, @{thm rev_subsetD})
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   583
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   584
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   585
lemma subsetCE [elim]: "A \<subseteq>  B ==> (c \<notin> A ==> P) ==> (c \<in> B ==> P) ==> P"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   586
  -- {* Classical elimination rule. *}
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   587
  by (unfold mem_def) blast
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   588
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   589
lemma subset_eq: "A \<le> B = (\<forall>x\<in>A. x \<in> B)" by blast
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   590
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   591
text {*
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   592
  \medskip Takes assumptions @{prop "A \<subseteq> B"}; @{prop "c \<in> A"} and
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   593
  creates the assumption @{prop "c \<in> B"}.
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   594
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   595
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   596
ML {*
22139
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
   597
  fun set_mp_tac i = etac @{thm subsetCE} i THEN mp_tac i
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   598
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   599
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   600
lemma contra_subsetD: "A \<subseteq> B ==> c \<notin> B ==> c \<notin> A"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   601
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   602
19175
c6e4b073f6a5 subset_refl now included using the atp attribute
paulson
parents: 18851
diff changeset
   603
lemma subset_refl [simp,atp]: "A \<subseteq> A"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   604
  by fast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   605
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   606
lemma subset_trans: "A \<subseteq> B ==> B \<subseteq> C ==> A \<subseteq> C"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   607
  by blast
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   608
2261
d926157c0a6a added "op :", "op ~:" syntax;
wenzelm
parents: 2006
diff changeset
   609
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   610
subsubsection {* Equality *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   611
13865
0a6bf71955b0 moved one proof, added another
paulson
parents: 13860
diff changeset
   612
lemma set_ext: assumes prem: "(!!x. (x:A) = (x:B))" shows "A = B"
0a6bf71955b0 moved one proof, added another
paulson
parents: 13860
diff changeset
   613
  apply (rule prem [THEN ext, THEN arg_cong, THEN box_equals])
0a6bf71955b0 moved one proof, added another
paulson
parents: 13860
diff changeset
   614
   apply (rule Collect_mem_eq)
0a6bf71955b0 moved one proof, added another
paulson
parents: 13860
diff changeset
   615
  apply (rule Collect_mem_eq)
0a6bf71955b0 moved one proof, added another
paulson
parents: 13860
diff changeset
   616
  done
0a6bf71955b0 moved one proof, added another
paulson
parents: 13860
diff changeset
   617
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15535
diff changeset
   618
(* Due to Brian Huffman *)
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15535
diff changeset
   619
lemma expand_set_eq: "(A = B) = (ALL x. (x:A) = (x:B))"
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15535
diff changeset
   620
by(auto intro:set_ext)
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15535
diff changeset
   621
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   622
lemma subset_antisym [intro!]: "A \<subseteq> B ==> B \<subseteq> A ==> A = B"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   623
  -- {* Anti-symmetry of the subset relation. *}
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
   624
  by (iprover intro: set_ext subsetD)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   625
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   626
lemmas equalityI [intro!] = subset_antisym
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   627
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   628
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   629
  \medskip Equality rules from ZF set theory -- are they appropriate
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   630
  here?
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   631
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   632
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   633
lemma equalityD1: "A = B ==> A \<subseteq> B"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   634
  by (simp add: subset_refl)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   635
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   636
lemma equalityD2: "A = B ==> B \<subseteq> A"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   637
  by (simp add: subset_refl)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   638
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   639
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   640
  \medskip Be careful when adding this to the claset as @{text
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   641
  subset_empty} is in the simpset: @{prop "A = {}"} goes to @{prop "{}
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   642
  \<subseteq> A"} and @{prop "A \<subseteq> {}"} and then back to @{prop "A = {}"}!
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   643
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   644
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   645
lemma equalityE: "A = B ==> (A \<subseteq> B ==> B \<subseteq> A ==> P) ==> P"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   646
  by (simp add: subset_refl)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   647
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   648
lemma equalityCE [elim]:
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   649
    "A = B ==> (c \<in> A ==> c \<in> B ==> P) ==> (c \<notin> A ==> c \<notin> B ==> P) ==> P"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   650
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   651
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   652
lemma eqset_imp_iff: "A = B ==> (x : A) = (x : B)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   653
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   654
13865
0a6bf71955b0 moved one proof, added another
paulson
parents: 13860
diff changeset
   655
lemma eqelem_imp_iff: "x = y ==> (x : A) = (y : A)"
0a6bf71955b0 moved one proof, added another
paulson
parents: 13860
diff changeset
   656
  by simp
0a6bf71955b0 moved one proof, added another
paulson
parents: 13860
diff changeset
   657
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   658
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   659
subsubsection {* The universal set -- UNIV *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   660
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   661
lemma UNIV_I [simp]: "x : UNIV"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   662
  by (simp add: UNIV_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   663
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   664
declare UNIV_I [intro]  -- {* unsafe makes it less likely to cause problems *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   665
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   666
lemma UNIV_witness [intro?]: "EX x. x : UNIV"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   667
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   668
18144
4edcb5fdc3b0 duplicate axioms in ATP linkup, and general fixes
paulson
parents: 17875
diff changeset
   669
lemma subset_UNIV [simp]: "A \<subseteq> UNIV"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   670
  by (rule subsetI) (rule UNIV_I)
2388
d1f0505fc602 added set inclusion symbol syntax;
wenzelm
parents: 2372
diff changeset
   671
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   672
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   673
  \medskip Eta-contracting these two rules (to remove @{text P})
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   674
  causes them to be ignored because of their interaction with
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   675
  congruence rules.
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   676
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   677
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   678
lemma ball_UNIV [simp]: "Ball UNIV P = All P"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   679
  by (simp add: Ball_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   680
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   681
lemma bex_UNIV [simp]: "Bex UNIV P = Ex P"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   682
  by (simp add: Bex_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   683
26150
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
   684
lemma UNIV_eq_I: "(\<And>x. x \<in> A) \<Longrightarrow> UNIV = A"
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
   685
  by auto
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
   686
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   687
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   688
subsubsection {* The empty set *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   689
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   690
lemma empty_iff [simp]: "(c : {}) = False"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   691
  by (simp add: empty_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   692
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   693
lemma emptyE [elim!]: "a : {} ==> P"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   694
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   695
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   696
lemma empty_subsetI [iff]: "{} \<subseteq> A"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   697
    -- {* One effect is to delete the ASSUMPTION @{prop "{} <= A"} *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   698
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   699
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   700
lemma equals0I: "(!!y. y \<in> A ==> False) ==> A = {}"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   701
  by blast
2388
d1f0505fc602 added set inclusion symbol syntax;
wenzelm
parents: 2372
diff changeset
   702
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   703
lemma equals0D: "A = {} ==> a \<notin> A"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   704
    -- {* Use for reasoning about disjointness: @{prop "A Int B = {}"} *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   705
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   706
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   707
lemma ball_empty [simp]: "Ball {} P = True"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   708
  by (simp add: Ball_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   709
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   710
lemma bex_empty [simp]: "Bex {} P = False"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   711
  by (simp add: Bex_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   712
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   713
lemma UNIV_not_empty [iff]: "UNIV ~= {}"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   714
  by (blast elim: equalityE)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   715
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   716
12023
wenzelm
parents: 12020
diff changeset
   717
subsubsection {* The Powerset operator -- Pow *}
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   718
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   719
lemma Pow_iff [iff]: "(A \<in> Pow B) = (A \<subseteq> B)"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   720
  by (simp add: Pow_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   721
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   722
lemma PowI: "A \<subseteq> B ==> A \<in> Pow B"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   723
  by (simp add: Pow_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   724
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   725
lemma PowD: "A \<in> Pow B ==> A \<subseteq> B"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   726
  by (simp add: Pow_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   727
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   728
lemma Pow_bottom: "{} \<in> Pow B"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   729
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   730
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   731
lemma Pow_top: "A \<in> Pow A"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   732
  by (simp add: subset_refl)
2684
9781d63ef063 added proper subset symbols syntax;
wenzelm
parents: 2412
diff changeset
   733
2388
d1f0505fc602 added set inclusion symbol syntax;
wenzelm
parents: 2372
diff changeset
   734
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   735
subsubsection {* Set complement *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   736
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   737
lemma Compl_iff [simp]: "(c \<in> -A) = (c \<notin> A)"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   738
  by (simp add: mem_def fun_Compl_def bool_Compl_def)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   739
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   740
lemma ComplI [intro!]: "(c \<in> A ==> False) ==> c \<in> -A"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   741
  by (unfold mem_def fun_Compl_def bool_Compl_def) blast
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   742
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   743
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   744
  \medskip This form, with negated conclusion, works well with the
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   745
  Classical prover.  Negated assumptions behave like formulae on the
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   746
  right side of the notional turnstile ... *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   747
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 17002
diff changeset
   748
lemma ComplD [dest!]: "c : -A ==> c~:A"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   749
  by (simp add: mem_def fun_Compl_def bool_Compl_def)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   750
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 17002
diff changeset
   751
lemmas ComplE = ComplD [elim_format]
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   752
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   753
lemma Compl_eq: "- A = {x. ~ x : A}" by blast
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   754
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   755
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   756
subsubsection {* Binary union -- Un *}
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   757
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   758
lemma Un_iff [simp]: "(c : A Un B) = (c:A | c:B)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   759
  by (unfold Un_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   760
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   761
lemma UnI1 [elim?]: "c:A ==> c : A Un B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   762
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   763
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   764
lemma UnI2 [elim?]: "c:B ==> c : A Un B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   765
  by simp
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   766
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   767
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   768
  \medskip Classical introduction rule: no commitment to @{prop A} vs
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   769
  @{prop B}.
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   770
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   771
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   772
lemma UnCI [intro!]: "(c~:B ==> c:A) ==> c : A Un B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   773
  by auto
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   774
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   775
lemma UnE [elim!]: "c : A Un B ==> (c:A ==> P) ==> (c:B ==> P) ==> P"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   776
  by (unfold Un_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   777
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   778
12023
wenzelm
parents: 12020
diff changeset
   779
subsubsection {* Binary intersection -- Int *}
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   780
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   781
lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   782
  by (unfold Int_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   783
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   784
lemma IntI [intro!]: "c:A ==> c:B ==> c : A Int B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   785
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   786
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   787
lemma IntD1: "c : A Int B ==> c:A"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   788
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   789
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   790
lemma IntD2: "c : A Int B ==> c:B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   791
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   792
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   793
lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   794
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   795
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   796
12023
wenzelm
parents: 12020
diff changeset
   797
subsubsection {* Set difference *}
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   798
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   799
lemma Diff_iff [simp]: "(c : A - B) = (c:A & c~:B)"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   800
  by (simp add: mem_def fun_diff_def bool_diff_def)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   801
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   802
lemma DiffI [intro!]: "c : A ==> c ~: B ==> c : A - B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   803
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   804
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   805
lemma DiffD1: "c : A - B ==> c : A"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   806
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   807
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   808
lemma DiffD2: "c : A - B ==> c : B ==> P"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   809
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   810
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   811
lemma DiffE [elim!]: "c : A - B ==> (c:A ==> c~:B ==> P) ==> P"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   812
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   813
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   814
lemma set_diff_eq: "A - B = {x. x : A & ~ x : B}" by blast
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
   815
29901
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29691
diff changeset
   816
lemma Compl_eq_Diff_UNIV: "-A = (UNIV - A)"
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29691
diff changeset
   817
by blast
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29691
diff changeset
   818
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   819
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   820
subsubsection {* Augmenting a set -- insert *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   821
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   822
lemma insert_iff [simp]: "(a : insert b A) = (a = b | a:A)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   823
  by (unfold insert_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   824
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   825
lemma insertI1: "a : insert a B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   826
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   827
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   828
lemma insertI2: "a : B ==> a : insert b B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   829
  by simp
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   830
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   831
lemma insertE [elim!]: "a : insert b A ==> (a = b ==> P) ==> (a:A ==> P) ==> P"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   832
  by (unfold insert_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   833
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   834
lemma insertCI [intro!]: "(a~:B ==> a = b) ==> a: insert b B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   835
  -- {* Classical introduction rule. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   836
  by auto
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   837
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   838
lemma subset_insert_iff: "(A \<subseteq> insert x B) = (if x:A then A - {x} \<subseteq> B else A \<subseteq> B)"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   839
  by auto
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   840
24730
a87d8d31abc0 convenient obtain rule for sets
haftmann
parents: 24658
diff changeset
   841
lemma set_insert:
a87d8d31abc0 convenient obtain rule for sets
haftmann
parents: 24658
diff changeset
   842
  assumes "x \<in> A"
a87d8d31abc0 convenient obtain rule for sets
haftmann
parents: 24658
diff changeset
   843
  obtains B where "A = insert x B" and "x \<notin> B"
a87d8d31abc0 convenient obtain rule for sets
haftmann
parents: 24658
diff changeset
   844
proof
a87d8d31abc0 convenient obtain rule for sets
haftmann
parents: 24658
diff changeset
   845
  from assms show "A = insert x (A - {x})" by blast
a87d8d31abc0 convenient obtain rule for sets
haftmann
parents: 24658
diff changeset
   846
next
a87d8d31abc0 convenient obtain rule for sets
haftmann
parents: 24658
diff changeset
   847
  show "x \<notin> A - {x}" by blast
a87d8d31abc0 convenient obtain rule for sets
haftmann
parents: 24658
diff changeset
   848
qed
a87d8d31abc0 convenient obtain rule for sets
haftmann
parents: 24658
diff changeset
   849
25287
094dab519ff5 added lemmas
nipkow
parents: 24730
diff changeset
   850
lemma insert_ident: "x ~: A ==> x ~: B ==> (insert x A = insert x B) = (A = B)"
094dab519ff5 added lemmas
nipkow
parents: 24730
diff changeset
   851
by auto
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   852
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   853
subsubsection {* Singletons, using insert *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   854
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   855
lemma singletonI [intro!,noatp]: "a : {a}"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   856
    -- {* Redundant? But unlike @{text insertCI}, it proves the subgoal immediately! *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   857
  by (rule insertI1)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   858
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   859
lemma singletonD [dest!,noatp]: "b : {a} ==> b = a"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   860
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   861
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 17002
diff changeset
   862
lemmas singletonE = singletonD [elim_format]
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   863
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   864
lemma singleton_iff: "(b : {a}) = (b = a)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   865
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   866
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   867
lemma singleton_inject [dest!]: "{a} = {b} ==> a = b"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   868
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   869
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   870
lemma singleton_insert_inj_eq [iff,noatp]:
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   871
     "({b} = insert a A) = (a = b & A \<subseteq> {b})"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   872
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   873
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   874
lemma singleton_insert_inj_eq' [iff,noatp]:
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   875
     "(insert a A = {b}) = (a = b & A \<subseteq> {b})"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   876
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   877
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   878
lemma subset_singletonD: "A \<subseteq> {x} ==> A = {} | A = {x}"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   879
  by fast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   880
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   881
lemma singleton_conv [simp]: "{x. x = a} = {a}"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   882
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   883
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   884
lemma singleton_conv2 [simp]: "{x. a = x} = {a}"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   885
  by blast
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   886
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
   887
lemma diff_single_insert: "A - {x} \<subseteq> B ==> x \<in> A ==> A \<subseteq> insert x B"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   888
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   889
19870
ef037d1b32d1 new results
paulson
parents: 19656
diff changeset
   890
lemma doubleton_eq_iff: "({a,b} = {c,d}) = (a=c & b=d | a=d & b=c)"
ef037d1b32d1 new results
paulson
parents: 19656
diff changeset
   891
  by (blast elim: equalityE)
ef037d1b32d1 new results
paulson
parents: 19656
diff changeset
   892
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   893
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   894
subsubsection {* Unions of families *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   895
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   896
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   897
  @{term [source] "UN x:A. B x"} is @{term "Union (B`A)"}.
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   898
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   899
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   900
declare UNION_def [noatp]
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   901
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   902
lemma UN_iff [simp]: "(b: (UN x:A. B x)) = (EX x:A. b: B x)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   903
  by (unfold UNION_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   904
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   905
lemma UN_I [intro]: "a:A ==> b: B a ==> b: (UN x:A. B x)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   906
  -- {* The order of the premises presupposes that @{term A} is rigid;
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   907
    @{term b} may be flexible. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   908
  by auto
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   909
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   910
lemma UN_E [elim!]: "b : (UN x:A. B x) ==> (!!x. x:A ==> b: B x ==> R) ==> R"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   911
  by (unfold UNION_def) blast
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   912
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   913
lemma UN_cong [cong]:
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   914
    "A = B ==> (!!x. x:B ==> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   915
  by (simp add: UNION_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   916
29691
9f03b5f847cd Added strong congruence rule for UN.
berghofe
parents: 28562
diff changeset
   917
lemma strong_UN_cong:
9f03b5f847cd Added strong congruence rule for UN.
berghofe
parents: 28562
diff changeset
   918
    "A = B ==> (!!x. x:B =simp=> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)"
9f03b5f847cd Added strong congruence rule for UN.
berghofe
parents: 28562
diff changeset
   919
  by (simp add: UNION_def simp_implies_def)
9f03b5f847cd Added strong congruence rule for UN.
berghofe
parents: 28562
diff changeset
   920
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   921
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   922
subsubsection {* Intersections of families *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   923
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   924
text {* @{term [source] "INT x:A. B x"} is @{term "Inter (B`A)"}. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   925
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   926
lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   927
  by (unfold INTER_def) blast
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   928
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   929
lemma INT_I [intro!]: "(!!x. x:A ==> b: B x) ==> b : (INT x:A. B x)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   930
  by (unfold INTER_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   931
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   932
lemma INT_D [elim]: "b : (INT x:A. B x) ==> a:A ==> b: B a"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   933
  by auto
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   934
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   935
lemma INT_E [elim]: "b : (INT x:A. B x) ==> (b: B a ==> R) ==> (a~:A ==> R) ==> R"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   936
  -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a:A"}. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   937
  by (unfold INTER_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   938
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   939
lemma INT_cong [cong]:
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   940
    "A = B ==> (!!x. x:B ==> C x = D x) ==> (INT x:A. C x) = (INT x:B. D x)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   941
  by (simp add: INTER_def)
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 5931
diff changeset
   942
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   943
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   944
subsubsection {* Union *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   945
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   946
lemma Union_iff [simp,noatp]: "(A : Union C) = (EX X:C. A:X)"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   947
  by (unfold Union_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   948
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   949
lemma UnionI [intro]: "X:C ==> A:X ==> A : Union C"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   950
  -- {* The order of the premises presupposes that @{term C} is rigid;
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   951
    @{term A} may be flexible. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   952
  by auto
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   953
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   954
lemma UnionE [elim!]: "A : Union C ==> (!!X. A:X ==> X:C ==> R) ==> R"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   955
  by (unfold Union_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   956
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   957
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   958
subsubsection {* Inter *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   959
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   960
lemma Inter_iff [simp,noatp]: "(A : Inter C) = (ALL X:C. A:X)"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   961
  by (unfold Inter_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   962
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   963
lemma InterI [intro!]: "(!!X. X:C ==> A:X) ==> A : Inter C"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   964
  by (simp add: Inter_def)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   965
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   966
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   967
  \medskip A ``destruct'' rule -- every @{term X} in @{term C}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   968
  contains @{term A} as an element, but @{prop "A:X"} can hold when
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   969
  @{prop "X:C"} does not!  This rule is analogous to @{text spec}.
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   970
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   971
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   972
lemma InterD [elim]: "A : Inter C ==> X:C ==> A:X"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   973
  by auto
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   974
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   975
lemma InterE [elim]: "A : Inter C ==> (X~:C ==> R) ==> (A:X ==> R) ==> R"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   976
  -- {* ``Classical'' elimination rule -- does not require proving
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   977
    @{prop "X:C"}. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   978
  by (unfold Inter_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   980
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   981
  \medskip Image of a set under a function.  Frequently @{term b} does
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   982
  not have the syntactic form of @{term "f x"}.
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   983
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   984
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   985
declare image_def [noatp]
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
   986
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   987
lemma image_eqI [simp, intro]: "b = f x ==> x:A ==> b : f`A"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   988
  by (unfold image_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   989
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   990
lemma imageI: "x : A ==> f x : f ` A"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   991
  by (rule image_eqI) (rule refl)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   992
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   993
lemma rev_image_eqI: "x:A ==> b = f x ==> b : f`A"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   994
  -- {* This version's more effective when we already have the
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   995
    required @{term x}. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   996
  by (unfold image_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   997
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   998
lemma imageE [elim!]:
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
   999
  "b : (%x. f x)`A ==> (!!x. b = f x ==> x:A ==> P) ==> P"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1000
  -- {* The eta-expansion gives variable-name preservation. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1001
  by (unfold image_def) blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1002
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1003
lemma image_Un: "f`(A Un B) = f`A Un f`B"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1004
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1005
26150
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  1006
lemma image_eq_UN: "f`A = (UN x:A. {f x})"
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  1007
  by blast
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  1008
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1009
lemma image_iff: "(z : f`A) = (EX x:A. z = f x)"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1010
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1011
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1012
lemma image_subset_iff: "(f`A \<subseteq> B) = (\<forall>x\<in>A. f x \<in> B)"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1013
  -- {* This rewrite rule would confuse users if made default. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1014
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1015
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1016
lemma subset_image_iff: "(B \<subseteq> f`A) = (EX AA. AA \<subseteq> A & B = f`AA)"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1017
  apply safe
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1018
   prefer 2 apply fast
14208
144f45277d5a misc tidying
paulson
parents: 14098
diff changeset
  1019
  apply (rule_tac x = "{a. a : A & f a : B}" in exI, fast)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1020
  done
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1021
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1022
lemma image_subsetI: "(!!x. x \<in> A ==> f x \<in> B) ==> f`A \<subseteq> B"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1023
  -- {* Replaces the three steps @{text subsetI}, @{text imageE},
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1024
    @{text hypsubst}, but breaks too many existing proofs. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1025
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1026
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1027
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1028
  \medskip Range of a function -- just a translation for image!
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1029
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1030
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1031
lemma range_eqI: "b = f x ==> b \<in> range f"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1032
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1033
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1034
lemma rangeI: "f x \<in> range f"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1035
  by simp
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1036
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1037
lemma rangeE [elim?]: "b \<in> range (\<lambda>x. f x) ==> (!!x. b = f x ==> P) ==> P"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1038
  by blast
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1039
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1040
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1041
subsubsection {* Set reasoning tools *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1042
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1043
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1044
  Rewrite rules for boolean case-splitting: faster than @{text
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1045
  "split_if [split]"}.
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1046
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1047
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1048
lemma split_if_eq1: "((if Q then x else y) = b) = ((Q --> x = b) & (~ Q --> y = b))"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1049
  by (rule split_if)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1050
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1051
lemma split_if_eq2: "(a = (if Q then x else y)) = ((Q --> a = x) & (~ Q --> a = y))"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1052
  by (rule split_if)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1053
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1054
text {*
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1055
  Split ifs on either side of the membership relation.  Not for @{text
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1056
  "[simp]"} -- can cause goals to blow up!
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1057
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1058
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1059
lemma split_if_mem1: "((if Q then x else y) : b) = ((Q --> x : b) & (~ Q --> y : b))"
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1060
  by (rule split_if)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1061
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1062
lemma split_if_mem2: "(a : (if Q then x else y)) = ((Q --> a : x) & (~ Q --> a : y))"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  1063
  by (rule split_if [where P="%S. a : S"])
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1064
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1065
lemmas split_ifs = if_bool_eq_conj split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1066
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1067
lemmas mem_simps =
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1068
  insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1069
  mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1070
  -- {* Each of these has ALREADY been added @{text "[simp]"} above. *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1071
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1072
(*Would like to add these, but the existing code only searches for the
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1073
  outer-level constant, which in this case is just "op :"; we instead need
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1074
  to use term-nets to associate patterns with rules.  Also, if a rule fails to
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1075
  apply, then the formula should be kept.
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19175
diff changeset
  1076
  [("HOL.uminus", Compl_iff RS iffD1), ("HOL.minus", [Diff_iff RS iffD1]),
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  1077
   ("Int", [IntD1,IntD2]),
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1078
   ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])]
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1079
 *)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1080
26339
7825c83c9eff eliminated change_claset/simpset;
wenzelm
parents: 26150
diff changeset
  1081
ML {*
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  1082
  val mksimps_pairs = [(@{const_name Ball}, @{thms bspec})] @ mksimps_pairs;
26339
7825c83c9eff eliminated change_claset/simpset;
wenzelm
parents: 26150
diff changeset
  1083
*}
7825c83c9eff eliminated change_claset/simpset;
wenzelm
parents: 26150
diff changeset
  1084
declaration {* fn _ =>
7825c83c9eff eliminated change_claset/simpset;
wenzelm
parents: 26150
diff changeset
  1085
  Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs))
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1086
*}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1087
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1088
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1089
subsubsection {* The ``proper subset'' relation *}
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1090
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1091
lemma psubsetI [intro!,noatp]: "A \<subseteq> B ==> A \<noteq> B ==> A \<subset> B"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  1092
  by (unfold less_le) blast
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1093
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1094
lemma psubsetE [elim!,noatp]: 
13624
17684cf64fda added the new elim rule psubsetE
paulson
parents: 13550
diff changeset
  1095
    "[|A \<subset> B;  [|A \<subseteq> B; ~ (B\<subseteq>A)|] ==> R|] ==> R"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  1096
  by (unfold less_le) blast
13624
17684cf64fda added the new elim rule psubsetE
paulson
parents: 13550
diff changeset
  1097
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1098
lemma psubset_insert_iff:
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1099
  "(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)"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  1100
  by (auto simp add: less_le subset_insert_iff)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1101
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1102
lemma psubset_eq: "(A \<subset> B) = (A \<subseteq> B & A \<noteq> B)"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  1103
  by (simp only: less_le)
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1104
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1105
lemma psubset_imp_subset: "A \<subset> B ==> A \<subseteq> B"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1106
  by (simp add: psubset_eq)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1107
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14302
diff changeset
  1108
lemma psubset_trans: "[| A \<subset> B; B \<subset> C |] ==> A \<subset> C"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  1109
apply (unfold less_le)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14302
diff changeset
  1110
apply (auto dest: subset_antisym)
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14302
diff changeset
  1111
done
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14302
diff changeset
  1112
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14302
diff changeset
  1113
lemma psubsetD: "[| A \<subset> B; c \<in> A |] ==> c \<in> B"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  1114
apply (unfold less_le)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14302
diff changeset
  1115
apply (auto dest: subsetD)
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14302
diff changeset
  1116
done
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14302
diff changeset
  1117
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1118
lemma psubset_subset_trans: "A \<subset> B ==> B \<subseteq> C ==> A \<subset> C"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1119
  by (auto simp add: psubset_eq)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1120
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1121
lemma subset_psubset_trans: "A \<subseteq> B ==> B \<subset> C ==> A \<subset> C"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1122
  by (auto simp add: psubset_eq)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1123
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1124
lemma psubset_imp_ex_mem: "A \<subset> B ==> \<exists>b. b \<in> (B - A)"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  1125
  by (unfold less_le) blast
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1126
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1127
lemma atomize_ball:
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1128
    "(!!x. x \<in> A ==> P x) == Trueprop (\<forall>x\<in>A. P x)"
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1129
  by (simp only: Ball_def atomize_all atomize_imp)
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1130
18832
6ab4de872a70 declare 'defn' rules;
wenzelm
parents: 18674
diff changeset
  1131
lemmas [symmetric, rulify] = atomize_ball
6ab4de872a70 declare 'defn' rules;
wenzelm
parents: 18674
diff changeset
  1132
  and [symmetric, defn] = atomize_ball
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1133
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  1134
22455
f6f22aba2e0e added instance of sets as distributive lattices
haftmann
parents: 22439
diff changeset
  1135
subsection {* Further set-theory lemmas *}
f6f22aba2e0e added instance of sets as distributive lattices
haftmann
parents: 22439
diff changeset
  1136
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1137
subsubsection {* Derived rules involving subsets. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1138
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1139
text {* @{text insert}. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1140
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1141
lemma subset_insertI: "B \<subseteq> insert a B"
23878
bd651ecd4b8a simplified HOL bootstrap
haftmann
parents: 23857
diff changeset
  1142
  by (rule subsetI) (erule insertI2)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1143
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1144
lemma subset_insertI2: "A \<subseteq> B \<Longrightarrow> A \<subseteq> insert b B"
23878
bd651ecd4b8a simplified HOL bootstrap
haftmann
parents: 23857
diff changeset
  1145
  by blast
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1146
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1147
lemma subset_insert: "x \<notin> A ==> (A \<subseteq> insert x B) = (A \<subseteq> B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1148
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1149
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1150
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1151
text {* \medskip Big Union -- least upper bound of a set. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1152
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1153
lemma Union_upper: "B \<in> A ==> B \<subseteq> Union A"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  1154
  by (iprover intro: subsetI UnionI)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1155
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1156
lemma Union_least: "(!!X. X \<in> A ==> X \<subseteq> C) ==> Union A \<subseteq> C"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  1157
  by (iprover intro: subsetI elim: UnionE dest: subsetD)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1158
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1159
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1160
text {* \medskip General union. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1161
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1162
lemma UN_upper: "a \<in> A ==> B a \<subseteq> (\<Union>x\<in>A. B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1163
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1164
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1165
lemma UN_least: "(!!x. x \<in> A ==> B x \<subseteq> C) ==> (\<Union>x\<in>A. B x) \<subseteq> C"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  1166
  by (iprover intro: subsetI elim: UN_E dest: subsetD)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1167
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1168
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1169
text {* \medskip Big Intersection -- greatest lower bound of a set. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1170
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1171
lemma Inter_lower: "B \<in> A ==> Inter A \<subseteq> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1172
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1173
14551
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents: 14479
diff changeset
  1174
lemma Inter_subset:
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents: 14479
diff changeset
  1175
  "[| !!X. X \<in> A ==> X \<subseteq> B; A ~= {} |] ==> \<Inter>A \<subseteq> B"
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents: 14479
diff changeset
  1176
  by blast
2cb6ff394bfb Various changes to HOL-Algebra;
ballarin
parents: 14479
diff changeset
  1177
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1178
lemma Inter_greatest: "(!!X. X \<in> A ==> C \<subseteq> X) ==> C \<subseteq> Inter A"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  1179
  by (iprover intro: InterI subsetI dest: subsetD)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1180
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1181
lemma INT_lower: "a \<in> A ==> (\<Inter>x\<in>A. B x) \<subseteq> B a"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1182
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1183
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1184
lemma INT_greatest: "(!!x. x \<in> A ==> C \<subseteq> B x) ==> C \<subseteq> (\<Inter>x\<in>A. B x)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  1185
  by (iprover intro: INT_I subsetI dest: subsetD)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1186
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1187
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1188
text {* \medskip Finite Union -- the least upper bound of two sets. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1189
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1190
lemma Un_upper1: "A \<subseteq> A \<union> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1191
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1192
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1193
lemma Un_upper2: "B \<subseteq> A \<union> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1194
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1195
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1196
lemma Un_least: "A \<subseteq> C ==> B \<subseteq> C ==> A \<union> B \<subseteq> C"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1197
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1198
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1199
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1200
text {* \medskip Finite Intersection -- the greatest lower bound of two sets. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1201
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1202
lemma Int_lower1: "A \<inter> B \<subseteq> A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1203
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1204
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1205
lemma Int_lower2: "A \<inter> B \<subseteq> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1206
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1207
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1208
lemma Int_greatest: "C \<subseteq> A ==> C \<subseteq> B ==> C \<subseteq> A \<inter> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1209
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1210
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1211
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1212
text {* \medskip Set difference. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1213
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1214
lemma Diff_subset: "A - B \<subseteq> A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1215
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1216
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1217
lemma Diff_subset_conv: "(A - B \<subseteq> C) = (A \<subseteq> B \<union> C)"
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1218
by blast
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1219
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1220
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1221
subsubsection {* Equalities involving union, intersection, inclusion, etc. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1222
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1223
text {* @{text "{}"}. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1224
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1225
lemma Collect_const [simp]: "{s. P} = (if P then UNIV else {})"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1226
  -- {* supersedes @{text "Collect_False_empty"} *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1227
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1228
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1229
lemma subset_empty [simp]: "(A \<subseteq> {}) = (A = {})"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1230
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1231
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1232
lemma not_psubset_empty [iff]: "\<not> (A < {})"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  1233
  by (unfold less_le) blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1234
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1235
lemma Collect_empty_eq [simp]: "(Collect P = {}) = (\<forall>x. \<not> P x)"
18423
d7859164447f new lemmas
nipkow
parents: 18413
diff changeset
  1236
by blast
d7859164447f new lemmas
nipkow
parents: 18413
diff changeset
  1237
d7859164447f new lemmas
nipkow
parents: 18413
diff changeset
  1238
lemma empty_Collect_eq [simp]: "({} = Collect P) = (\<forall>x. \<not> P x)"
d7859164447f new lemmas
nipkow
parents: 18413
diff changeset
  1239
by blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1240
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1241
lemma Collect_neg_eq: "{x. \<not> P x} = - {x. P x}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1242
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1243
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1244
lemma Collect_disj_eq: "{x. P x | Q x} = {x. P x} \<union> {x. Q x}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1245
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1246
14812
4b2c006d0534 new theorem Collect_imp_eq
paulson
parents: 14804
diff changeset
  1247
lemma Collect_imp_eq: "{x. P x --> Q x} = -{x. P x} \<union> {x. Q x}"
4b2c006d0534 new theorem Collect_imp_eq
paulson
parents: 14804
diff changeset
  1248
  by blast
4b2c006d0534 new theorem Collect_imp_eq
paulson
parents: 14804
diff changeset
  1249
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1250
lemma Collect_conj_eq: "{x. P x & Q x} = {x. P x} \<inter> {x. Q x}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1251
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1252
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1253
lemma Collect_all_eq: "{x. \<forall>y. P x y} = (\<Inter>y. {x. P x y})"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1254
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1255
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1256
lemma Collect_ball_eq: "{x. \<forall>y\<in>A. P x y} = (\<Inter>y\<in>A. {x. P x y})"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1257
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1258
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1259
lemma Collect_ex_eq [noatp]: "{x. \<exists>y. P x y} = (\<Union>y. {x. P x y})"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1260
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1261
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1262
lemma Collect_bex_eq [noatp]: "{x. \<exists>y\<in>A. P x y} = (\<Union>y\<in>A. {x. P x y})"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1263
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1264
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1265
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1266
text {* \medskip @{text insert}. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1267
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1268
lemma insert_is_Un: "insert a A = {a} Un A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1269
  -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a {}"} *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1270
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1271
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1272
lemma insert_not_empty [simp]: "insert a A \<noteq> {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1273
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1274
17715
68583762ebdb a name for empty_not_insert
paulson
parents: 17702
diff changeset
  1275
lemmas empty_not_insert = insert_not_empty [symmetric, standard]
68583762ebdb a name for empty_not_insert
paulson
parents: 17702
diff changeset
  1276
declare empty_not_insert [simp]
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1277
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1278
lemma insert_absorb: "a \<in> A ==> insert a A = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1279
  -- {* @{text "[simp]"} causes recursive calls when there are nested inserts *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1280
  -- {* with \emph{quadratic} running time *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1281
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1282
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1283
lemma insert_absorb2 [simp]: "insert x (insert x A) = insert x A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1284
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1285
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1286
lemma insert_commute: "insert x (insert y A) = insert y (insert x A)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1287
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1288
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1289
lemma insert_subset [simp]: "(insert x A \<subseteq> B) = (x \<in> B & A \<subseteq> B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1290
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1291
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1292
lemma mk_disjoint_insert: "a \<in> A ==> \<exists>B. A = insert a B & a \<notin> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1293
  -- {* use new @{text B} rather than @{text "A - {a}"} to avoid infinite unfolding *}
14208
144f45277d5a misc tidying
paulson
parents: 14098
diff changeset
  1294
  apply (rule_tac x = "A - {a}" in exI, blast)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1295
  done
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1296
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1297
lemma insert_Collect: "insert a (Collect P) = {u. u \<noteq> a --> P u}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1298
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1299
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1300
lemma UN_insert_distrib: "u \<in> A ==> (\<Union>x\<in>A. insert a (B x)) = insert a (\<Union>x\<in>A. B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1301
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1302
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1303
lemma insert_inter_insert[simp]: "insert a A \<inter> insert a B = insert a (A \<inter> B)"
14742
dde816115d6a New simp rules added:
mehta
parents: 14692
diff changeset
  1304
  by blast
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1305
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1306
lemma insert_disjoint [simp,noatp]:
13103
66659a4b16f6 Added insert_disjoint and disjoint_insert [simp], and simplified proofs
nipkow
parents: 12937
diff changeset
  1307
 "(insert a A \<inter> B = {}) = (a \<notin> B \<and> A \<inter> B = {})"
14742
dde816115d6a New simp rules added:
mehta
parents: 14692
diff changeset
  1308
 "({} = insert a A \<inter> B) = (a \<notin> B \<and> {} = A \<inter> B)"
16773
33c4d8fe6f78 tweaked
paulson
parents: 16636
diff changeset
  1309
  by auto
13103
66659a4b16f6 Added insert_disjoint and disjoint_insert [simp], and simplified proofs
nipkow
parents: 12937
diff changeset
  1310
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1311
lemma disjoint_insert [simp,noatp]:
13103
66659a4b16f6 Added insert_disjoint and disjoint_insert [simp], and simplified proofs
nipkow
parents: 12937
diff changeset
  1312
 "(B \<inter> insert a A = {}) = (a \<notin> B \<and> B \<inter> A = {})"
14742
dde816115d6a New simp rules added:
mehta
parents: 14692
diff changeset
  1313
 "({} = A \<inter> insert b B) = (b \<notin> A \<and> {} = A \<inter> B)"
16773
33c4d8fe6f78 tweaked
paulson
parents: 16636
diff changeset
  1314
  by auto
14742
dde816115d6a New simp rules added:
mehta
parents: 14692
diff changeset
  1315
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1316
text {* \medskip @{text image}. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1317
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1318
lemma image_empty [simp]: "f`{} = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1319
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1320
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1321
lemma image_insert [simp]: "f ` insert a B = insert (f a) (f`B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1322
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1323
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1324
lemma image_constant: "x \<in> A ==> (\<lambda>x. c) ` A = {c}"
16773
33c4d8fe6f78 tweaked
paulson
parents: 16636
diff changeset
  1325
  by auto
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1326
21316
4d913b8bccf1 image_constant_conv no longer [simp]
nipkow
parents: 21312
diff changeset
  1327
lemma image_constant_conv: "(%x. c) ` A = (if A = {} then {} else {c})"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21210
diff changeset
  1328
by auto
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21210
diff changeset
  1329
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1330
lemma image_image: "f ` (g ` A) = (\<lambda>x. f (g x)) ` A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1331
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1332
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1333
lemma insert_image [simp]: "x \<in> A ==> insert (f x) (f`A) = f`A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1334
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1335
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1336
lemma image_is_empty [iff]: "(f`A = {}) = (A = {})"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1337
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1338
16773
33c4d8fe6f78 tweaked
paulson
parents: 16636
diff changeset
  1339
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1340
lemma image_Collect [noatp]: "f ` {x. P x} = {f x | x. P x}"
16773
33c4d8fe6f78 tweaked
paulson
parents: 16636
diff changeset
  1341
  -- {* NOT suitable as a default simprule: the RHS isn't simpler than the LHS,
33c4d8fe6f78 tweaked
paulson
parents: 16636
diff changeset
  1342
      with its implicit quantifier and conjunction.  Also image enjoys better
33c4d8fe6f78 tweaked
paulson
parents: 16636
diff changeset
  1343
      equational properties than does the RHS. *}
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1344
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1345
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1346
lemma if_image_distrib [simp]:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1347
  "(\<lambda>x. if P x then f x else g x) ` S
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1348
    = (f ` (S \<inter> {x. P x})) \<union> (g ` (S \<inter> {x. \<not> P x}))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1349
  by (auto simp add: image_def)
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1350
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1351
lemma image_cong: "M = N ==> (!!x. x \<in> N ==> f x = g x) ==> f`M = g`N"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1352
  by (simp add: image_def)
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1353
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1354
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1355
text {* \medskip @{text range}. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1356
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1357
lemma full_SetCompr_eq [noatp]: "{u. \<exists>x. u = f x} = range f"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1358
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1359
27418
564117b58d73 remove simp attribute from range_composition
huffman
parents: 27106
diff changeset
  1360
lemma range_composition: "range (\<lambda>x. f (g x)) = f`range g"
14208
144f45277d5a misc tidying
paulson
parents: 14098
diff changeset
  1361
by (subst image_image, simp)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1362
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1363
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1364
text {* \medskip @{text Int} *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1365
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1366
lemma Int_absorb [simp]: "A \<inter> A = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1367
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1368
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1369
lemma Int_left_absorb: "A \<inter> (A \<inter> B) = A \<inter> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1370
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1371
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1372
lemma Int_commute: "A \<inter> B = B \<inter> A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1373
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1374
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1375
lemma Int_left_commute: "A \<inter> (B \<inter> C) = B \<inter> (A \<inter> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1376
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1377
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1378
lemma Int_assoc: "(A \<inter> B) \<inter> C = A \<inter> (B \<inter> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1379
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1380
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1381
lemmas Int_ac = Int_assoc Int_left_absorb Int_commute Int_left_commute
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1382
  -- {* Intersection is an AC-operator *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1383
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1384
lemma Int_absorb1: "B \<subseteq> A ==> A \<inter> B = B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1385
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1386
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1387
lemma Int_absorb2: "A \<subseteq> B ==> A \<inter> B = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1388
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1389
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1390
lemma Int_empty_left [simp]: "{} \<inter> B = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1391
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1392
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1393
lemma Int_empty_right [simp]: "A \<inter> {} = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1394
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1395
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1396
lemma disjoint_eq_subset_Compl: "(A \<inter> B = {}) = (A \<subseteq> -B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1397
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1398
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1399
lemma disjoint_iff_not_equal: "(A \<inter> B = {}) = (\<forall>x\<in>A. \<forall>y\<in>B. x \<noteq> y)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1400
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1401
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1402
lemma Int_UNIV_left [simp]: "UNIV \<inter> B = B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1403
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1404
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1405
lemma Int_UNIV_right [simp]: "A \<inter> UNIV = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1406
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1407
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1408
lemma Int_eq_Inter: "A \<inter> B = \<Inter>{A, B}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1409
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1410
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1411
lemma Int_Un_distrib: "A \<inter> (B \<union> C) = (A \<inter> B) \<union> (A \<inter> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1412
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1413
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1414
lemma Int_Un_distrib2: "(B \<union> C) \<inter> A = (B \<inter> A) \<union> (C \<inter> A)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1415
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1416
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1417
lemma Int_UNIV [simp,noatp]: "(A \<inter> B = UNIV) = (A = UNIV & B = UNIV)"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1418
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1419
15102
04b0e943fcc9 new simprules Int_subset_iff and Un_subset_iff
paulson
parents: 14981
diff changeset
  1420
lemma Int_subset_iff [simp]: "(C \<subseteq> A \<inter> B) = (C \<subseteq> A & C \<subseteq> B)"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1421
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1422
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1423
lemma Int_Collect: "(x \<in> A \<inter> {x. P x}) = (x \<in> A & P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1424
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1425
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1426
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1427
text {* \medskip @{text Un}. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1428
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1429
lemma Un_absorb [simp]: "A \<union> A = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1430
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1431
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1432
lemma Un_left_absorb: "A \<union> (A \<union> B) = A \<union> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1433
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1434
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1435
lemma Un_commute: "A \<union> B = B \<union> A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1436
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1437
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1438
lemma Un_left_commute: "A \<union> (B \<union> C) = B \<union> (A \<union> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1439
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1440
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1441
lemma Un_assoc: "(A \<union> B) \<union> C = A \<union> (B \<union> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1442
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1443
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1444
lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1445
  -- {* Union is an AC-operator *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1446
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1447
lemma Un_absorb1: "A \<subseteq> B ==> A \<union> B = B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1448
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1449
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1450
lemma Un_absorb2: "B \<subseteq> A ==> A \<union> B = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1451
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1452
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1453
lemma Un_empty_left [simp]: "{} \<union> B = B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1454
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1455
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1456
lemma Un_empty_right [simp]: "A \<union> {} = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1457
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1458
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1459
lemma Un_UNIV_left [simp]: "UNIV \<union> B = UNIV"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1460
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1461
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1462
lemma Un_UNIV_right [simp]: "A \<union> UNIV = UNIV"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1463
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1464
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1465
lemma Un_eq_Union: "A \<union> B = \<Union>{A, B}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1466
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1467
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1468
lemma Un_insert_left [simp]: "(insert a B) \<union> C = insert a (B \<union> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1469
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1470
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1471
lemma Un_insert_right [simp]: "A \<union> (insert a B) = insert a (A \<union> B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1472
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1473
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1474
lemma Int_insert_left:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1475
    "(insert a B) Int C = (if a \<in> C then insert a (B \<inter> C) else B \<inter> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1476
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1477
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1478
lemma Int_insert_right:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1479
    "A \<inter> (insert a B) = (if a \<in> A then insert a (A \<inter> B) else A \<inter> B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1480
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1481
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1482
lemma Un_Int_distrib: "A \<union> (B \<inter> C) = (A \<union> B) \<inter> (A \<union> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1483
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1484
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1485
lemma Un_Int_distrib2: "(B \<inter> C) \<union> A = (B \<union> A) \<inter> (C \<union> A)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1486
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1487
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1488
lemma Un_Int_crazy:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1489
    "(A \<inter> B) \<union> (B \<inter> C) \<union> (C \<inter> A) = (A \<union> B) \<inter> (B \<union> C) \<inter> (C \<union> A)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1490
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1491
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1492
lemma subset_Un_eq: "(A \<subseteq> B) = (A \<union> B = B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1493
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1494
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1495
lemma Un_empty [iff]: "(A \<union> B = {}) = (A = {} & B = {})"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1496
  by blast
15102
04b0e943fcc9 new simprules Int_subset_iff and Un_subset_iff
paulson
parents: 14981
diff changeset
  1497
04b0e943fcc9 new simprules Int_subset_iff and Un_subset_iff
paulson
parents: 14981
diff changeset
  1498
lemma Un_subset_iff [simp]: "(A \<union> B \<subseteq> C) = (A \<subseteq> C & B \<subseteq> C)"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1499
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1500
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1501
lemma Un_Diff_Int: "(A - B) \<union> (A \<inter> B) = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1502
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1503
22172
e7d6cb237b5e some new lemmas
paulson
parents: 22139
diff changeset
  1504
lemma Diff_Int2: "A \<inter> C - B \<inter> C = A \<inter> C - B"
e7d6cb237b5e some new lemmas
paulson
parents: 22139
diff changeset
  1505
  by blast
e7d6cb237b5e some new lemmas
paulson
parents: 22139
diff changeset
  1506
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1507
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1508
text {* \medskip Set complement *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1509
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1510
lemma Compl_disjoint [simp]: "A \<inter> -A = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1511
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1512
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1513
lemma Compl_disjoint2 [simp]: "-A \<inter> A = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1514
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1515
13818
274fda8cca4b new theorem Compl_partition2
paulson
parents: 13764
diff changeset
  1516
lemma Compl_partition: "A \<union> -A = UNIV"
274fda8cca4b new theorem Compl_partition2
paulson
parents: 13764
diff changeset
  1517
  by blast
274fda8cca4b new theorem Compl_partition2
paulson
parents: 13764
diff changeset
  1518
274fda8cca4b new theorem Compl_partition2
paulson
parents: 13764
diff changeset
  1519
lemma Compl_partition2: "-A \<union> A = UNIV"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1520
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1521
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1522
lemma double_complement [simp]: "- (-A) = (A::'a set)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1523
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1524
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1525
lemma Compl_Un [simp]: "-(A \<union> B) = (-A) \<inter> (-B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1526
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1527
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1528
lemma Compl_Int [simp]: "-(A \<inter> B) = (-A) \<union> (-B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1529
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1530
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1531
lemma Compl_UN [simp]: "-(\<Union>x\<in>A. B x) = (\<Inter>x\<in>A. -B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1532
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1533
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1534
lemma Compl_INT [simp]: "-(\<Inter>x\<in>A. B x) = (\<Union>x\<in>A. -B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1535
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1536
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1537
lemma subset_Compl_self_eq: "(A \<subseteq> -A) = (A = {})"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1538
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1539
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1540
lemma Un_Int_assoc_eq: "((A \<inter> B) \<union> C = A \<inter> (B \<union> C)) = (C \<subseteq> A)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1541
  -- {* Halmos, Naive Set Theory, page 16. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1542
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1543
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1544
lemma Compl_UNIV_eq [simp]: "-UNIV = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1545
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1546
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1547
lemma Compl_empty_eq [simp]: "-{} = UNIV"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1548
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1549
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1550
lemma Compl_subset_Compl_iff [iff]: "(-A \<subseteq> -B) = (B \<subseteq> A)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1551
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1552
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1553
lemma Compl_eq_Compl_iff [iff]: "(-A = -B) = (A = (B::'a set))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1554
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1555
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1556
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1557
text {* \medskip @{text Union}. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1558
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1559
lemma Union_empty [simp]: "Union({}) = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1560
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1561
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1562
lemma Union_UNIV [simp]: "Union UNIV = UNIV"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1563
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1564
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1565
lemma Union_insert [simp]: "Union (insert a B) = a \<union> \<Union>B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1566
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1567
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1568
lemma Union_Un_distrib [simp]: "\<Union>(A Un B) = \<Union>A \<union> \<Union>B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1569
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1570
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1571
lemma Union_Int_subset: "\<Union>(A \<inter> B) \<subseteq> \<Union>A \<inter> \<Union>B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1572
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1573
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1574
lemma Union_empty_conv [simp,noatp]: "(\<Union>A = {}) = (\<forall>x\<in>A. x = {})"
13653
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1575
  by blast
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1576
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1577
lemma empty_Union_conv [simp,noatp]: "({} = \<Union>A) = (\<forall>x\<in>A. x = {})"
13653
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1578
  by blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1579
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1580
lemma Union_disjoint: "(\<Union>C \<inter> A = {}) = (\<forall>B\<in>C. B \<inter> A = {})"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1581
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1582
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1583
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1584
text {* \medskip @{text Inter}. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1585
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1586
lemma Inter_empty [simp]: "\<Inter>{} = UNIV"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1587
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1588
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1589
lemma Inter_UNIV [simp]: "\<Inter>UNIV = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1590
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1591
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1592
lemma Inter_insert [simp]: "\<Inter>(insert a B) = a \<inter> \<Inter>B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1593
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1594
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1595
lemma Inter_Un_subset: "\<Inter>A \<union> \<Inter>B \<subseteq> \<Inter>(A \<inter> B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1596
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1597
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1598
lemma Inter_Un_distrib: "\<Inter>(A \<union> B) = \<Inter>A \<inter> \<Inter>B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1599
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1600
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1601
lemma Inter_UNIV_conv [simp,noatp]:
13653
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1602
  "(\<Inter>A = UNIV) = (\<forall>x\<in>A. x = UNIV)"
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1603
  "(UNIV = \<Inter>A) = (\<forall>x\<in>A. x = UNIV)"
14208
144f45277d5a misc tidying
paulson
parents: 14098
diff changeset
  1604
  by blast+
13653
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1605
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1606
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1607
text {*
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1608
  \medskip @{text UN} and @{text INT}.
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1609
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1610
  Basic identities: *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1611
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1612
lemma UN_empty [simp,noatp]: "(\<Union>x\<in>{}. B x) = {}"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1613
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1614
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1615
lemma UN_empty2 [simp]: "(\<Union>x\<in>A. {}) = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1616
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1617
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1618
lemma UN_singleton [simp]: "(\<Union>x\<in>A. {x}) = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1619
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1620
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1621
lemma UN_absorb: "k \<in> I ==> A k \<union> (\<Union>i\<in>I. A i) = (\<Union>i\<in>I. A i)"
15102
04b0e943fcc9 new simprules Int_subset_iff and Un_subset_iff
paulson
parents: 14981
diff changeset
  1622
  by auto
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1623
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1624
lemma INT_empty [simp]: "(\<Inter>x\<in>{}. B x) = UNIV"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1625
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1626
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1627
lemma INT_absorb: "k \<in> I ==> A k \<inter> (\<Inter>i\<in>I. A i) = (\<Inter>i\<in>I. A i)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1628
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1629
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1630
lemma UN_insert [simp]: "(\<Union>x\<in>insert a A. B x) = B a \<union> UNION A B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1631
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1632
24331
76f7a8c6e842 Made UN_Un simp
nipkow
parents: 24303
diff changeset
  1633
lemma UN_Un[simp]: "(\<Union>i \<in> A \<union> B. M i) = (\<Union>i\<in>A. M i) \<union> (\<Union>i\<in>B. M i)"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1634
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1635
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1636
lemma UN_UN_flatten: "(\<Union>x \<in> (\<Union>y\<in>A. B y). C x) = (\<Union>y\<in>A. \<Union>x\<in>B y. C x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1637
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1638
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1639
lemma UN_subset_iff: "((\<Union>i\<in>I. A i) \<subseteq> B) = (\<forall>i\<in>I. A i \<subseteq> B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1640
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1641
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1642
lemma INT_subset_iff: "(B \<subseteq> (\<Inter>i\<in>I. A i)) = (\<forall>i\<in>I. B \<subseteq> A i)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1643
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1644
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1645
lemma INT_insert [simp]: "(\<Inter>x \<in> insert a A. B x) = B a \<inter> INTER A B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1646
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1647
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1648
lemma INT_Un: "(\<Inter>i \<in> A \<union> B. M i) = (\<Inter>i \<in> A. M i) \<inter> (\<Inter>i\<in>B. M i)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1649
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1650
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1651
lemma INT_insert_distrib:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1652
    "u \<in> A ==> (\<Inter>x\<in>A. insert a (B x)) = insert a (\<Inter>x\<in>A. B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1653
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1654
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1655
lemma Union_image_eq [simp]: "\<Union>(B`A) = (\<Union>x\<in>A. B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1656
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1657
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1658
lemma image_Union: "f ` \<Union>S = (\<Union>x\<in>S. f ` x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1659
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1660
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1661
lemma Inter_image_eq [simp]: "\<Inter>(B`A) = (\<Inter>x\<in>A. B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1662
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1663
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1664
lemma UN_constant [simp]: "(\<Union>y\<in>A. c) = (if A = {} then {} else c)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1665
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1666
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1667
lemma INT_constant [simp]: "(\<Inter>y\<in>A. c) = (if A = {} then UNIV else c)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1668
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1669
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1670
lemma UN_eq: "(\<Union>x\<in>A. B x) = \<Union>({Y. \<exists>x\<in>A. Y = B x})"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1671
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1672
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1673
lemma INT_eq: "(\<Inter>x\<in>A. B x) = \<Inter>({Y. \<exists>x\<in>A. Y = B x})"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1674
  -- {* Look: it has an \emph{existential} quantifier *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1675
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1676
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
  1677
lemma UNION_empty_conv[simp]:
13653
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1678
  "({} = (UN x:A. B x)) = (\<forall>x\<in>A. B x = {})"
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1679
  "((UN x:A. B x) = {}) = (\<forall>x\<in>A. B x = {})"
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1680
by blast+
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1681
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
  1682
lemma INTER_UNIV_conv[simp]:
13653
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1683
 "(UNIV = (INT x:A. B x)) = (\<forall>x\<in>A. B x = UNIV)"
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1684
 "((INT x:A. B x) = UNIV) = (\<forall>x\<in>A. B x = UNIV)"
ef123b9e8089 Added a few thms about UN/INT/{}/UNIV
nipkow
parents: 13624
diff changeset
  1685
by blast+
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1686
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1687
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1688
text {* \medskip Distributive laws: *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1689
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1690
lemma Int_Union: "A \<inter> \<Union>B = (\<Union>C\<in>B. A \<inter> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1691
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1692
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1693
lemma Int_Union2: "\<Union>B \<inter> A = (\<Union>C\<in>B. C \<inter> A)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1694
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1695
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1696
lemma Un_Union_image: "(\<Union>x\<in>C. A x \<union> B x) = \<Union>(A`C) \<union> \<Union>(B`C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1697
  -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1698
  -- {* Union of a family of unions *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1699
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1700
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1701
lemma UN_Un_distrib: "(\<Union>i\<in>I. A i \<union> B i) = (\<Union>i\<in>I. A i) \<union> (\<Union>i\<in>I. B i)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1702
  -- {* Equivalent version *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1703
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1704
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1705
lemma Un_Inter: "A \<union> \<Inter>B = (\<Inter>C\<in>B. A \<union> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1706
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1707
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1708
lemma Int_Inter_image: "(\<Inter>x\<in>C. A x \<inter> B x) = \<Inter>(A`C) \<inter> \<Inter>(B`C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1709
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1710
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1711
lemma INT_Int_distrib: "(\<Inter>i\<in>I. A i \<inter> B i) = (\<Inter>i\<in>I. A i) \<inter> (\<Inter>i\<in>I. B i)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1712
  -- {* Equivalent version *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1713
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1714
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1715
lemma Int_UN_distrib: "B \<inter> (\<Union>i\<in>I. A i) = (\<Union>i\<in>I. B \<inter> A i)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1716
  -- {* Halmos, Naive Set Theory, page 35. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1717
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1718
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1719
lemma Un_INT_distrib: "B \<union> (\<Inter>i\<in>I. A i) = (\<Inter>i\<in>I. B \<union> A i)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1720
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1721
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1722
lemma Int_UN_distrib2: "(\<Union>i\<in>I. A i) \<inter> (\<Union>j\<in>J. B j) = (\<Union>i\<in>I. \<Union>j\<in>J. A i \<inter> B j)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1723
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1724
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1725
lemma Un_INT_distrib2: "(\<Inter>i\<in>I. A i) \<union> (\<Inter>j\<in>J. B j) = (\<Inter>i\<in>I. \<Inter>j\<in>J. A i \<union> B j)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1726
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1727
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1728
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1729
text {* \medskip Bounded quantifiers.
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1730
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1731
  The following are not added to the default simpset because
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1732
  (a) they duplicate the body and (b) there are no similar rules for @{text Int}. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1733
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1734
lemma ball_Un: "(\<forall>x \<in> A \<union> B. P x) = ((\<forall>x\<in>A. P x) & (\<forall>x\<in>B. P x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1735
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1736
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1737
lemma bex_Un: "(\<exists>x \<in> A \<union> B. P x) = ((\<exists>x\<in>A. P x) | (\<exists>x\<in>B. P x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1738
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1739
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1740
lemma ball_UN: "(\<forall>z \<in> UNION A B. P z) = (\<forall>x\<in>A. \<forall>z \<in> B x. P z)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1741
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1742
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1743
lemma bex_UN: "(\<exists>z \<in> UNION A B. P z) = (\<exists>x\<in>A. \<exists>z\<in>B x. P z)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1744
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1745
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1746
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1747
text {* \medskip Set difference. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1748
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1749
lemma Diff_eq: "A - B = A \<inter> (-B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1750
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1751
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1752
lemma Diff_eq_empty_iff [simp,noatp]: "(A - B = {}) = (A \<subseteq> B)"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1753
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1754
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1755
lemma Diff_cancel [simp]: "A - A = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1756
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1757
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1758
lemma Diff_idemp [simp]: "(A - B) - B = A - (B::'a set)"
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1759
by blast
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1760
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1761
lemma Diff_triv: "A \<inter> B = {} ==> A - B = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1762
  by (blast elim: equalityE)
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1763
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1764
lemma empty_Diff [simp]: "{} - A = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1765
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1766
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1767
lemma Diff_empty [simp]: "A - {} = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1768
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1769
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1770
lemma Diff_UNIV [simp]: "A - UNIV = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1771
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1772
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1773
lemma Diff_insert0 [simp,noatp]: "x \<notin> A ==> A - insert x B = A - B"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1774
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1775
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1776
lemma Diff_insert: "A - insert a B = A - B - {a}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1777
  -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"} *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1778
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1779
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1780
lemma Diff_insert2: "A - insert a B = A - {a} - B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1781
  -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"} *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1782
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1783
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1784
lemma insert_Diff_if: "insert x A - B = (if x \<in> B then A - B else insert x (A - B))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1785
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1786
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1787
lemma insert_Diff1 [simp]: "x \<in> B ==> insert x A - B = A - B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1788
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1789
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1790
lemma insert_Diff_single[simp]: "insert a (A - {a}) = insert a A"
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1791
by blast
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
  1792
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1793
lemma insert_Diff: "a \<in> A ==> insert a (A - {a}) = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1794
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1795
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1796
lemma Diff_insert_absorb: "x \<notin> A ==> (insert x A) - {x} = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1797
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1798
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1799
lemma Diff_disjoint [simp]: "A \<inter> (B - A) = {}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1800
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1801
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1802
lemma Diff_partition: "A \<subseteq> B ==> A \<union> (B - A) = B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1803
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1804
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1805
lemma double_diff: "A \<subseteq> B ==> B \<subseteq> C ==> B - (C - A) = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1806
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1807
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1808
lemma Un_Diff_cancel [simp]: "A \<union> (B - A) = A \<union> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1809
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1810
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1811
lemma Un_Diff_cancel2 [simp]: "(B - A) \<union> A = B \<union> A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1812
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1813
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1814
lemma Diff_Un: "A - (B \<union> C) = (A - B) \<inter> (A - C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1815
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1816
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1817
lemma Diff_Int: "A - (B \<inter> C) = (A - B) \<union> (A - C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1818
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1819
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1820
lemma Un_Diff: "(A \<union> B) - C = (A - C) \<union> (B - C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1821
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1822
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1823
lemma Int_Diff: "(A \<inter> B) - C = A \<inter> (B - C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1824
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1825
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1826
lemma Diff_Int_distrib: "C \<inter> (A - B) = (C \<inter> A) - (C \<inter> B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1827
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1828
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1829
lemma Diff_Int_distrib2: "(A - B) \<inter> C = (A \<inter> C) - (B \<inter> C)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1830
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1831
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1832
lemma Diff_Compl [simp]: "A - (- B) = A \<inter> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1833
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1834
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1835
lemma Compl_Diff_eq [simp]: "- (A - B) = -A \<union> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1836
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1837
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1838
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1839
text {* \medskip Quantification over type @{typ bool}. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1840
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1841
lemma bool_induct: "P True \<Longrightarrow> P False \<Longrightarrow> P x"
21549
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1842
  by (cases x) auto
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1843
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1844
lemma all_bool_eq: "(\<forall>b. P b) \<longleftrightarrow> P True \<and> P False"
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1845
  by (auto intro: bool_induct)
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1846
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1847
lemma bool_contrapos: "P x \<Longrightarrow> \<not> P False \<Longrightarrow> P True"
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1848
  by (cases x) auto
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1849
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1850
lemma ex_bool_eq: "(\<exists>b. P b) \<longleftrightarrow> P True \<or> P False"
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1851
  by (auto intro: bool_contrapos)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1852
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1853
lemma Un_eq_UN: "A \<union> B = (\<Union>b. if b then A else B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1854
  by (auto simp add: split_if_mem2)
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1855
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1856
lemma UN_bool_eq: "(\<Union>b::bool. A b) = (A True \<union> A False)"
21549
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1857
  by (auto intro: bool_contrapos)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1858
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1859
lemma INT_bool_eq: "(\<Inter>b::bool. A b) = (A True \<inter> A False)"
21549
12eff58b56a0 restructured some proofs
haftmann
parents: 21524
diff changeset
  1860
  by (auto intro: bool_induct)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1861
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1862
text {* \medskip @{text Pow} *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1863
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1864
lemma Pow_empty [simp]: "Pow {} = {{}}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1865
  by (auto simp add: Pow_def)
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1866
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1867
lemma Pow_insert: "Pow (insert a A) = Pow A \<union> (insert a ` Pow A)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1868
  by (blast intro: image_eqI [where ?x = "u - {a}", standard])
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1869
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1870
lemma Pow_Compl: "Pow (- A) = {-B | B. A \<in> Pow B}"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1871
  by (blast intro: exI [where ?x = "- u", standard])
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1872
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1873
lemma Pow_UNIV [simp]: "Pow UNIV = UNIV"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1874
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1875
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1876
lemma Un_Pow_subset: "Pow A \<union> Pow B \<subseteq> Pow (A \<union> B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1877
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1878
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1879
lemma UN_Pow_subset: "(\<Union>x\<in>A. Pow (B x)) \<subseteq> Pow (\<Union>x\<in>A. B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1880
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1881
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1882
lemma subset_Pow_Union: "A \<subseteq> Pow (\<Union>A)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1883
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1884
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1885
lemma Union_Pow_eq [simp]: "\<Union>(Pow A) = A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1886
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1887
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1888
lemma Pow_Int_eq [simp]: "Pow (A \<inter> B) = Pow A \<inter> Pow B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1889
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1890
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1891
lemma Pow_INT_eq: "Pow (\<Inter>x\<in>A. B x) = (\<Inter>x\<in>A. Pow (B x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1892
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1893
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1894
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1895
text {* \medskip Miscellany. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1896
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1897
lemma set_eq_subset: "(A = B) = (A \<subseteq> B & B \<subseteq> A)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1898
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1899
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1900
lemma subset_iff: "(A \<subseteq> B) = (\<forall>t. t \<in> A --> t \<in> B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1901
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1902
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1903
lemma subset_iff_psubset_eq: "(A \<subseteq> B) = ((A \<subset> B) | (A = B))"
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  1904
  by (unfold less_le) blast
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1905
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 18423
diff changeset
  1906
lemma all_not_in_conv [simp]: "(\<forall>x. x \<notin> A) = (A = {})"
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1907
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1908
13831
ab27b36aba99 new lemma
paulson
parents: 13826
diff changeset
  1909
lemma ex_in_conv: "(\<exists>x. x \<in> A) = (A \<noteq> {})"
ab27b36aba99 new lemma
paulson
parents: 13826
diff changeset
  1910
  by blast
ab27b36aba99 new lemma
paulson
parents: 13826
diff changeset
  1911
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1912
lemma distinct_lemma: "f x \<noteq> f y ==> x \<noteq> y"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  1913
  by iprover
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1914
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1915
13860
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1916
text {* \medskip Miniscoping: pushing in quantifiers and big Unions
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1917
           and Intersections. *}
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1918
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1919
lemma UN_simps [simp]:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1920
  "!!a B C. (UN x:C. insert a (B x)) = (if C={} then {} else insert a (UN x:C. B x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1921
  "!!A B C. (UN x:C. A x Un B)   = ((if C={} then {} else (UN x:C. A x) Un B))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1922
  "!!A B C. (UN x:C. A Un B x)   = ((if C={} then {} else A Un (UN x:C. B x)))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1923
  "!!A B C. (UN x:C. A x Int B)  = ((UN x:C. A x) Int B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1924
  "!!A B C. (UN x:C. A Int B x)  = (A Int (UN x:C. B x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1925
  "!!A B C. (UN x:C. A x - B)    = ((UN x:C. A x) - B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1926
  "!!A B C. (UN x:C. A - B x)    = (A - (INT x:C. B x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1927
  "!!A B. (UN x: Union A. B x) = (UN y:A. UN x:y. B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1928
  "!!A B C. (UN z: UNION A B. C z) = (UN  x:A. UN z: B(x). C z)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1929
  "!!A B f. (UN x:f`A. B x)     = (UN a:A. B (f a))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1930
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1931
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1932
lemma INT_simps [simp]:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1933
  "!!A B C. (INT x:C. A x Int B) = (if C={} then UNIV else (INT x:C. A x) Int B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1934
  "!!A B C. (INT x:C. A Int B x) = (if C={} then UNIV else A Int (INT x:C. B x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1935
  "!!A B C. (INT x:C. A x - B)   = (if C={} then UNIV else (INT x:C. A x) - B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1936
  "!!A B C. (INT x:C. A - B x)   = (if C={} then UNIV else A - (UN x:C. B x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1937
  "!!a B C. (INT x:C. insert a (B x)) = insert a (INT x:C. B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1938
  "!!A B C. (INT x:C. A x Un B)  = ((INT x:C. A x) Un B)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1939
  "!!A B C. (INT x:C. A Un B x)  = (A Un (INT x:C. B x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1940
  "!!A B. (INT x: Union A. B x) = (INT y:A. INT x:y. B x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1941
  "!!A B C. (INT z: UNION A B. C z) = (INT x:A. INT z: B(x). C z)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1942
  "!!A B f. (INT x:f`A. B x)    = (INT a:A. B (f a))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1943
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1944
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1945
lemma ball_simps [simp,noatp]:
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1946
  "!!A P Q. (ALL x:A. P x | Q) = ((ALL x:A. P x) | Q)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1947
  "!!A P Q. (ALL x:A. P | Q x) = (P | (ALL x:A. Q x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1948
  "!!A P Q. (ALL x:A. P --> Q x) = (P --> (ALL x:A. Q x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1949
  "!!A P Q. (ALL x:A. P x --> Q) = ((EX x:A. P x) --> Q)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1950
  "!!P. (ALL x:{}. P x) = True"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1951
  "!!P. (ALL x:UNIV. P x) = (ALL x. P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1952
  "!!a B P. (ALL x:insert a B. P x) = (P a & (ALL x:B. P x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1953
  "!!A P. (ALL x:Union A. P x) = (ALL y:A. ALL x:y. P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1954
  "!!A B P. (ALL x: UNION A B. P x) = (ALL a:A. ALL x: B a. P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1955
  "!!P Q. (ALL x:Collect Q. P x) = (ALL x. Q x --> P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1956
  "!!A P f. (ALL x:f`A. P x) = (ALL x:A. P (f x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1957
  "!!A P. (~(ALL x:A. P x)) = (EX x:A. ~P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1958
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1959
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24280
diff changeset
  1960
lemma bex_simps [simp,noatp]:
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1961
  "!!A P Q. (EX x:A. P x & Q) = ((EX x:A. P x) & Q)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1962
  "!!A P Q. (EX x:A. P & Q x) = (P & (EX x:A. Q x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1963
  "!!P. (EX x:{}. P x) = False"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1964
  "!!P. (EX x:UNIV. P x) = (EX x. P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1965
  "!!a B P. (EX x:insert a B. P x) = (P(a) | (EX x:B. P x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1966
  "!!A P. (EX x:Union A. P x) = (EX y:A. EX x:y. P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1967
  "!!A B P. (EX x: UNION A B. P x) = (EX a:A. EX x:B a. P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1968
  "!!P Q. (EX x:Collect Q. P x) = (EX x. Q x & P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1969
  "!!A P f. (EX x:f`A. P x) = (EX x:A. P (f x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1970
  "!!A P. (~(EX x:A. P x)) = (ALL x:A. ~P x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1971
  by auto
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1972
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1973
lemma ball_conj_distrib:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1974
  "(ALL x:A. P x & Q x) = ((ALL x:A. P x) & (ALL x:A. Q x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1975
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1976
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1977
lemma bex_disj_distrib:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1978
  "(EX x:A. P x | Q x) = ((EX x:A. P x) | (EX x:A. Q x))"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1979
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1980
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  1981
13860
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1982
text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *}
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1983
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1984
lemma UN_extend_simps:
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1985
  "!!a B C. insert a (UN x:C. B x) = (if C={} then {a} else (UN x:C. insert a (B x)))"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1986
  "!!A B C. (UN x:C. A x) Un B    = (if C={} then B else (UN x:C. A x Un B))"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1987
  "!!A B C. A Un (UN x:C. B x)   = (if C={} then A else (UN x:C. A Un B x))"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1988
  "!!A B C. ((UN x:C. A x) Int B) = (UN x:C. A x Int B)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1989
  "!!A B C. (A Int (UN x:C. B x)) = (UN x:C. A Int B x)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1990
  "!!A B C. ((UN x:C. A x) - B) = (UN x:C. A x - B)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1991
  "!!A B C. (A - (INT x:C. B x)) = (UN x:C. A - B x)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1992
  "!!A B. (UN y:A. UN x:y. B x) = (UN x: Union A. B x)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1993
  "!!A B C. (UN  x:A. UN z: B(x). C z) = (UN z: UNION A B. C z)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1994
  "!!A B f. (UN a:A. B (f a)) = (UN x:f`A. B x)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1995
  by auto
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1996
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1997
lemma INT_extend_simps:
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1998
  "!!A B C. (INT x:C. A x) Int B = (if C={} then B else (INT x:C. A x Int B))"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  1999
  "!!A B C. A Int (INT x:C. B x) = (if C={} then A else (INT x:C. A Int B x))"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2000
  "!!A B C. (INT x:C. A x) - B   = (if C={} then UNIV-B else (INT x:C. A x - B))"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2001
  "!!A B C. A - (UN x:C. B x)   = (if C={} then A else (INT x:C. A - B x))"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2002
  "!!a B C. insert a (INT x:C. B x) = (INT x:C. insert a (B x))"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2003
  "!!A B C. ((INT x:C. A x) Un B)  = (INT x:C. A x Un B)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2004
  "!!A B C. A Un (INT x:C. B x)  = (INT x:C. A Un B x)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2005
  "!!A B. (INT y:A. INT x:y. B x) = (INT x: Union A. B x)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2006
  "!!A B C. (INT x:A. INT z: B(x). C z) = (INT z: UNION A B. C z)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2007
  "!!A B f. (INT a:A. B (f a))    = (INT x:f`A. B x)"
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2008
  by auto
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2009
b681a3cb0beb new UN/INT simprules
paulson
parents: 13858
diff changeset
  2010
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2011
subsubsection {* Monotonicity of various operations *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2012
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2013
lemma image_mono: "A \<subseteq> B ==> f`A \<subseteq> f`B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2014
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2015
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2016
lemma Pow_mono: "A \<subseteq> B ==> Pow A \<subseteq> Pow B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2017
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2018
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2019
lemma Union_mono: "A \<subseteq> B ==> \<Union>A \<subseteq> \<Union>B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2020
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2021
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2022
lemma Inter_anti_mono: "B \<subseteq> A ==> \<Inter>A \<subseteq> \<Inter>B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2023
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2024
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2025
lemma UN_mono:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2026
  "A \<subseteq> B ==> (!!x. x \<in> A ==> f x \<subseteq> g x) ==>
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2027
    (\<Union>x\<in>A. f x) \<subseteq> (\<Union>x\<in>B. g x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2028
  by (blast dest: subsetD)
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2029
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2030
lemma INT_anti_mono:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2031
  "B \<subseteq> A ==> (!!x. x \<in> A ==> f x \<subseteq> g x) ==>
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2032
    (\<Inter>x\<in>A. f x) \<subseteq> (\<Inter>x\<in>A. g x)"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2033
  -- {* The last inclusion is POSITIVE! *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2034
  by (blast dest: subsetD)
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2035
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2036
lemma insert_mono: "C \<subseteq> D ==> insert a C \<subseteq> insert a D"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2037
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2038
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2039
lemma Un_mono: "A \<subseteq> C ==> B \<subseteq> D ==> A \<union> B \<subseteq> C \<union> D"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2040
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2041
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2042
lemma Int_mono: "A \<subseteq> C ==> B \<subseteq> D ==> A \<inter> B \<subseteq> C \<inter> D"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2043
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2044
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2045
lemma Diff_mono: "A \<subseteq> C ==> D \<subseteq> B ==> A - B \<subseteq> C - D"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2046
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2047
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2048
lemma Compl_anti_mono: "A \<subseteq> B ==> -B \<subseteq> -A"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2049
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2050
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2051
text {* \medskip Monotonicity of implications. *}
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2052
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2053
lemma in_mono: "A \<subseteq> B ==> x \<in> A --> x \<in> B"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2054
  apply (rule impI)
14208
144f45277d5a misc tidying
paulson
parents: 14098
diff changeset
  2055
  apply (erule subsetD, assumption)
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2056
  done
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2057
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2058
lemma conj_mono: "P1 --> Q1 ==> P2 --> Q2 ==> (P1 & P2) --> (Q1 & Q2)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  2059
  by iprover
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2060
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2061
lemma disj_mono: "P1 --> Q1 ==> P2 --> Q2 ==> (P1 | P2) --> (Q1 | Q2)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  2062
  by iprover
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2063
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2064
lemma imp_mono: "Q1 --> P1 ==> P2 --> Q2 ==> (P1 --> P2) --> (Q1 --> Q2)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  2065
  by iprover
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2066
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2067
lemma imp_refl: "P --> P" ..
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2068
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2069
lemma ex_mono: "(!!x. P x --> Q x) ==> (EX x. P x) --> (EX x. Q x)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  2070
  by iprover
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2071
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2072
lemma all_mono: "(!!x. P x --> Q x) ==> (ALL x. P x) --> (ALL x. Q x)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  2073
  by iprover
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2074
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2075
lemma Collect_mono: "(!!x. P x --> Q x) ==> Collect P \<subseteq> Collect Q"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2076
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2077
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2078
lemma Int_Collect_mono:
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2079
    "A \<subseteq> B ==> (!!x. x \<in> A ==> P x --> Q x) ==> A \<inter> Collect P \<subseteq> B \<inter> Collect Q"
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2080
  by blast
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2081
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2082
lemmas basic_monos =
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2083
  subset_refl imp_refl disj_mono conj_mono
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2084
  ex_mono Collect_mono in_mono
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2085
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2086
lemma eq_to_mono: "a = b ==> c = d ==> b --> d ==> a --> c"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  2087
  by iprover
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2088
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2089
lemma eq_to_mono2: "a = b ==> c = d ==> ~ b --> ~ d ==> ~ a --> ~ c"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17508
diff changeset
  2090
  by iprover
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  2091
12020
a24373086908 theory Calculation move to Set;
wenzelm
parents: 11982
diff changeset
  2092
12257
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2093
subsection {* Inverse image of a function *}
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2094
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2095
constdefs
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2096
  vimage :: "('a => 'b) => 'b set => 'a set"    (infixr "-`" 90)
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27824
diff changeset
  2097
  [code del]: "f -` B == {x. f x : B}"
12257
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2098
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2099
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2100
subsubsection {* Basic rules *}
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2101
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2102
lemma vimage_eq [simp]: "(a : f -` B) = (f a : B)"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2103
  by (unfold vimage_def) blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2104
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2105
lemma vimage_singleton_eq: "(a : f -` {b}) = (f a = b)"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2106
  by simp
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2107
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2108
lemma vimageI [intro]: "f a = b ==> b:B ==> a : f -` B"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2109
  by (unfold vimage_def) blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2110
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2111
lemma vimageI2: "f a : A ==> a : f -` A"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2112
  by (unfold vimage_def) fast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2113
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2114
lemma vimageE [elim!]: "a: f -` B ==> (!!x. f a = x ==> x:B ==> P) ==> P"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2115
  by (unfold vimage_def) blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2116
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2117
lemma vimageD: "a : f -` A ==> f a : A"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2118
  by (unfold vimage_def) fast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2119
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2120
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2121
subsubsection {* Equations *}
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2122
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2123
lemma vimage_empty [simp]: "f -` {} = {}"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2124
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2125
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2126
lemma vimage_Compl: "f -` (-A) = -(f -` A)"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2127
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2128
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2129
lemma vimage_Un [simp]: "f -` (A Un B) = (f -` A) Un (f -` B)"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2130
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2131
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2132
lemma vimage_Int [simp]: "f -` (A Int B) = (f -` A) Int (f -` B)"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2133
  by fast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2134
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2135
lemma vimage_Union: "f -` (Union A) = (UN X:A. f -` X)"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2136
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2137
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2138
lemma vimage_UN: "f-`(UN x:A. B x) = (UN x:A. f -` B x)"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2139
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2140
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2141
lemma vimage_INT: "f-`(INT x:A. B x) = (INT x:A. f -` B x)"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2142
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2143
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2144
lemma vimage_Collect_eq [simp]: "f -` Collect P = {y. P (f y)}"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2145
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2146
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2147
lemma vimage_Collect: "(!!x. P (f x) = Q x) ==> f -` (Collect P) = Collect Q"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2148
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2149
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2150
lemma vimage_insert: "f-`(insert a B) = (f-`{a}) Un (f-`B)"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2151
  -- {* NOT suitable for rewriting because of the recurrence of @{term "{a}"}. *}
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2152
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2153
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2154
lemma vimage_Diff: "f -` (A - B) = (f -` A) - (f -` B)"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2155
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2156
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2157
lemma vimage_UNIV [simp]: "f -` UNIV = UNIV"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2158
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2159
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2160
lemma vimage_eq_UN: "f-`B = (UN y: B. f-`{y})"
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2161
  -- {* NOT suitable for rewriting *}
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2162
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2163
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2164
lemma vimage_mono: "A \<subseteq> B ==> f -` A \<subseteq> f -` B"
12257
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2165
  -- {* monotonicity *}
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2166
  by blast
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2167
26150
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2168
lemma vimage_image_eq [noatp]: "f -` (f ` A) = {y. EX x:A. f x = f y}"
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2169
by (blast intro: sym)
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2170
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2171
lemma image_vimage_subset: "f ` (f -` A) <= A"
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2172
by blast
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2173
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2174
lemma image_vimage_eq [simp]: "f ` (f -` A) = A Int range f"
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2175
by blast
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2176
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2177
lemma image_Int_subset: "f`(A Int B) <= f`A Int f`B"
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2178
by blast
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2179
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2180
lemma image_diff_subset: "f`A - f`B <= f`(A - B)"
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2181
by blast
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2182
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2183
lemma image_UN: "(f ` (UNION A B)) = (UN x:A.(f ` (B x)))"
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2184
by blast
f6bd8686b71e moved some set lemmas from Set.thy here
haftmann
parents: 25965
diff changeset
  2185
12257
e3f7d6fb55d7 theory Inverse_Image converted and moved to Set;
wenzelm
parents: 12114
diff changeset
  2186
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14398
diff changeset
  2187
subsection {* Getting the Contents of a Singleton Set *}
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14398
diff changeset
  2188
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2189
definition contents :: "'a set \<Rightarrow> 'a" where
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27824
diff changeset
  2190
  [code del]: "contents X = (THE x. X = {x})"
24658
49adbdcc52e2 clarified code lemmas
haftmann
parents: 24420
diff changeset
  2191
49adbdcc52e2 clarified code lemmas
haftmann
parents: 24420
diff changeset
  2192
lemma contents_eq [simp]: "contents {x} = x"
49adbdcc52e2 clarified code lemmas
haftmann
parents: 24420
diff changeset
  2193
  by (simp add: contents_def)
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14398
diff changeset
  2194
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14398
diff changeset
  2195
12023
wenzelm
parents: 12020
diff changeset
  2196
subsection {* Transitivity rules for calculational reasoning *}
12020
a24373086908 theory Calculation move to Set;
wenzelm
parents: 11982
diff changeset
  2197
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2198
lemma set_rev_mp: "x:A ==> A \<subseteq> B ==> x:B"
12020
a24373086908 theory Calculation move to Set;
wenzelm
parents: 11982
diff changeset
  2199
  by (rule subsetD)
a24373086908 theory Calculation move to Set;
wenzelm
parents: 11982
diff changeset
  2200
12897
f4d10ad0ea7b converted/deleted equalities.ML, mono.ML, subset.ML (see Set.thy);
wenzelm
parents: 12633
diff changeset
  2201
lemma set_mp: "A \<subseteq> B ==> x:A ==> x:B"
12020
a24373086908 theory Calculation move to Set;
wenzelm
parents: 11982
diff changeset
  2202
  by (rule subsetD)
a24373086908 theory Calculation move to Set;
wenzelm
parents: 11982
diff changeset
  2203
26800
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2204
lemmas basic_trans_rules [trans] =
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2205
  order_trans_rules set_rev_mp set_mp
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2206
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2207
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2208
subsection {* Least value operator *}
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2209
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2210
lemma Least_mono:
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2211
  "mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2212
    ==> (LEAST y. y : f ` S) = f (LEAST x. x : S)"
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2213
    -- {* Courtesy of Stephan Merz *}
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2214
  apply clarify
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2215
  apply (erule_tac P = "%x. x : S" in LeastI2_order, fast)
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2216
  apply (rule LeastI2_order)
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2217
  apply (auto elim: monoD intro!: order_antisym)
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2218
  done
dcf1dfc915a7 - Now uses Orderings as parent theory
berghofe
parents: 26732
diff changeset
  2219
24420
9fa337721689 made sets executable
haftmann
parents: 24331
diff changeset
  2220
27824
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2221
subsection {* Rudimentary code generation *}
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2222
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27824
diff changeset
  2223
lemma empty_code [code]: "{} x \<longleftrightarrow> False"
27824
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2224
  unfolding empty_def Collect_def ..
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2225
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27824
diff changeset
  2226
lemma UNIV_code [code]: "UNIV x \<longleftrightarrow> True"
27824
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2227
  unfolding UNIV_def Collect_def ..
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2228
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27824
diff changeset
  2229
lemma insert_code [code]: "insert y A x \<longleftrightarrow> y = x \<or> A x"
27824
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2230
  unfolding insert_def Collect_def mem_def Un_def by auto
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2231
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27824
diff changeset
  2232
lemma inter_code [code]: "(A \<inter> B) x \<longleftrightarrow> A x \<and> B x"
27824
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2233
  unfolding Int_def Collect_def mem_def ..
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2234
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27824
diff changeset
  2235
lemma union_code [code]: "(A \<union> B) x \<longleftrightarrow> A x \<or> B x"
27824
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2236
  unfolding Un_def Collect_def mem_def ..
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2237
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27824
diff changeset
  2238
lemma vimage_code [code]: "(f -` A) x = A (f x)"
27824
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2239
  unfolding vimage_def Collect_def mem_def ..
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2240
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2241
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2242
subsection {* Complete lattices *}
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2243
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2244
notation
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2245
  less_eq  (infix "\<sqsubseteq>" 50) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2246
  less (infix "\<sqsubset>" 50) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2247
  inf  (infixl "\<sqinter>" 70) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2248
  sup  (infixl "\<squnion>" 65)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2249
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2250
class complete_lattice = lattice + bot + top +
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2251
  fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2252
    and Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2253
  assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> x"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2254
     and Inf_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> \<Sqinter>A"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2255
  assumes Sup_upper: "x \<in> A \<Longrightarrow> x \<sqsubseteq> \<Squnion>A"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2256
     and Sup_least: "(\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq> z) \<Longrightarrow> \<Squnion>A \<sqsubseteq> z"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2257
begin
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2258
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2259
lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<le> a}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2260
  by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2261
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2262
lemma Sup_Inf:  "\<Squnion>A = \<Sqinter>{b. \<forall>a \<in> A. a \<le> b}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2263
  by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2264
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2265
lemma Inf_Univ: "\<Sqinter>UNIV = \<Squnion>{}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2266
  unfolding Sup_Inf by auto
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2267
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2268
lemma Sup_Univ: "\<Squnion>UNIV = \<Sqinter>{}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2269
  unfolding Inf_Sup by auto
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2270
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2271
lemma Inf_insert: "\<Sqinter>insert a A = a \<sqinter> \<Sqinter>A"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2272
  by (auto intro: antisym Inf_greatest Inf_lower)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2273
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2274
lemma Sup_insert: "\<Squnion>insert a A = a \<squnion> \<Squnion>A"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2275
  by (auto intro: antisym Sup_least Sup_upper)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2276
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2277
lemma Inf_singleton [simp]:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2278
  "\<Sqinter>{a} = a"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2279
  by (auto intro: antisym Inf_lower Inf_greatest)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2280
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2281
lemma Sup_singleton [simp]:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2282
  "\<Squnion>{a} = a"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2283
  by (auto intro: antisym Sup_upper Sup_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2284
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2285
lemma Inf_insert_simp:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2286
  "\<Sqinter>insert a A = (if A = {} then a else a \<sqinter> \<Sqinter>A)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2287
  by (cases "A = {}") (simp_all, simp add: Inf_insert)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2288
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2289
lemma Sup_insert_simp:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2290
  "\<Squnion>insert a A = (if A = {} then a else a \<squnion> \<Squnion>A)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2291
  by (cases "A = {}") (simp_all, simp add: Sup_insert)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2292
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2293
lemma Inf_binary:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2294
  "\<Sqinter>{a, b} = a \<sqinter> b"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2295
  by (simp add: Inf_insert_simp)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2296
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2297
lemma Sup_binary:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2298
  "\<Squnion>{a, b} = a \<squnion> b"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2299
  by (simp add: Sup_insert_simp)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2300
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2301
lemma bot_def:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2302
  "bot = \<Squnion>{}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2303
  by (auto intro: antisym Sup_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2305
lemma top_def:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2306
  "top = \<Sqinter>{}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2307
  by (auto intro: antisym Inf_greatest)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2308
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2309
lemma sup_bot [simp]:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2310
  "x \<squnion> bot = x"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2311
  using bot_least [of x] by (simp add: le_iff_sup sup_commute)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2312
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2313
lemma inf_top [simp]:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2314
  "x \<sqinter> top = x"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2315
  using top_greatest [of x] by (simp add: le_iff_inf inf_commute)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2316
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2317
definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2318
  "SUPR A f == \<Squnion> (f ` A)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2319
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2320
definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2321
  "INFI A f == \<Sqinter> (f ` A)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2322
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2323
end
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2324
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2325
syntax
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2326
  "_SUP1"     :: "pttrns => 'b => 'b"           ("(3SUP _./ _)" [0, 10] 10)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2327
  "_SUP"      :: "pttrn => 'a set => 'b => 'b"  ("(3SUP _:_./ _)" [0, 10] 10)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2328
  "_INF1"     :: "pttrns => 'b => 'b"           ("(3INF _./ _)" [0, 10] 10)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2329
  "_INF"      :: "pttrn => 'a set => 'b => 'b"  ("(3INF _:_./ _)" [0, 10] 10)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2330
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2331
translations
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2332
  "SUP x y. B"   == "SUP x. SUP y. B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2333
  "SUP x. B"     == "CONST SUPR CONST UNIV (%x. B)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2334
  "SUP x. B"     == "SUP x:CONST UNIV. B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2335
  "SUP x:A. B"   == "CONST SUPR A (%x. B)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2336
  "INF x y. B"   == "INF x. INF y. B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2337
  "INF x. B"     == "CONST INFI CONST UNIV (%x. B)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2338
  "INF x. B"     == "INF x:CONST UNIV. B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2339
  "INF x:A. B"   == "CONST INFI A (%x. B)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2340
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2341
(* To avoid eta-contraction of body: *)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2342
print_translation {*
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2343
let
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2344
  fun btr' syn (A :: Abs abs :: ts) =
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2345
    let val (x,t) = atomic_abs_tr' abs
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2346
    in list_comb (Syntax.const syn $ x $ A $ t, ts) end
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2347
  val const_syntax_name = Sign.const_syntax_name @{theory} o fst o dest_Const
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2348
in
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2349
[(const_syntax_name @{term SUPR}, btr' "_SUP"),(const_syntax_name @{term "INFI"}, btr' "_INF")]
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2350
end
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2351
*}
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2352
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2353
context complete_lattice
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2354
begin
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2355
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2356
lemma le_SUPI: "i : A \<Longrightarrow> M i \<le> (SUP i:A. M i)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2357
  by (auto simp add: SUPR_def intro: Sup_upper)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2358
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2359
lemma SUP_leI: "(\<And>i. i : A \<Longrightarrow> M i \<le> u) \<Longrightarrow> (SUP i:A. M i) \<le> u"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2360
  by (auto simp add: SUPR_def intro: Sup_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2361
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2362
lemma INF_leI: "i : A \<Longrightarrow> (INF i:A. M i) \<le> M i"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2363
  by (auto simp add: INFI_def intro: Inf_lower)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2364
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2365
lemma le_INFI: "(\<And>i. i : A \<Longrightarrow> u \<le> M i) \<Longrightarrow> u \<le> (INF i:A. M i)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2366
  by (auto simp add: INFI_def intro: Inf_greatest)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2367
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2368
lemma SUP_const[simp]: "A \<noteq> {} \<Longrightarrow> (SUP i:A. M) = M"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2369
  by (auto intro: antisym SUP_leI le_SUPI)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2370
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2371
lemma INF_const[simp]: "A \<noteq> {} \<Longrightarrow> (INF i:A. M) = M"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2372
  by (auto intro: antisym INF_leI le_INFI)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2373
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2374
end
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2375
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2376
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2377
subsection {* Bool as complete lattice *}
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2378
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2379
instantiation bool :: complete_lattice
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2380
begin
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2381
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2382
definition
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2383
  Inf_bool_def: "\<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2384
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2385
definition
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2386
  Sup_bool_def: "\<Squnion>A \<longleftrightarrow> (\<exists>x\<in>A. x)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2387
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2388
instance
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2389
  by intro_classes (auto simp add: Inf_bool_def Sup_bool_def le_bool_def)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2390
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2391
end
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2392
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2393
lemma Inf_empty_bool [simp]:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2394
  "\<Sqinter>{}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2395
  unfolding Inf_bool_def by auto
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2396
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2397
lemma not_Sup_empty_bool [simp]:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2398
  "\<not> Sup {}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2399
  unfolding Sup_bool_def by auto
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2400
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2401
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2402
subsection {* Fun as complete lattice *}
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2403
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2404
instantiation "fun" :: (type, complete_lattice) complete_lattice
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2405
begin
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2406
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2407
definition
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2408
  Inf_fun_def [code del]: "\<Sqinter>A = (\<lambda>x. \<Sqinter>{y. \<exists>f\<in>A. y = f x})"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2409
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2410
definition
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2411
  Sup_fun_def [code del]: "\<Squnion>A = (\<lambda>x. \<Squnion>{y. \<exists>f\<in>A. y = f x})"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2412
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2413
instance
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2414
  by intro_classes
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2415
    (auto simp add: Inf_fun_def Sup_fun_def le_fun_def
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2416
      intro: Inf_lower Sup_upper Inf_greatest Sup_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2417
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2418
end
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2419
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2420
lemma Inf_empty_fun:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2421
  "\<Sqinter>{} = (\<lambda>_. \<Sqinter>{})"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2422
  by rule (auto simp add: Inf_fun_def)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2423
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2424
lemma Sup_empty_fun:
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2425
  "\<Squnion>{} = (\<lambda>_. \<Squnion>{})"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2426
  by rule (auto simp add: Sup_fun_def)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2427
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2428
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2429
subsection {* Set as lattice *}
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2430
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2431
lemma inf_set_eq: "A \<sqinter> B = A \<inter> B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2432
  apply (rule subset_antisym)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2433
  apply (rule Int_greatest)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2434
  apply (rule inf_le1)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2435
  apply (rule inf_le2)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2436
  apply (rule inf_greatest)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2437
  apply (rule Int_lower1)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2438
  apply (rule Int_lower2)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2439
  done
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2440
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2441
lemma sup_set_eq: "A \<squnion> B = A \<union> B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2442
  apply (rule subset_antisym)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2443
  apply (rule sup_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2444
  apply (rule Un_upper1)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2445
  apply (rule Un_upper2)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2446
  apply (rule Un_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2447
  apply (rule sup_ge1)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2448
  apply (rule sup_ge2)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2449
  done
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2450
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2451
lemma mono_Int: "mono f \<Longrightarrow> f (A \<inter> B) \<subseteq> f A \<inter> f B"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2452
  apply (fold inf_set_eq sup_set_eq)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2453
  apply (erule mono_inf)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2454
  done
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2455
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2456
lemma mono_Un: "mono f \<Longrightarrow> f A \<union> f B \<subseteq> f (A \<union> B)"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2457
  apply (fold inf_set_eq sup_set_eq)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2458
  apply (erule mono_sup)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2459
  done
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2460
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2461
lemma Inf_set_eq: "\<Sqinter>S = \<Inter>S"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2462
  apply (rule subset_antisym)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2463
  apply (rule Inter_greatest)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2464
  apply (erule Inf_lower)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2465
  apply (rule Inf_greatest)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2466
  apply (erule Inter_lower)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2467
  done
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2468
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2469
lemma Sup_set_eq: "\<Squnion>S = \<Union>S"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2470
  apply (rule subset_antisym)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2471
  apply (rule Sup_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2472
  apply (erule Union_upper)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2473
  apply (rule Union_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2474
  apply (erule Sup_upper)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2475
  done
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2476
  
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2477
lemma top_set_eq: "top = UNIV"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2478
  by (iprover intro!: subset_antisym subset_UNIV top_greatest)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2479
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2480
lemma bot_set_eq: "bot = {}"
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2481
  by (iprover intro!: subset_antisym empty_subsetI bot_least)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2482
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2483
no_notation
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2484
  less_eq  (infix "\<sqsubseteq>" 50) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2485
  less (infix "\<sqsubset>" 50) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2486
  inf  (infixl "\<sqinter>" 70) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2487
  sup  (infixl "\<squnion>" 65) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2488
  Inf  ("\<Sqinter>_" [900] 900) and
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2489
  Sup  ("\<Squnion>_" [900] 900)
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29901
diff changeset
  2490
27824
97d2a3797ce0 rudimentary code setup for set operations
haftmann
parents: 27418
diff changeset
  2491
21669
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21549
diff changeset
  2492
subsection {* Basic ML bindings *}
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21549
diff changeset
  2493
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21549
diff changeset
  2494
ML {*
22139
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2495
val Ball_def = @{thm Ball_def}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2496
val Bex_def = @{thm Bex_def}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2497
val CollectD = @{thm CollectD}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2498
val CollectE = @{thm CollectE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2499
val CollectI = @{thm CollectI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2500
val Collect_conj_eq = @{thm Collect_conj_eq}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2501
val Collect_mem_eq = @{thm Collect_mem_eq}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2502
val IntD1 = @{thm IntD1}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2503
val IntD2 = @{thm IntD2}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2504
val IntE = @{thm IntE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2505
val IntI = @{thm IntI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2506
val Int_Collect = @{thm Int_Collect}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2507
val UNIV_I = @{thm UNIV_I}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2508
val UNIV_witness = @{thm UNIV_witness}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2509
val UnE = @{thm UnE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2510
val UnI1 = @{thm UnI1}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2511
val UnI2 = @{thm UnI2}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2512
val ballE = @{thm ballE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2513
val ballI = @{thm ballI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2514
val bexCI = @{thm bexCI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2515
val bexE = @{thm bexE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2516
val bexI = @{thm bexI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2517
val bex_triv = @{thm bex_triv}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2518
val bspec = @{thm bspec}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2519
val contra_subsetD = @{thm contra_subsetD}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2520
val distinct_lemma = @{thm distinct_lemma}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2521
val eq_to_mono = @{thm eq_to_mono}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2522
val eq_to_mono2 = @{thm eq_to_mono2}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2523
val equalityCE = @{thm equalityCE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2524
val equalityD1 = @{thm equalityD1}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2525
val equalityD2 = @{thm equalityD2}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2526
val equalityE = @{thm equalityE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2527
val equalityI = @{thm equalityI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2528
val imageE = @{thm imageE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2529
val imageI = @{thm imageI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2530
val image_Un = @{thm image_Un}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2531
val image_insert = @{thm image_insert}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2532
val insert_commute = @{thm insert_commute}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2533
val insert_iff = @{thm insert_iff}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2534
val mem_Collect_eq = @{thm mem_Collect_eq}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2535
val rangeE = @{thm rangeE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2536
val rangeI = @{thm rangeI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2537
val range_eqI = @{thm range_eqI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2538
val subsetCE = @{thm subsetCE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2539
val subsetD = @{thm subsetD}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2540
val subsetI = @{thm subsetI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2541
val subset_refl = @{thm subset_refl}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2542
val subset_trans = @{thm subset_trans}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2543
val vimageD = @{thm vimageD}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2544
val vimageE = @{thm vimageE}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2545
val vimageI = @{thm vimageI}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2546
val vimageI2 = @{thm vimageI2}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2547
val vimage_Collect = @{thm vimage_Collect}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2548
val vimage_Int = @{thm vimage_Int}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21833
diff changeset
  2549
val vimage_Un = @{thm vimage_Un}
21669
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21549
diff changeset
  2550
*}
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21549
diff changeset
  2551
11979
0a3dace545c5 converted theory "Set";
wenzelm
parents: 11752
diff changeset
  2552
end