clasohm@923: (* Title: HOL/Set.thy wenzelm@12257: Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel clasohm@923: *) clasohm@923: wenzelm@11979: header {* Set theory for higher-order logic *} wenzelm@11979: nipkow@15131: theory Set haftmann@30304: imports Lattices nipkow@15131: begin wenzelm@11979: wenzelm@11979: text {* A set in HOL is simply a predicate. *} clasohm@923: haftmann@30531: haftmann@30531: subsection {* Basic syntax *} haftmann@30531: wenzelm@3947: global wenzelm@3947: berghofe@26800: types 'a set = "'a => bool" wenzelm@3820: clasohm@923: consts haftmann@30531: Collect :: "('a => bool) => 'a set" -- "comprehension" haftmann@30531: "op :" :: "'a => 'a set => bool" -- "membership" haftmann@30531: insert :: "'a => 'a set => 'a set" haftmann@30531: Ball :: "'a set => ('a => bool) => bool" -- "bounded universal quantifiers" haftmann@30531: Bex :: "'a set => ('a => bool) => bool" -- "bounded existential quantifiers" haftmann@30531: Bex1 :: "'a set => ('a => bool) => bool" -- "bounded unique existential quantifiers" haftmann@30531: Pow :: "'a set => 'a set set" -- "powerset" haftmann@30531: image :: "('a => 'b) => 'a set => 'b set" (infixr "`" 90) haftmann@30304: haftmann@30304: local wenzelm@19656: wenzelm@21210: notation wenzelm@21404: "op :" ("op :") and wenzelm@19656: "op :" ("(_/ : _)" [50, 51] 50) wenzelm@11979: wenzelm@19656: abbreviation wenzelm@21404: "not_mem x A == ~ (x : A)" -- "non-membership" wenzelm@19656: wenzelm@21210: notation wenzelm@21404: not_mem ("op ~:") and wenzelm@19656: not_mem ("(_/ ~: _)" [50, 51] 50) wenzelm@19656: wenzelm@21210: notation (xsymbols) wenzelm@21404: "op :" ("op \") and wenzelm@21404: "op :" ("(_/ \ _)" [50, 51] 50) and wenzelm@21404: not_mem ("op \") and haftmann@30304: not_mem ("(_/ \ _)" [50, 51] 50) wenzelm@19656: wenzelm@21210: notation (HTML output) wenzelm@21404: "op :" ("op \") and wenzelm@21404: "op :" ("(_/ \ _)" [50, 51] 50) and wenzelm@21404: not_mem ("op \") and wenzelm@19656: not_mem ("(_/ \ _)" [50, 51] 50) wenzelm@19656: haftmann@30531: syntax haftmann@30531: "@Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})") haftmann@30531: haftmann@30531: translations haftmann@30531: "{x. P}" == "Collect (%x. P)" haftmann@30531: haftmann@30531: definition empty :: "'a set" ("{}") where haftmann@30531: "empty \ {x. False}" haftmann@30531: haftmann@30531: definition UNIV :: "'a set" where haftmann@30531: "UNIV \ {x. True}" haftmann@30531: haftmann@30531: syntax haftmann@30531: "@Finset" :: "args => 'a set" ("{(_)}") haftmann@30531: haftmann@30531: translations haftmann@30531: "{x, xs}" == "insert x {xs}" haftmann@30531: "{x}" == "insert x {}" haftmann@30531: haftmann@30531: definition Int :: "'a set \ 'a set \ 'a set" (infixl "Int" 70) where haftmann@30531: "A Int B \ {x. x \ A \ x \ B}" haftmann@30531: haftmann@30531: definition Un :: "'a set \ 'a set \ 'a set" (infixl "Un" 65) where haftmann@30531: "A Un B \ {x. x \ A \ x \ B}" haftmann@30531: haftmann@30531: notation (xsymbols) haftmann@30531: "Int" (infixl "\" 70) and haftmann@30531: "Un" (infixl "\" 65) haftmann@30531: haftmann@30531: notation (HTML output) haftmann@30531: "Int" (infixl "\" 70) and haftmann@30531: "Un" (infixl "\" 65) haftmann@30531: haftmann@30531: syntax haftmann@30531: "_Ball" :: "pttrn => 'a set => bool => bool" ("(3ALL _:_./ _)" [0, 0, 10] 10) haftmann@30531: "_Bex" :: "pttrn => 'a set => bool => bool" ("(3EX _:_./ _)" [0, 0, 10] 10) haftmann@30531: "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3EX! _:_./ _)" [0, 0, 10] 10) haftmann@30531: "_Bleast" :: "id => 'a set => bool => 'a" ("(3LEAST _:_./ _)" [0, 0, 10] 10) haftmann@30531: haftmann@30531: syntax (HOL) haftmann@30531: "_Ball" :: "pttrn => 'a set => bool => bool" ("(3! _:_./ _)" [0, 0, 10] 10) haftmann@30531: "_Bex" :: "pttrn => 'a set => bool => bool" ("(3? _:_./ _)" [0, 0, 10] 10) haftmann@30531: "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3?! _:_./ _)" [0, 0, 10] 10) haftmann@30531: haftmann@30531: syntax (xsymbols) haftmann@30531: "_Ball" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@30531: "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@30531: "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) haftmann@30531: "_Bleast" :: "id => 'a set => bool => 'a" ("(3LEAST_\_./ _)" [0, 0, 10] 10) haftmann@30531: haftmann@30531: syntax (HTML output) haftmann@30531: "_Ball" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@30531: "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@30531: "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) haftmann@30531: haftmann@30531: translations haftmann@30531: "ALL x:A. P" == "Ball A (%x. P)" haftmann@30531: "EX x:A. P" == "Bex A (%x. P)" haftmann@30531: "EX! x:A. P" == "Bex1 A (%x. P)" haftmann@30531: "LEAST x:A. P" => "LEAST x. x:A & P" haftmann@30531: haftmann@30531: definition INTER :: "'a set \ ('a \ 'b set) \ 'b set" where haftmann@30531: "INTER A B \ {y. \x\A. y \ B x}" haftmann@30531: haftmann@30531: definition UNION :: "'a set \ ('a \ 'b set) \ 'b set" where haftmann@30531: "UNION A B \ {y. \x\A. y \ B x}" haftmann@30531: haftmann@30531: definition Inter :: "'a set set \ 'a set" where haftmann@30531: "Inter S \ INTER S (\x. x)" haftmann@30531: haftmann@30531: definition Union :: "'a set set \ 'a set" where haftmann@30531: "Union S \ UNION S (\x. x)" haftmann@30531: haftmann@30531: notation (xsymbols) haftmann@30531: Inter ("\_" [90] 90) and haftmann@30531: Union ("\_" [90] 90) haftmann@30531: haftmann@30531: haftmann@30531: subsection {* Additional concrete syntax *} haftmann@30531: haftmann@30531: syntax haftmann@30531: "@SetCompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})") haftmann@30531: "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ :/ _./ _})") haftmann@30531: "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) haftmann@30531: "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) haftmann@30531: "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 10] 10) haftmann@30531: "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 10] 10) haftmann@30531: haftmann@30531: syntax (xsymbols) haftmann@30531: "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ \/ _./ _})") haftmann@30531: "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) haftmann@30531: "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) haftmann@30531: "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) haftmann@30531: "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) haftmann@30531: haftmann@30531: syntax (latex output) haftmann@30531: "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) haftmann@30531: "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) haftmann@30531: "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) haftmann@30531: "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) haftmann@30531: haftmann@30531: translations haftmann@30531: "{x:A. P}" => "{x. x:A & P}" haftmann@30531: "INT x y. B" == "INT x. INT y. B" haftmann@30531: "INT x. B" == "CONST INTER CONST UNIV (%x. B)" haftmann@30531: "INT x. B" == "INT x:CONST UNIV. B" haftmann@30531: "INT x:A. B" == "CONST INTER A (%x. B)" haftmann@30531: "UN x y. B" == "UN x. UN y. B" haftmann@30531: "UN x. B" == "CONST UNION CONST UNIV (%x. B)" haftmann@30531: "UN x. B" == "UN x:CONST UNIV. B" haftmann@30531: "UN x:A. B" == "CONST UNION A (%x. B)" haftmann@30531: haftmann@30531: text {* haftmann@30531: Note the difference between ordinary xsymbol syntax of indexed haftmann@30531: unions and intersections (e.g.\ @{text"\a\<^isub>1\A\<^isub>1. B"}) haftmann@30531: and their \LaTeX\ rendition: @{term"\a\<^isub>1\A\<^isub>1. B"}. The haftmann@30531: former does not make the index expression a subscript of the haftmann@30531: union/intersection symbol because this leads to problems with nested haftmann@30531: subscripts in Proof General. haftmann@30531: *} wenzelm@2261: haftmann@21333: abbreviation wenzelm@21404: subset :: "'a set \ 'a set \ bool" where haftmann@21819: "subset \ less" wenzelm@21404: wenzelm@21404: abbreviation wenzelm@21404: subset_eq :: "'a set \ 'a set \ bool" where haftmann@21819: "subset_eq \ less_eq" haftmann@21333: haftmann@21333: notation (output) wenzelm@21404: subset ("op <") and wenzelm@21404: subset ("(_/ < _)" [50, 51] 50) and wenzelm@21404: subset_eq ("op <=") and haftmann@21333: subset_eq ("(_/ <= _)" [50, 51] 50) haftmann@21333: haftmann@21333: notation (xsymbols) wenzelm@21404: subset ("op \") and wenzelm@21404: subset ("(_/ \ _)" [50, 51] 50) and wenzelm@21404: subset_eq ("op \") and haftmann@21333: subset_eq ("(_/ \ _)" [50, 51] 50) haftmann@21333: haftmann@21333: notation (HTML output) wenzelm@21404: subset ("op \") and wenzelm@21404: subset ("(_/ \ _)" [50, 51] 50) and wenzelm@21404: subset_eq ("op \") and haftmann@21333: subset_eq ("(_/ \ _)" [50, 51] 50) haftmann@21333: haftmann@21333: abbreviation (input) haftmann@21819: supset :: "'a set \ 'a set \ bool" where haftmann@21819: "supset \ greater" wenzelm@21404: wenzelm@21404: abbreviation (input) haftmann@21819: supset_eq :: "'a set \ 'a set \ bool" where haftmann@21819: "supset_eq \ greater_eq" haftmann@21819: haftmann@21819: notation (xsymbols) haftmann@21819: supset ("op \") and haftmann@21819: supset ("(_/ \ _)" [50, 51] 50) and haftmann@21819: supset_eq ("op \") and haftmann@21819: supset_eq ("(_/ \ _)" [50, 51] 50) haftmann@21333: haftmann@30531: abbreviation haftmann@30531: range :: "('a => 'b) => 'b set" where -- "of function" haftmann@30531: "range f == f ` UNIV" haftmann@30531: haftmann@30531: haftmann@30531: subsubsection "Bounded quantifiers" nipkow@14804: wenzelm@19656: syntax (output) nipkow@14804: "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) nipkow@14804: "_setlessEx" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10) nipkow@14804: "_setleAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) nipkow@14804: "_setleEx" :: "[idt, 'a, bool] => bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) webertj@20217: "_setleEx1" :: "[idt, 'a, bool] => bool" ("(3EX! _<=_./ _)" [0, 0, 10] 10) nipkow@14804: nipkow@14804: syntax (xsymbols) nipkow@14804: "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) nipkow@14804: "_setlessEx" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) nipkow@14804: "_setleAll" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) nipkow@14804: "_setleEx" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) webertj@20217: "_setleEx1" :: "[idt, 'a, bool] => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) nipkow@14804: wenzelm@19656: syntax (HOL output) nipkow@14804: "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3! _<_./ _)" [0, 0, 10] 10) nipkow@14804: "_setlessEx" :: "[idt, 'a, bool] => bool" ("(3? _<_./ _)" [0, 0, 10] 10) nipkow@14804: "_setleAll" :: "[idt, 'a, bool] => bool" ("(3! _<=_./ _)" [0, 0, 10] 10) nipkow@14804: "_setleEx" :: "[idt, 'a, bool] => bool" ("(3? _<=_./ _)" [0, 0, 10] 10) webertj@20217: "_setleEx1" :: "[idt, 'a, bool] => bool" ("(3?! _<=_./ _)" [0, 0, 10] 10) nipkow@14804: nipkow@14804: syntax (HTML output) nipkow@14804: "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) nipkow@14804: "_setlessEx" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) nipkow@14804: "_setleAll" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) nipkow@14804: "_setleEx" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) webertj@20217: "_setleEx1" :: "[idt, 'a, bool] => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) nipkow@14804: nipkow@14804: translations haftmann@30531: "\A\B. P" => "ALL A. A \ B --> P" haftmann@30531: "\A\B. P" => "EX A. A \ B & P" haftmann@30531: "\A\B. P" => "ALL A. A \ B --> P" haftmann@30531: "\A\B. P" => "EX A. A \ B & P" haftmann@30531: "\!A\B. P" => "EX! A. A \ B & P" nipkow@14804: nipkow@14804: print_translation {* nipkow@14804: let wenzelm@22377: val Type (set_type, _) = @{typ "'a set"}; wenzelm@22377: val All_binder = Syntax.binder_name @{const_syntax "All"}; wenzelm@22377: val Ex_binder = Syntax.binder_name @{const_syntax "Ex"}; wenzelm@22377: val impl = @{const_syntax "op -->"}; wenzelm@22377: val conj = @{const_syntax "op &"}; wenzelm@22377: val sbset = @{const_syntax "subset"}; wenzelm@22377: val sbset_eq = @{const_syntax "subset_eq"}; haftmann@21819: haftmann@21819: val trans = haftmann@21819: [((All_binder, impl, sbset), "_setlessAll"), haftmann@21819: ((All_binder, impl, sbset_eq), "_setleAll"), haftmann@21819: ((Ex_binder, conj, sbset), "_setlessEx"), haftmann@21819: ((Ex_binder, conj, sbset_eq), "_setleEx")]; haftmann@21819: haftmann@21819: fun mk v v' c n P = haftmann@21819: if v = v' andalso not (Term.exists_subterm (fn Free (x, _) => x = v | _ => false) n) haftmann@21819: then Syntax.const c $ Syntax.mark_bound v' $ n $ P else raise Match; haftmann@21819: haftmann@21819: fun tr' q = (q, haftmann@21819: fn [Const ("_bound", _) $ Free (v, Type (T, _)), Const (c, _) $ (Const (d, _) $ (Const ("_bound", _) $ Free (v', _)) $ n) $ P] => haftmann@21819: if T = (set_type) then case AList.lookup (op =) trans (q, c, d) haftmann@21819: of NONE => raise Match haftmann@21819: | SOME l => mk v v' l n P haftmann@21819: else raise Match haftmann@21819: | _ => raise Match); nipkow@14804: in haftmann@21819: [tr' All_binder, tr' Ex_binder] nipkow@14804: end nipkow@14804: *} nipkow@14804: haftmann@30531: wenzelm@11979: text {* wenzelm@11979: \medskip Translate between @{text "{e | x1...xn. P}"} and @{text wenzelm@11979: "{u. EX x1..xn. u = e & P}"}; @{text "{y. EX x1..xn. y = e & P}"} is wenzelm@11979: only translated if @{text "[0..n] subset bvs(e)"}. wenzelm@11979: *} wenzelm@11979: wenzelm@11979: parse_translation {* wenzelm@11979: let wenzelm@11979: val ex_tr = snd (mk_binder_tr ("EX ", "Ex")); wenzelm@3947: wenzelm@11979: fun nvars (Const ("_idts", _) $ _ $ idts) = nvars idts + 1 wenzelm@11979: | nvars _ = 1; wenzelm@11979: wenzelm@11979: fun setcompr_tr [e, idts, b] = wenzelm@11979: let wenzelm@11979: val eq = Syntax.const "op =" $ Bound (nvars idts) $ e; wenzelm@11979: val P = Syntax.const "op &" $ eq $ b; wenzelm@11979: val exP = ex_tr [idts, P]; wenzelm@17784: in Syntax.const "Collect" $ Term.absdummy (dummyT, exP) end; wenzelm@11979: wenzelm@11979: in [("@SetCompr", setcompr_tr)] end; wenzelm@11979: *} clasohm@923: haftmann@30531: (* To avoid eta-contraction of body: *) haftmann@30531: print_translation {* haftmann@30531: let haftmann@30531: fun btr' syn [A, Abs abs] = haftmann@30531: let val (x, t) = atomic_abs_tr' abs haftmann@30531: in Syntax.const syn $ x $ A $ t end haftmann@30531: in haftmann@30531: [(@{const_syntax Ball}, btr' "_Ball"), (@{const_syntax Bex}, btr' "_Bex"), haftmann@30531: (@{const_syntax UNION}, btr' "@UNION"),(@{const_syntax INTER}, btr' "@INTER")] haftmann@30531: end haftmann@30531: *} haftmann@30531: nipkow@13763: print_translation {* nipkow@13763: let nipkow@13763: val ex_tr' = snd (mk_binder_tr' ("Ex", "DUMMY")); nipkow@13763: nipkow@13763: fun setcompr_tr' [Abs (abs as (_, _, P))] = nipkow@13763: let nipkow@13763: fun check (Const ("Ex", _) $ Abs (_, _, P), n) = check (P, n + 1) nipkow@13763: | check (Const ("op &", _) $ (Const ("op =", _) $ Bound m $ e) $ P, n) = nipkow@13763: n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso nipkow@13763: ((0 upto (n - 1)) subset add_loose_bnos (e, 0, [])) nipkow@13764: | check _ = false clasohm@923: wenzelm@11979: fun tr' (_ $ abs) = wenzelm@11979: let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr' [abs] wenzelm@11979: in Syntax.const "@SetCompr" $ e $ idts $ Q end; nipkow@13763: in if check (P, 0) then tr' P nipkow@15535: else let val (x as _ $ Free(xN,_), t) = atomic_abs_tr' abs nipkow@15535: val M = Syntax.const "@Coll" $ x $ t nipkow@15535: in case t of nipkow@15535: Const("op &",_) nipkow@15535: $ (Const("op :",_) $ (Const("_bound",_) $ Free(yN,_)) $ A) nipkow@15535: $ P => nipkow@15535: if xN=yN then Syntax.const "@Collect" $ x $ A $ P else M nipkow@15535: | _ => M nipkow@15535: end nipkow@13763: end; wenzelm@11979: in [("Collect", setcompr_tr')] end; wenzelm@11979: *} wenzelm@11979: haftmann@30531: haftmann@30531: subsection {* Rules and definitions *} haftmann@30531: haftmann@30531: text {* Isomorphisms between predicates and sets. *} haftmann@30531: haftmann@30531: defs haftmann@30531: mem_def [code]: "x : S == S x" haftmann@30531: Collect_def [code]: "Collect P == P" haftmann@30531: haftmann@30531: defs haftmann@30531: Ball_def: "Ball A P == ALL x. x:A --> P(x)" haftmann@30531: Bex_def: "Bex A P == EX x. x:A & P(x)" haftmann@30531: Bex1_def: "Bex1 A P == EX! x. x:A & P(x)" haftmann@30531: haftmann@30531: instantiation "fun" :: (type, minus) minus haftmann@30531: begin haftmann@30531: haftmann@30531: definition haftmann@30531: fun_diff_def: "A - B = (%x. A x - B x)" haftmann@30531: haftmann@30531: instance .. haftmann@30531: haftmann@30531: end haftmann@30531: haftmann@30531: instantiation bool :: minus haftmann@30531: begin haftmann@30531: haftmann@30531: definition haftmann@30531: bool_diff_def: "A - B = (A & ~ B)" haftmann@30531: haftmann@30531: instance .. haftmann@30531: haftmann@30531: end haftmann@30531: haftmann@30531: instantiation "fun" :: (type, uminus) uminus haftmann@30531: begin haftmann@30531: haftmann@30531: definition haftmann@30531: fun_Compl_def: "- A = (%x. - A x)" haftmann@30531: haftmann@30531: instance .. haftmann@30531: haftmann@30531: end haftmann@30531: haftmann@30531: instantiation bool :: uminus haftmann@30531: begin haftmann@30531: haftmann@30531: definition haftmann@30531: bool_Compl_def: "- A = (~ A)" haftmann@30531: haftmann@30531: instance .. haftmann@30531: haftmann@30531: end haftmann@30531: haftmann@30531: defs haftmann@30531: Pow_def: "Pow A == {B. B <= A}" haftmann@30531: insert_def: "insert a B == {x. x=a} Un B" haftmann@30531: image_def: "f`A == {y. EX x:A. y = f(x)}" haftmann@30531: haftmann@30531: haftmann@30531: subsection {* Lemmas and proof tool setup *} haftmann@30531: haftmann@30531: subsubsection {* Relating predicates and sets *} haftmann@30531: haftmann@30531: lemma mem_Collect_eq [iff]: "(a : {x. P(x)}) = P(a)" haftmann@30531: by (simp add: Collect_def mem_def) haftmann@30531: haftmann@30531: lemma Collect_mem_eq [simp]: "{x. x:A} = A" haftmann@30531: by (simp add: Collect_def mem_def) haftmann@30531: haftmann@30531: lemma CollectI: "P(a) ==> a : {x. P(x)}" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma CollectD: "a : {x. P(x)} ==> P(a)" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma Collect_cong: "(!!x. P x = Q x) ==> {x. P(x)} = {x. Q(x)}" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemmas CollectE = CollectD [elim_format] haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* Bounded quantifiers *} haftmann@30531: wenzelm@11979: lemma ballI [intro!]: "(!!x. x:A ==> P x) ==> ALL x:A. P x" wenzelm@11979: by (simp add: Ball_def) wenzelm@11979: wenzelm@11979: lemmas strip = impI allI ballI wenzelm@11979: wenzelm@11979: lemma bspec [dest?]: "ALL x:A. P x ==> x:A ==> P x" wenzelm@11979: by (simp add: Ball_def) wenzelm@11979: wenzelm@11979: lemma ballE [elim]: "ALL x:A. P x ==> (P x ==> Q) ==> (x ~: A ==> Q) ==> Q" wenzelm@11979: by (unfold Ball_def) blast wenzelm@22139: wenzelm@22139: ML {* bind_thm ("rev_ballE", permute_prems 1 1 @{thm ballE}) *} wenzelm@11979: wenzelm@11979: text {* wenzelm@11979: \medskip This tactic takes assumptions @{prop "ALL x:A. P x"} and wenzelm@11979: @{prop "a:A"}; creates assumption @{prop "P a"}. wenzelm@11979: *} wenzelm@11979: wenzelm@11979: ML {* wenzelm@22139: fun ball_tac i = etac @{thm ballE} i THEN contr_tac (i + 1) wenzelm@11979: *} wenzelm@11979: wenzelm@11979: text {* wenzelm@11979: Gives better instantiation for bound: wenzelm@11979: *} wenzelm@11979: wenzelm@26339: declaration {* fn _ => wenzelm@26339: Classical.map_cs (fn cs => cs addbefore ("bspec", datac @{thm bspec} 1)) wenzelm@11979: *} wenzelm@11979: wenzelm@11979: lemma bexI [intro]: "P x ==> x:A ==> EX x:A. P x" wenzelm@11979: -- {* Normally the best argument order: @{prop "P x"} constrains the wenzelm@11979: choice of @{prop "x:A"}. *} wenzelm@11979: by (unfold Bex_def) blast wenzelm@11979: wenzelm@13113: lemma rev_bexI [intro?]: "x:A ==> P x ==> EX x:A. P x" wenzelm@11979: -- {* The best argument order when there is only one @{prop "x:A"}. *} wenzelm@11979: by (unfold Bex_def) blast wenzelm@11979: wenzelm@11979: lemma bexCI: "(ALL x:A. ~P x ==> P a) ==> a:A ==> EX x:A. P x" wenzelm@11979: by (unfold Bex_def) blast wenzelm@11979: wenzelm@11979: lemma bexE [elim!]: "EX x:A. P x ==> (!!x. x:A ==> P x ==> Q) ==> Q" wenzelm@11979: by (unfold Bex_def) blast wenzelm@11979: wenzelm@11979: lemma ball_triv [simp]: "(ALL x:A. P) = ((EX x. x:A) --> P)" wenzelm@11979: -- {* Trival rewrite rule. *} wenzelm@11979: by (simp add: Ball_def) wenzelm@11979: wenzelm@11979: lemma bex_triv [simp]: "(EX x:A. P) = ((EX x. x:A) & P)" wenzelm@11979: -- {* Dual form for existentials. *} wenzelm@11979: by (simp add: Bex_def) wenzelm@11979: wenzelm@11979: lemma bex_triv_one_point1 [simp]: "(EX x:A. x = a) = (a:A)" wenzelm@11979: by blast wenzelm@11979: wenzelm@11979: lemma bex_triv_one_point2 [simp]: "(EX x:A. a = x) = (a:A)" wenzelm@11979: by blast wenzelm@11979: wenzelm@11979: lemma bex_one_point1 [simp]: "(EX x:A. x = a & P x) = (a:A & P a)" wenzelm@11979: by blast wenzelm@11979: wenzelm@11979: lemma bex_one_point2 [simp]: "(EX x:A. a = x & P x) = (a:A & P a)" wenzelm@11979: by blast wenzelm@11979: wenzelm@11979: lemma ball_one_point1 [simp]: "(ALL x:A. x = a --> P x) = (a:A --> P a)" wenzelm@11979: by blast wenzelm@11979: wenzelm@11979: lemma ball_one_point2 [simp]: "(ALL x:A. a = x --> P x) = (a:A --> P a)" wenzelm@11979: by blast wenzelm@11979: wenzelm@26480: ML {* wenzelm@13462: local wenzelm@22139: val unfold_bex_tac = unfold_tac @{thms "Bex_def"}; wenzelm@18328: fun prove_bex_tac ss = unfold_bex_tac ss THEN Quantifier1.prove_one_point_ex_tac; wenzelm@11979: val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac; wenzelm@11979: wenzelm@22139: val unfold_ball_tac = unfold_tac @{thms "Ball_def"}; wenzelm@18328: fun prove_ball_tac ss = unfold_ball_tac ss THEN Quantifier1.prove_one_point_all_tac; wenzelm@11979: val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac; wenzelm@11979: in wenzelm@18328: val defBEX_regroup = Simplifier.simproc (the_context ()) wenzelm@13462: "defined BEX" ["EX x:A. P x & Q x"] rearrange_bex; wenzelm@18328: val defBALL_regroup = Simplifier.simproc (the_context ()) wenzelm@13462: "defined BALL" ["ALL x:A. P x --> Q x"] rearrange_ball; wenzelm@11979: end; wenzelm@13462: wenzelm@13462: Addsimprocs [defBALL_regroup, defBEX_regroup]; wenzelm@11979: *} wenzelm@11979: haftmann@30531: haftmann@30531: subsubsection {* Congruence rules *} wenzelm@11979: berghofe@16636: lemma ball_cong: wenzelm@11979: "A = B ==> (!!x. x:B ==> P x = Q x) ==> wenzelm@11979: (ALL x:A. P x) = (ALL x:B. Q x)" wenzelm@11979: by (simp add: Ball_def) wenzelm@11979: berghofe@16636: lemma strong_ball_cong [cong]: berghofe@16636: "A = B ==> (!!x. x:B =simp=> P x = Q x) ==> berghofe@16636: (ALL x:A. P x) = (ALL x:B. Q x)" berghofe@16636: by (simp add: simp_implies_def Ball_def) berghofe@16636: berghofe@16636: lemma bex_cong: wenzelm@11979: "A = B ==> (!!x. x:B ==> P x = Q x) ==> wenzelm@11979: (EX x:A. P x) = (EX x:B. Q x)" wenzelm@11979: by (simp add: Bex_def cong: conj_cong) regensbu@1273: berghofe@16636: lemma strong_bex_cong [cong]: berghofe@16636: "A = B ==> (!!x. x:B =simp=> P x = Q x) ==> berghofe@16636: (EX x:A. P x) = (EX x:B. Q x)" berghofe@16636: by (simp add: simp_implies_def Bex_def cong: conj_cong) berghofe@16636: haftmann@30531: haftmann@30531: subsubsection {* Subsets *} haftmann@30531: haftmann@30531: lemma subsetI [atp,intro!]: "(!!x. x:A ==> x:B) ==> A \ B" haftmann@30531: by (auto simp add: mem_def intro: predicate1I) haftmann@30352: wenzelm@11979: text {* haftmann@30531: \medskip Map the type @{text "'a set => anything"} to just @{typ haftmann@30531: 'a}; for overloading constants whose first argument has type @{typ haftmann@30531: "'a set"}. wenzelm@11979: *} wenzelm@11979: haftmann@30596: lemma subsetD [elim, intro?]: "A \ B ==> c \ A ==> c \ B" haftmann@30531: -- {* Rule in Modus Ponens style. *} haftmann@30531: by (unfold mem_def) blast haftmann@30531: haftmann@30596: lemma rev_subsetD [intro?]: "c \ A ==> A \ B ==> c \ B" haftmann@30531: -- {* The same, with reversed premises for use with @{text erule} -- haftmann@30531: cf @{text rev_mp}. *} haftmann@30531: by (rule subsetD) haftmann@30531: wenzelm@11979: text {* haftmann@30531: \medskip Converts @{prop "A \ B"} to @{prop "x \ A ==> x \ B"}. haftmann@30531: *} haftmann@30531: haftmann@30531: ML {* haftmann@30531: fun impOfSubs th = th RSN (2, @{thm rev_subsetD}) wenzelm@11979: *} wenzelm@11979: haftmann@30531: lemma subsetCE [elim]: "A \ B ==> (c \ A ==> P) ==> (c \ B ==> P) ==> P" haftmann@30531: -- {* Classical elimination rule. *} haftmann@30531: by (unfold mem_def) blast haftmann@30531: haftmann@30531: lemma subset_eq: "A \ B = (\x\A. x \ B)" by blast wenzelm@2388: wenzelm@11979: text {* haftmann@30531: \medskip Takes assumptions @{prop "A \ B"}; @{prop "c \ A"} and haftmann@30531: creates the assumption @{prop "c \ B"}. haftmann@30352: *} haftmann@30352: haftmann@30352: ML {* haftmann@30531: fun set_mp_tac i = etac @{thm subsetCE} i THEN mp_tac i wenzelm@11979: *} wenzelm@11979: haftmann@30531: lemma contra_subsetD: "A \ B ==> c \ B ==> c \ A" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma subset_refl [simp,atp]: "A \ A" haftmann@30531: by fast haftmann@30531: haftmann@30531: lemma subset_trans: "A \ B ==> B \ C ==> A \ C" haftmann@30531: by blast haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* Equality *} haftmann@30531: haftmann@30531: lemma set_ext: assumes prem: "(!!x. (x:A) = (x:B))" shows "A = B" haftmann@30531: apply (rule prem [THEN ext, THEN arg_cong, THEN box_equals]) haftmann@30531: apply (rule Collect_mem_eq) haftmann@30531: apply (rule Collect_mem_eq) haftmann@30531: done haftmann@30531: haftmann@30531: (* Due to Brian Huffman *) haftmann@30531: lemma expand_set_eq: "(A = B) = (ALL x. (x:A) = (x:B))" haftmann@30531: by(auto intro:set_ext) haftmann@30531: haftmann@30531: lemma subset_antisym [intro!]: "A \ B ==> B \ A ==> A = B" haftmann@30531: -- {* Anti-symmetry of the subset relation. *} haftmann@30531: by (iprover intro: set_ext subsetD) haftmann@30531: haftmann@30531: text {* haftmann@30531: \medskip Equality rules from ZF set theory -- are they appropriate haftmann@30531: here? haftmann@30531: *} haftmann@30531: haftmann@30531: lemma equalityD1: "A = B ==> A \ B" haftmann@30531: by (simp add: subset_refl) haftmann@30531: haftmann@30531: lemma equalityD2: "A = B ==> B \ A" haftmann@30531: by (simp add: subset_refl) haftmann@30531: haftmann@30531: text {* haftmann@30531: \medskip Be careful when adding this to the claset as @{text haftmann@30531: subset_empty} is in the simpset: @{prop "A = {}"} goes to @{prop "{} haftmann@30531: \ A"} and @{prop "A \ {}"} and then back to @{prop "A = {}"}! haftmann@30352: *} haftmann@30352: haftmann@30531: lemma equalityE: "A = B ==> (A \ B ==> B \ A ==> P) ==> P" haftmann@30531: by (simp add: subset_refl) haftmann@30531: haftmann@30531: lemma equalityCE [elim]: haftmann@30531: "A = B ==> (c \ A ==> c \ B ==> P) ==> (c \ A ==> c \ B ==> P) ==> P" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma eqset_imp_iff: "A = B ==> (x : A) = (x : B)" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma eqelem_imp_iff: "x = y ==> (x : A) = (y : A)" haftmann@30531: by simp haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* The universal set -- UNIV *} haftmann@30531: haftmann@30531: lemma UNIV_I [simp]: "x : UNIV" haftmann@30531: by (simp add: UNIV_def) haftmann@30531: haftmann@30531: declare UNIV_I [intro] -- {* unsafe makes it less likely to cause problems *} haftmann@30531: haftmann@30531: lemma UNIV_witness [intro?]: "EX x. x : UNIV" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma subset_UNIV [simp]: "A \ UNIV" haftmann@30531: by (rule subsetI) (rule UNIV_I) haftmann@30531: haftmann@30531: text {* haftmann@30531: \medskip Eta-contracting these two rules (to remove @{text P}) haftmann@30531: causes them to be ignored because of their interaction with haftmann@30531: congruence rules. haftmann@30531: *} haftmann@30531: haftmann@30531: lemma ball_UNIV [simp]: "Ball UNIV P = All P" haftmann@30531: by (simp add: Ball_def) haftmann@30531: haftmann@30531: lemma bex_UNIV [simp]: "Bex UNIV P = Ex P" haftmann@30531: by (simp add: Bex_def) haftmann@30531: haftmann@30531: lemma UNIV_eq_I: "(\x. x \ A) \ UNIV = A" haftmann@30531: by auto haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* The empty set *} haftmann@30531: haftmann@30531: lemma empty_iff [simp]: "(c : {}) = False" haftmann@30531: by (simp add: empty_def) haftmann@30531: haftmann@30531: lemma emptyE [elim!]: "a : {} ==> P" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma empty_subsetI [iff]: "{} \ A" haftmann@30531: -- {* One effect is to delete the ASSUMPTION @{prop "{} <= A"} *} haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma equals0I: "(!!y. y \ A ==> False) ==> A = {}" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma equals0D: "A = {} ==> a \ A" haftmann@30531: -- {* Use for reasoning about disjointness: @{prop "A Int B = {}"} *} haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma ball_empty [simp]: "Ball {} P = True" haftmann@30531: by (simp add: Ball_def) haftmann@30531: haftmann@30531: lemma bex_empty [simp]: "Bex {} P = False" haftmann@30531: by (simp add: Bex_def) haftmann@30531: haftmann@30531: lemma UNIV_not_empty [iff]: "UNIV ~= {}" haftmann@30531: by (blast elim: equalityE) haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* The Powerset operator -- Pow *} haftmann@30531: haftmann@30531: lemma Pow_iff [iff]: "(A \ Pow B) = (A \ B)" haftmann@30531: by (simp add: Pow_def) haftmann@30531: haftmann@30531: lemma PowI: "A \ B ==> A \ Pow B" haftmann@30531: by (simp add: Pow_def) haftmann@30531: haftmann@30531: lemma PowD: "A \ Pow B ==> A \ B" haftmann@30531: by (simp add: Pow_def) haftmann@30531: haftmann@30531: lemma Pow_bottom: "{} \ Pow B" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma Pow_top: "A \ Pow A" haftmann@30531: by (simp add: subset_refl) haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* Set complement *} haftmann@30531: haftmann@30531: lemma Compl_iff [simp]: "(c \ -A) = (c \ A)" haftmann@30531: by (simp add: mem_def fun_Compl_def bool_Compl_def) haftmann@30531: haftmann@30531: lemma ComplI [intro!]: "(c \ A ==> False) ==> c \ -A" haftmann@30531: by (unfold mem_def fun_Compl_def bool_Compl_def) blast clasohm@923: wenzelm@11979: text {* haftmann@30531: \medskip This form, with negated conclusion, works well with the haftmann@30531: Classical prover. Negated assumptions behave like formulae on the haftmann@30531: right side of the notional turnstile ... *} haftmann@30531: haftmann@30531: lemma ComplD [dest!]: "c : -A ==> c~:A" haftmann@30531: by (simp add: mem_def fun_Compl_def bool_Compl_def) haftmann@30531: haftmann@30531: lemmas ComplE = ComplD [elim_format] haftmann@30531: haftmann@30531: lemma Compl_eq: "- A = {x. ~ x : A}" by blast haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* Binary union -- Un *} haftmann@30531: haftmann@30531: lemma Un_iff [simp]: "(c : A Un B) = (c:A | c:B)" haftmann@30531: by (unfold Un_def) blast haftmann@30531: haftmann@30531: lemma UnI1 [elim?]: "c:A ==> c : A Un B" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma UnI2 [elim?]: "c:B ==> c : A Un B" haftmann@30531: by simp haftmann@30531: haftmann@30531: text {* haftmann@30531: \medskip Classical introduction rule: no commitment to @{prop A} vs haftmann@30531: @{prop B}. wenzelm@11979: *} wenzelm@11979: haftmann@30531: lemma UnCI [intro!]: "(c~:B ==> c:A) ==> c : A Un B" haftmann@30531: by auto haftmann@30531: haftmann@30531: lemma UnE [elim!]: "c : A Un B ==> (c:A ==> P) ==> (c:B ==> P) ==> P" haftmann@30531: by (unfold Un_def) blast haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* Binary intersection -- Int *} haftmann@30531: haftmann@30531: lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)" haftmann@30531: by (unfold Int_def) blast haftmann@30531: haftmann@30531: lemma IntI [intro!]: "c:A ==> c:B ==> c : A Int B" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma IntD1: "c : A Int B ==> c:A" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma IntD2: "c : A Int B ==> c:B" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P" haftmann@30531: by simp haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* Set difference *} haftmann@30531: haftmann@30531: lemma Diff_iff [simp]: "(c : A - B) = (c:A & c~:B)" haftmann@30531: by (simp add: mem_def fun_diff_def bool_diff_def) haftmann@30531: haftmann@30531: lemma DiffI [intro!]: "c : A ==> c ~: B ==> c : A - B" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma DiffD1: "c : A - B ==> c : A" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma DiffD2: "c : A - B ==> c : B ==> P" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma DiffE [elim!]: "c : A - B ==> (c:A ==> c~:B ==> P) ==> P" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma set_diff_eq: "A - B = {x. x : A & ~ x : B}" by blast haftmann@30531: haftmann@30531: lemma Compl_eq_Diff_UNIV: "-A = (UNIV - A)" haftmann@30531: by blast haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* Augmenting a set -- insert *} haftmann@30531: haftmann@30531: lemma insert_iff [simp]: "(a : insert b A) = (a = b | a:A)" haftmann@30531: by (unfold insert_def) blast haftmann@30531: haftmann@30531: lemma insertI1: "a : insert a B" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma insertI2: "a : B ==> a : insert b B" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma insertE [elim!]: "a : insert b A ==> (a = b ==> P) ==> (a:A ==> P) ==> P" haftmann@30531: by (unfold insert_def) blast haftmann@30531: haftmann@30531: lemma insertCI [intro!]: "(a~:B ==> a = b) ==> a: insert b B" haftmann@30531: -- {* Classical introduction rule. *} haftmann@30531: by auto haftmann@30531: haftmann@30531: lemma subset_insert_iff: "(A \ insert x B) = (if x:A then A - {x} \ B else A \ B)" haftmann@30531: by auto haftmann@30531: haftmann@30531: lemma set_insert: haftmann@30531: assumes "x \ A" haftmann@30531: obtains B where "A = insert x B" and "x \ B" haftmann@30531: proof haftmann@30531: from assms show "A = insert x (A - {x})" by blast haftmann@30531: next haftmann@30531: show "x \ A - {x}" by blast haftmann@30531: qed haftmann@30531: haftmann@30531: lemma insert_ident: "x ~: A ==> x ~: B ==> (insert x A = insert x B) = (A = B)" haftmann@30531: by auto haftmann@30531: haftmann@30531: subsubsection {* Singletons, using insert *} haftmann@30531: haftmann@30531: lemma singletonI [intro!,noatp]: "a : {a}" haftmann@30531: -- {* Redundant? But unlike @{text insertCI}, it proves the subgoal immediately! *} haftmann@30531: by (rule insertI1) haftmann@30531: haftmann@30531: lemma singletonD [dest!,noatp]: "b : {a} ==> b = a" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemmas singletonE = singletonD [elim_format] haftmann@30531: haftmann@30531: lemma singleton_iff: "(b : {a}) = (b = a)" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma singleton_inject [dest!]: "{a} = {b} ==> a = b" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma singleton_insert_inj_eq [iff,noatp]: haftmann@30531: "({b} = insert a A) = (a = b & A \ {b})" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma singleton_insert_inj_eq' [iff,noatp]: haftmann@30531: "(insert a A = {b}) = (a = b & A \ {b})" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma subset_singletonD: "A \ {x} ==> A = {} | A = {x}" haftmann@30531: by fast haftmann@30531: haftmann@30531: lemma singleton_conv [simp]: "{x. x = a} = {a}" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma singleton_conv2 [simp]: "{x. a = x} = {a}" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma diff_single_insert: "A - {x} \ B ==> x \ A ==> A \ insert x B" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma doubleton_eq_iff: "({a,b} = {c,d}) = (a=c & b=d | a=d & b=c)" haftmann@30531: by (blast elim: equalityE) haftmann@30531: wenzelm@11979: wenzelm@11979: subsubsection {* Unions of families *} wenzelm@11979: wenzelm@11979: text {* wenzelm@11979: @{term [source] "UN x:A. B x"} is @{term "Union (B`A)"}. wenzelm@11979: *} wenzelm@11979: paulson@24286: declare UNION_def [noatp] paulson@24286: wenzelm@11979: lemma UN_iff [simp]: "(b: (UN x:A. B x)) = (EX x:A. b: B x)" wenzelm@11979: by (unfold UNION_def) blast wenzelm@11979: wenzelm@11979: lemma UN_I [intro]: "a:A ==> b: B a ==> b: (UN x:A. B x)" wenzelm@11979: -- {* The order of the premises presupposes that @{term A} is rigid; wenzelm@11979: @{term b} may be flexible. *} wenzelm@11979: by auto wenzelm@11979: wenzelm@11979: lemma UN_E [elim!]: "b : (UN x:A. B x) ==> (!!x. x:A ==> b: B x ==> R) ==> R" wenzelm@11979: by (unfold UNION_def) blast clasohm@923: wenzelm@11979: lemma UN_cong [cong]: wenzelm@11979: "A = B ==> (!!x. x:B ==> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" wenzelm@11979: by (simp add: UNION_def) wenzelm@11979: berghofe@29691: lemma strong_UN_cong: berghofe@29691: "A = B ==> (!!x. x:B =simp=> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" berghofe@29691: by (simp add: UNION_def simp_implies_def) berghofe@29691: wenzelm@11979: wenzelm@11979: subsubsection {* Intersections of families *} wenzelm@11979: wenzelm@11979: text {* @{term [source] "INT x:A. B x"} is @{term "Inter (B`A)"}. *} wenzelm@11979: wenzelm@11979: lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)" wenzelm@11979: by (unfold INTER_def) blast clasohm@923: wenzelm@11979: lemma INT_I [intro!]: "(!!x. x:A ==> b: B x) ==> b : (INT x:A. B x)" wenzelm@11979: by (unfold INTER_def) blast wenzelm@11979: wenzelm@11979: lemma INT_D [elim]: "b : (INT x:A. B x) ==> a:A ==> b: B a" wenzelm@11979: by auto wenzelm@11979: wenzelm@11979: lemma INT_E [elim]: "b : (INT x:A. B x) ==> (b: B a ==> R) ==> (a~:A ==> R) ==> R" wenzelm@11979: -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a:A"}. *} wenzelm@11979: by (unfold INTER_def) blast wenzelm@11979: wenzelm@11979: lemma INT_cong [cong]: wenzelm@11979: "A = B ==> (!!x. x:B ==> C x = D x) ==> (INT x:A. C x) = (INT x:B. D x)" wenzelm@11979: by (simp add: INTER_def) wenzelm@7238: clasohm@923: wenzelm@11979: subsubsection {* Union *} wenzelm@11979: paulson@24286: lemma Union_iff [simp,noatp]: "(A : Union C) = (EX X:C. A:X)" wenzelm@11979: by (unfold Union_def) blast wenzelm@11979: wenzelm@11979: lemma UnionI [intro]: "X:C ==> A:X ==> A : Union C" wenzelm@11979: -- {* The order of the premises presupposes that @{term C} is rigid; wenzelm@11979: @{term A} may be flexible. *} wenzelm@11979: by auto wenzelm@11979: wenzelm@11979: lemma UnionE [elim!]: "A : Union C ==> (!!X. A:X ==> X:C ==> R) ==> R" wenzelm@11979: by (unfold Union_def) blast wenzelm@11979: wenzelm@11979: wenzelm@11979: subsubsection {* Inter *} wenzelm@11979: paulson@24286: lemma Inter_iff [simp,noatp]: "(A : Inter C) = (ALL X:C. A:X)" wenzelm@11979: by (unfold Inter_def) blast wenzelm@11979: wenzelm@11979: lemma InterI [intro!]: "(!!X. X:C ==> A:X) ==> A : Inter C" wenzelm@11979: by (simp add: Inter_def) wenzelm@11979: wenzelm@11979: text {* wenzelm@11979: \medskip A ``destruct'' rule -- every @{term X} in @{term C} wenzelm@11979: contains @{term A} as an element, but @{prop "A:X"} can hold when wenzelm@11979: @{prop "X:C"} does not! This rule is analogous to @{text spec}. wenzelm@11979: *} wenzelm@11979: wenzelm@11979: lemma InterD [elim]: "A : Inter C ==> X:C ==> A:X" wenzelm@11979: by auto wenzelm@11979: wenzelm@11979: lemma InterE [elim]: "A : Inter C ==> (X~:C ==> R) ==> (A:X ==> R) ==> R" wenzelm@11979: -- {* ``Classical'' elimination rule -- does not require proving wenzelm@11979: @{prop "X:C"}. *} wenzelm@11979: by (unfold Inter_def) blast wenzelm@11979: haftmann@30531: text {* haftmann@30531: \medskip Image of a set under a function. Frequently @{term b} does haftmann@30531: not have the syntactic form of @{term "f x"}. haftmann@30531: *} haftmann@30531: haftmann@30531: declare image_def [noatp] haftmann@30531: haftmann@30531: lemma image_eqI [simp, intro]: "b = f x ==> x:A ==> b : f`A" haftmann@30531: by (unfold image_def) blast haftmann@30531: haftmann@30531: lemma imageI: "x : A ==> f x : f ` A" haftmann@30531: by (rule image_eqI) (rule refl) haftmann@30531: haftmann@30531: lemma rev_image_eqI: "x:A ==> b = f x ==> b : f`A" haftmann@30531: -- {* This version's more effective when we already have the haftmann@30531: required @{term x}. *} haftmann@30531: by (unfold image_def) blast haftmann@30531: haftmann@30531: lemma imageE [elim!]: haftmann@30531: "b : (%x. f x)`A ==> (!!x. b = f x ==> x:A ==> P) ==> P" haftmann@30531: -- {* The eta-expansion gives variable-name preservation. *} haftmann@30531: by (unfold image_def) blast haftmann@30531: haftmann@30531: lemma image_Un: "f`(A Un B) = f`A Un f`B" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma image_eq_UN: "f`A = (UN x:A. {f x})" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma image_iff: "(z : f`A) = (EX x:A. z = f x)" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma image_subset_iff: "(f`A \ B) = (\x\A. f x \ B)" haftmann@30531: -- {* This rewrite rule would confuse users if made default. *} haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma subset_image_iff: "(B \ f`A) = (EX AA. AA \ A & B = f`AA)" haftmann@30531: apply safe haftmann@30531: prefer 2 apply fast haftmann@30531: apply (rule_tac x = "{a. a : A & f a : B}" in exI, fast) haftmann@30531: done haftmann@30531: haftmann@30531: lemma image_subsetI: "(!!x. x \ A ==> f x \ B) ==> f`A \ B" haftmann@30531: -- {* Replaces the three steps @{text subsetI}, @{text imageE}, haftmann@30531: @{text hypsubst}, but breaks too many existing proofs. *} haftmann@30531: by blast haftmann@30531: haftmann@30531: text {* haftmann@30531: \medskip Range of a function -- just a translation for image! haftmann@30531: *} haftmann@30531: haftmann@30531: lemma range_eqI: "b = f x ==> b \ range f" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma rangeI: "f x \ range f" haftmann@30531: by simp haftmann@30531: haftmann@30531: lemma rangeE [elim?]: "b \ range (\x. f x) ==> (!!x. b = f x ==> P) ==> P" haftmann@30531: by blast haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* Set reasoning tools *} haftmann@30531: nipkow@31166: text{* Elimination of @{text"{x. \ & x=t & \}"}. *} nipkow@31166: nipkow@31166: lemma singleton_conj_conv[simp]: "{x. x=a & P x} = (if P a then {a} else {})" nipkow@31166: by auto nipkow@31166: nipkow@31166: lemma singleton_conj_conv2[simp]: "{x. a=x & P x} = (if P a then {a} else {})" nipkow@31166: by auto nipkow@31166: nipkow@31166: ML{* nipkow@31166: local nipkow@31166: val Coll_perm_tac = rtac @{thm Collect_cong} 1 THEN rtac @{thm iffI} 1 THEN nipkow@31166: ALLGOALS(EVERY'[REPEAT_DETERM o (etac @{thm conjE}), nipkow@31166: DEPTH_SOLVE_1 o (ares_tac [@{thm conjI}])]) nipkow@31166: in nipkow@31166: val defColl_regroup = Simplifier.simproc (the_context ()) nipkow@31166: "defined Collect" ["{x. P x & Q x}"] nipkow@31166: (Quantifier1.rearrange_Coll Coll_perm_tac) nipkow@31166: end; nipkow@31166: nipkow@31166: Addsimprocs [defColl_regroup]; nipkow@31166: nipkow@31166: *} nipkow@31166: haftmann@30531: text {* haftmann@30531: Rewrite rules for boolean case-splitting: faster than @{text haftmann@30531: "split_if [split]"}. haftmann@30531: *} haftmann@30531: haftmann@30531: lemma split_if_eq1: "((if Q then x else y) = b) = ((Q --> x = b) & (~ Q --> y = b))" haftmann@30531: by (rule split_if) haftmann@30531: haftmann@30531: lemma split_if_eq2: "(a = (if Q then x else y)) = ((Q --> a = x) & (~ Q --> a = y))" haftmann@30531: by (rule split_if) haftmann@30531: haftmann@30531: text {* haftmann@30531: Split ifs on either side of the membership relation. Not for @{text haftmann@30531: "[simp]"} -- can cause goals to blow up! haftmann@30531: *} haftmann@30531: haftmann@30531: lemma split_if_mem1: "((if Q then x else y) : b) = ((Q --> x : b) & (~ Q --> y : b))" haftmann@30531: by (rule split_if) haftmann@30531: haftmann@30531: lemma split_if_mem2: "(a : (if Q then x else y)) = ((Q --> a : x) & (~ Q --> a : y))" haftmann@30531: by (rule split_if [where P="%S. a : S"]) haftmann@30531: haftmann@30531: lemmas split_ifs = if_bool_eq_conj split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2 haftmann@30531: haftmann@30531: (*Would like to add these, but the existing code only searches for the haftmann@30531: outer-level constant, which in this case is just "op :"; we instead need haftmann@30531: to use term-nets to associate patterns with rules. Also, if a rule fails to haftmann@30531: apply, then the formula should be kept. haftmann@30531: [("HOL.uminus", Compl_iff RS iffD1), ("HOL.minus", [Diff_iff RS iffD1]), haftmann@30531: ("Int", [IntD1,IntD2]), haftmann@30531: ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])] haftmann@30531: *) haftmann@30531: haftmann@30531: ML {* haftmann@30531: val mksimps_pairs = [(@{const_name Ball}, @{thms bspec})] @ mksimps_pairs; haftmann@30531: *} haftmann@30531: declaration {* fn _ => haftmann@30531: Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) haftmann@30531: *} haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* The ``proper subset'' relation *} haftmann@30531: haftmann@30531: lemma psubsetI [intro!,noatp]: "A \ B ==> A \ B ==> A \ B" haftmann@30531: by (unfold less_le) blast haftmann@30531: haftmann@30531: lemma psubsetE [elim!,noatp]: haftmann@30531: "[|A \ B; [|A \ B; ~ (B\A)|] ==> R|] ==> R" haftmann@30531: by (unfold less_le) blast haftmann@30531: haftmann@30531: lemma psubset_insert_iff: haftmann@30531: "(A \ insert x B) = (if x \ B then A \ B else if x \ A then A - {x} \ B else A \ B)" haftmann@30531: by (auto simp add: less_le subset_insert_iff) haftmann@30531: haftmann@30531: lemma psubset_eq: "(A \ B) = (A \ B & A \ B)" haftmann@30531: by (simp only: less_le) haftmann@30531: haftmann@30531: lemma psubset_imp_subset: "A \ B ==> A \ B" haftmann@30531: by (simp add: psubset_eq) haftmann@30531: haftmann@30531: lemma psubset_trans: "[| A \ B; B \ C |] ==> A \ C" haftmann@30531: apply (unfold less_le) haftmann@30531: apply (auto dest: subset_antisym) haftmann@30531: done haftmann@30531: haftmann@30531: lemma psubsetD: "[| A \ B; c \ A |] ==> c \ B" haftmann@30531: apply (unfold less_le) haftmann@30531: apply (auto dest: subsetD) haftmann@30531: done haftmann@30531: haftmann@30531: lemma psubset_subset_trans: "A \ B ==> B \ C ==> A \ C" haftmann@30531: by (auto simp add: psubset_eq) haftmann@30531: haftmann@30531: lemma subset_psubset_trans: "A \ B ==> B \ C ==> A \ C" haftmann@30531: by (auto simp add: psubset_eq) haftmann@30531: haftmann@30531: lemma psubset_imp_ex_mem: "A \ B ==> \b. b \ (B - A)" haftmann@30531: by (unfold less_le) blast haftmann@30531: haftmann@30531: lemma atomize_ball: haftmann@30531: "(!!x. x \ A ==> P x) == Trueprop (\x\A. P x)" haftmann@30531: by (simp only: Ball_def atomize_all atomize_imp) haftmann@30531: haftmann@30531: lemmas [symmetric, rulify] = atomize_ball haftmann@30531: and [symmetric, defn] = atomize_ball haftmann@30531: haftmann@30531: haftmann@30531: subsection {* Further set-theory lemmas *} haftmann@30531: haftmann@30531: subsubsection {* Derived rules involving subsets. *} haftmann@30531: haftmann@30531: text {* @{text insert}. *} haftmann@30531: haftmann@30531: lemma subset_insertI: "B \ insert a B" haftmann@30531: by (rule subsetI) (erule insertI2) haftmann@30531: haftmann@30531: lemma subset_insertI2: "A \ B \ A \ insert b B" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma subset_insert: "x \ A ==> (A \ insert x B) = (A \ B)" haftmann@30531: by blast wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip Big Union -- least upper bound of a set. *} wenzelm@12897: wenzelm@12897: lemma Union_upper: "B \ A ==> B \ Union A" nipkow@17589: by (iprover intro: subsetI UnionI) wenzelm@12897: wenzelm@12897: lemma Union_least: "(!!X. X \ A ==> X \ C) ==> Union A \ C" nipkow@17589: by (iprover intro: subsetI elim: UnionE dest: subsetD) wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip General union. *} wenzelm@12897: wenzelm@12897: lemma UN_upper: "a \ A ==> B a \ (\x\A. B x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_least: "(!!x. x \ A ==> B x \ C) ==> (\x\A. B x) \ C" nipkow@17589: by (iprover intro: subsetI elim: UN_E dest: subsetD) wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip Big Intersection -- greatest lower bound of a set. *} wenzelm@12897: wenzelm@12897: lemma Inter_lower: "B \ A ==> Inter A \ B" wenzelm@12897: by blast wenzelm@12897: ballarin@14551: lemma Inter_subset: ballarin@14551: "[| !!X. X \ A ==> X \ B; A ~= {} |] ==> \A \ B" ballarin@14551: by blast ballarin@14551: wenzelm@12897: lemma Inter_greatest: "(!!X. X \ A ==> C \ X) ==> C \ Inter A" nipkow@17589: by (iprover intro: InterI subsetI dest: subsetD) wenzelm@12897: wenzelm@12897: lemma INT_lower: "a \ A ==> (\x\A. B x) \ B a" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma INT_greatest: "(!!x. x \ A ==> C \ B x) ==> C \ (\x\A. B x)" nipkow@17589: by (iprover intro: INT_I subsetI dest: subsetD) wenzelm@12897: haftmann@30531: haftmann@30531: text {* \medskip Finite Union -- the least upper bound of two sets. *} haftmann@30531: haftmann@30531: lemma Un_upper1: "A \ A \ B" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma Un_upper2: "B \ A \ B" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma Un_least: "A \ C ==> B \ C ==> A \ B \ C" haftmann@30531: by blast haftmann@30531: haftmann@30531: haftmann@30531: text {* \medskip Finite Intersection -- the greatest lower bound of two sets. *} haftmann@30531: haftmann@30531: lemma Int_lower1: "A \ B \ A" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma Int_lower2: "A \ B \ B" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma Int_greatest: "C \ A ==> C \ B ==> C \ A \ B" haftmann@30531: by blast haftmann@30531: haftmann@30531: haftmann@30531: text {* \medskip Set difference. *} haftmann@30531: haftmann@30531: lemma Diff_subset: "A - B \ A" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma Diff_subset_conv: "(A - B \ C) = (A \ B \ C)" haftmann@30531: by blast haftmann@30531: haftmann@30531: haftmann@30531: subsubsection {* Equalities involving union, intersection, inclusion, etc. *} haftmann@30531: haftmann@30531: text {* @{text "{}"}. *} haftmann@30531: haftmann@30531: lemma Collect_const [simp]: "{s. P} = (if P then UNIV else {})" haftmann@30531: -- {* supersedes @{text "Collect_False_empty"} *} haftmann@30531: by auto haftmann@30531: haftmann@30531: lemma subset_empty [simp]: "(A \ {}) = (A = {})" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma not_psubset_empty [iff]: "\ (A < {})" haftmann@30531: by (unfold less_le) blast haftmann@30531: haftmann@30531: lemma Collect_empty_eq [simp]: "(Collect P = {}) = (\x. \ P x)" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma empty_Collect_eq [simp]: "({} = Collect P) = (\x. \ P x)" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma Collect_neg_eq: "{x. \ P x} = - {x. P x}" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma Collect_disj_eq: "{x. P x | Q x} = {x. P x} \ {x. Q x}" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma Collect_imp_eq: "{x. P x --> Q x} = -{x. P x} \ {x. Q x}" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma Collect_conj_eq: "{x. P x & Q x} = {x. P x} \ {x. Q x}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Collect_all_eq: "{x. \y. P x y} = (\y. {x. P x y})" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" wenzelm@12897: by blast wenzelm@12897: paulson@24286: lemma Collect_ex_eq [noatp]: "{x. \y. P x y} = (\y. {x. P x y})" wenzelm@12897: by blast wenzelm@12897: paulson@24286: lemma Collect_bex_eq [noatp]: "{x. \y\A. P x y} = (\y\A. {x. P x y})" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: haftmann@30531: text {* \medskip @{text insert}. *} haftmann@30531: haftmann@30531: lemma insert_is_Un: "insert a A = {a} Un A" haftmann@30531: -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a {}"} *} haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma insert_not_empty [simp]: "insert a A \ {}" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemmas empty_not_insert = insert_not_empty [symmetric, standard] haftmann@30531: declare empty_not_insert [simp] haftmann@30531: haftmann@30531: lemma insert_absorb: "a \ A ==> insert a A = A" haftmann@30531: -- {* @{text "[simp]"} causes recursive calls when there are nested inserts *} haftmann@30531: -- {* with \emph{quadratic} running time *} haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma insert_absorb2 [simp]: "insert x (insert x A) = insert x A" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma insert_commute: "insert x (insert y A) = insert y (insert x A)" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma insert_subset [simp]: "(insert x A \ B) = (x \ B & A \ B)" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma mk_disjoint_insert: "a \ A ==> \B. A = insert a B & a \ B" haftmann@30531: -- {* use new @{text B} rather than @{text "A - {a}"} to avoid infinite unfolding *} haftmann@30531: apply (rule_tac x = "A - {a}" in exI, blast) haftmann@30531: done haftmann@30531: haftmann@30531: lemma insert_Collect: "insert a (Collect P) = {u. u \ a --> P u}" haftmann@30531: by auto haftmann@30531: haftmann@30531: lemma UN_insert_distrib: "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma insert_inter_insert[simp]: "insert a A \ insert a B = insert a (A \ B)" mehta@14742: by blast nipkow@14302: haftmann@30531: lemma insert_disjoint [simp,noatp]: haftmann@30531: "(insert a A \ B = {}) = (a \ B \ A \ B = {})" haftmann@30531: "({} = insert a A \ B) = (a \ B \ {} = A \ B)" haftmann@30531: by auto haftmann@30531: haftmann@30531: lemma disjoint_insert [simp,noatp]: haftmann@30531: "(B \ insert a A = {}) = (a \ B \ B \ A = {})" haftmann@30531: "({} = A \ insert b B) = (b \ A \ {} = A \ B)" haftmann@30531: by auto haftmann@30531: haftmann@30531: text {* \medskip @{text image}. *} haftmann@30531: haftmann@30531: lemma image_empty [simp]: "f`{} = {}" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma image_insert [simp]: "f ` insert a B = insert (f a) (f`B)" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma image_constant: "x \ A ==> (\x. c) ` A = {c}" haftmann@30531: by auto haftmann@30531: haftmann@30531: lemma image_constant_conv: "(%x. c) ` A = (if A = {} then {} else {c})" haftmann@30531: by auto haftmann@30531: haftmann@30531: lemma image_image: "f ` (g ` A) = (\x. f (g x)) ` A" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma insert_image [simp]: "x \ A ==> insert (f x) (f`A) = f`A" haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma image_is_empty [iff]: "(f`A = {}) = (A = {})" haftmann@30531: by blast haftmann@30531: haftmann@30531: haftmann@30531: lemma image_Collect [noatp]: "f ` {x. P x} = {f x | x. P x}" haftmann@30531: -- {* NOT suitable as a default simprule: the RHS isn't simpler than the LHS, haftmann@30531: with its implicit quantifier and conjunction. Also image enjoys better haftmann@30531: equational properties than does the RHS. *} haftmann@30531: by blast haftmann@30531: haftmann@30531: lemma if_image_distrib [simp]: haftmann@30531: "(\x. if P x then f x else g x) ` S haftmann@30531: = (f ` (S \ {x. P x})) \ (g ` (S \ {x. \ P x}))" haftmann@30531: by (auto simp add: image_def) haftmann@30531: haftmann@30531: lemma image_cong: "M = N ==> (!!x. x \ N ==> f x = g x) ==> f`M = g`N" haftmann@30531: by (simp add: image_def) haftmann@30531: haftmann@30531: haftmann@30531: text {* \medskip @{text range}. *} haftmann@30531: paulson@24286: lemma full_SetCompr_eq [noatp]: "{u. \x. u = f x} = range f" wenzelm@12897: by auto wenzelm@12897: huffman@27418: lemma range_composition: "range (\x. f (g x)) = f`range g" paulson@14208: by (subst image_image, simp) wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip @{text Int} *} wenzelm@12897: wenzelm@12897: lemma Int_absorb [simp]: "A \ A = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_left_absorb: "A \ (A \ B) = A \ B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_commute: "A \ B = B \ A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_left_commute: "A \ (B \ C) = B \ (A \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_assoc: "(A \ B) \ C = A \ (B \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemmas Int_ac = Int_assoc Int_left_absorb Int_commute Int_left_commute wenzelm@12897: -- {* Intersection is an AC-operator *} wenzelm@12897: wenzelm@12897: lemma Int_absorb1: "B \ A ==> A \ B = B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_absorb2: "A \ B ==> A \ B = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_empty_left [simp]: "{} \ B = {}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_empty_right [simp]: "A \ {} = {}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma disjoint_eq_subset_Compl: "(A \ B = {}) = (A \ -B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma disjoint_iff_not_equal: "(A \ B = {}) = (\x\A. \y\B. x \ y)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_UNIV_left [simp]: "UNIV \ B = B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_UNIV_right [simp]: "A \ UNIV = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_eq_Inter: "A \ B = \{A, B}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_Un_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_Un_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" wenzelm@12897: by blast wenzelm@12897: paulson@24286: lemma Int_UNIV [simp,noatp]: "(A \ B = UNIV) = (A = UNIV & B = UNIV)" wenzelm@12897: by blast wenzelm@12897: paulson@15102: lemma Int_subset_iff [simp]: "(C \ A \ B) = (C \ A & C \ B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_Collect: "(x \ A \ {x. P x}) = (x \ A & P x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip @{text Un}. *} wenzelm@12897: wenzelm@12897: lemma Un_absorb [simp]: "A \ A = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_left_absorb: "A \ (A \ B) = A \ B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_commute: "A \ B = B \ A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_left_commute: "A \ (B \ C) = B \ (A \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_assoc: "(A \ B) \ C = A \ (B \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute wenzelm@12897: -- {* Union is an AC-operator *} wenzelm@12897: wenzelm@12897: lemma Un_absorb1: "A \ B ==> A \ B = B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_absorb2: "B \ A ==> A \ B = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_empty_left [simp]: "{} \ B = B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_empty_right [simp]: "A \ {} = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_UNIV_left [simp]: "UNIV \ B = UNIV" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_UNIV_right [simp]: "A \ UNIV = UNIV" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_eq_Union: "A \ B = \{A, B}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_insert_left [simp]: "(insert a B) \ C = insert a (B \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_insert_right [simp]: "A \ (insert a B) = insert a (A \ B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_insert_left: wenzelm@12897: "(insert a B) Int C = (if a \ C then insert a (B \ C) else B \ C)" wenzelm@12897: by auto wenzelm@12897: wenzelm@12897: lemma Int_insert_right: wenzelm@12897: "A \ (insert a B) = (if a \ A then insert a (A \ B) else A \ B)" wenzelm@12897: by auto wenzelm@12897: wenzelm@12897: lemma Un_Int_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Int_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Int_crazy: wenzelm@12897: "(A \ B) \ (B \ C) \ (C \ A) = (A \ B) \ (B \ C) \ (C \ A)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma subset_Un_eq: "(A \ B) = (A \ B = B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_empty [iff]: "(A \ B = {}) = (A = {} & B = {})" wenzelm@12897: by blast paulson@15102: paulson@15102: lemma Un_subset_iff [simp]: "(A \ B \ C) = (A \ C & B \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Diff_Int: "(A - B) \ (A \ B) = A" wenzelm@12897: by blast wenzelm@12897: paulson@22172: lemma Diff_Int2: "A \ C - B \ C = A \ C - B" paulson@22172: by blast paulson@22172: wenzelm@12897: wenzelm@12897: text {* \medskip Set complement *} wenzelm@12897: wenzelm@12897: lemma Compl_disjoint [simp]: "A \ -A = {}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Compl_disjoint2 [simp]: "-A \ A = {}" wenzelm@12897: by blast wenzelm@12897: paulson@13818: lemma Compl_partition: "A \ -A = UNIV" paulson@13818: by blast paulson@13818: paulson@13818: lemma Compl_partition2: "-A \ A = UNIV" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma double_complement [simp]: "- (-A) = (A::'a set)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Compl_Un [simp]: "-(A \ B) = (-A) \ (-B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Compl_Int [simp]: "-(A \ B) = (-A) \ (-B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Compl_UN [simp]: "-(\x\A. B x) = (\x\A. -B x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Compl_INT [simp]: "-(\x\A. B x) = (\x\A. -B x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma subset_Compl_self_eq: "(A \ -A) = (A = {})" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Int_assoc_eq: "((A \ B) \ C = A \ (B \ C)) = (C \ A)" wenzelm@12897: -- {* Halmos, Naive Set Theory, page 16. *} wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Compl_UNIV_eq [simp]: "-UNIV = {}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Compl_empty_eq [simp]: "-{} = UNIV" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Compl_subset_Compl_iff [iff]: "(-A \ -B) = (B \ A)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Compl_eq_Compl_iff [iff]: "(-A = -B) = (A = (B::'a set))" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip @{text Union}. *} wenzelm@12897: wenzelm@12897: lemma Union_empty [simp]: "Union({}) = {}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Union_UNIV [simp]: "Union UNIV = UNIV" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Union_insert [simp]: "Union (insert a B) = a \ \B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Union_Un_distrib [simp]: "\(A Un B) = \A \ \B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Union_Int_subset: "\(A \ B) \ \A \ \B" wenzelm@12897: by blast wenzelm@12897: paulson@24286: lemma Union_empty_conv [simp,noatp]: "(\A = {}) = (\x\A. x = {})" nipkow@13653: by blast nipkow@13653: paulson@24286: lemma empty_Union_conv [simp,noatp]: "({} = \A) = (\x\A. x = {})" nipkow@13653: by blast wenzelm@12897: wenzelm@12897: lemma Union_disjoint: "(\C \ A = {}) = (\B\C. B \ A = {})" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip @{text Inter}. *} wenzelm@12897: wenzelm@12897: lemma Inter_empty [simp]: "\{} = UNIV" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Inter_UNIV [simp]: "\UNIV = {}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Inter_insert [simp]: "\(insert a B) = a \ \B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" wenzelm@12897: by blast wenzelm@12897: paulson@24286: lemma Inter_UNIV_conv [simp,noatp]: nipkow@13653: "(\A = UNIV) = (\x\A. x = UNIV)" nipkow@13653: "(UNIV = \A) = (\x\A. x = UNIV)" paulson@14208: by blast+ nipkow@13653: wenzelm@12897: wenzelm@12897: text {* wenzelm@12897: \medskip @{text UN} and @{text INT}. wenzelm@12897: wenzelm@12897: Basic identities: *} wenzelm@12897: paulson@24286: lemma UN_empty [simp,noatp]: "(\x\{}. B x) = {}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_empty2 [simp]: "(\x\A. {}) = {}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_singleton [simp]: "(\x\A. {x}) = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_absorb: "k \ I ==> A k \ (\i\I. A i) = (\i\I. A i)" paulson@15102: by auto wenzelm@12897: wenzelm@12897: lemma INT_empty [simp]: "(\x\{}. B x) = UNIV" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma INT_absorb: "k \ I ==> A k \ (\i\I. A i) = (\i\I. A i)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" wenzelm@12897: by blast wenzelm@12897: nipkow@24331: lemma UN_Un[simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_subset_iff: "((\i\I. A i) \ B) = (\i\I. A i \ B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma INT_subset_iff: "(B \ (\i\I. A i)) = (\i\I. B \ A i)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ INTER A B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma INT_insert_distrib: wenzelm@12897: "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Union_image_eq [simp]: "\(B`A) = (\x\A. B x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma image_Union: "f ` \S = (\x\S. f ` x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Inter_image_eq [simp]: "\(B`A) = (\x\A. B x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" wenzelm@12897: by auto wenzelm@12897: wenzelm@12897: lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" wenzelm@12897: by auto wenzelm@12897: wenzelm@12897: lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" wenzelm@12897: -- {* Look: it has an \emph{existential} quantifier *} wenzelm@12897: by blast wenzelm@12897: paulson@18447: lemma UNION_empty_conv[simp]: nipkow@13653: "({} = (UN x:A. B x)) = (\x\A. B x = {})" nipkow@13653: "((UN x:A. B x) = {}) = (\x\A. B x = {})" nipkow@13653: by blast+ nipkow@13653: paulson@18447: lemma INTER_UNIV_conv[simp]: nipkow@13653: "(UNIV = (INT x:A. B x)) = (\x\A. B x = UNIV)" nipkow@13653: "((INT x:A. B x) = UNIV) = (\x\A. B x = UNIV)" nipkow@13653: by blast+ wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip Distributive laws: *} wenzelm@12897: wenzelm@12897: lemma Int_Union: "A \ \B = (\C\B. A \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_Union2: "\B \ A = (\C\B. C \ A)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Union_image: "(\x\C. A x \ B x) = \(A`C) \ \(B`C)" wenzelm@12897: -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *} wenzelm@12897: -- {* Union of a family of unions *} wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_Un_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" wenzelm@12897: -- {* Equivalent version *} wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Inter: "A \ \B = (\C\B. A \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A`C) \ \(B`C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma INT_Int_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" wenzelm@12897: -- {* Equivalent version *} wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_UN_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" wenzelm@12897: -- {* Halmos, Naive Set Theory, page 35. *} wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_INT_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_UN_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_INT_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip Bounded quantifiers. wenzelm@12897: wenzelm@12897: The following are not added to the default simpset because wenzelm@12897: (a) they duplicate the body and (b) there are no similar rules for @{text Int}. *} wenzelm@12897: wenzelm@12897: lemma ball_Un: "(\x \ A \ B. P x) = ((\x\A. P x) & (\x\B. P x))" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma bex_Un: "(\x \ A \ B. P x) = ((\x\A. P x) | (\x\B. P x))" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma ball_UN: "(\z \ UNION A B. P z) = (\x\A. \z \ B x. P z)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma bex_UN: "(\z \ UNION A B. P z) = (\x\A. \z\B x. P z)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip Set difference. *} wenzelm@12897: wenzelm@12897: lemma Diff_eq: "A - B = A \ (-B)" wenzelm@12897: by blast wenzelm@12897: paulson@24286: lemma Diff_eq_empty_iff [simp,noatp]: "(A - B = {}) = (A \ B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_cancel [simp]: "A - A = {}" wenzelm@12897: by blast wenzelm@12897: nipkow@14302: lemma Diff_idemp [simp]: "(A - B) - B = A - (B::'a set)" nipkow@14302: by blast nipkow@14302: wenzelm@12897: lemma Diff_triv: "A \ B = {} ==> A - B = A" wenzelm@12897: by (blast elim: equalityE) wenzelm@12897: wenzelm@12897: lemma empty_Diff [simp]: "{} - A = {}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_empty [simp]: "A - {} = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_UNIV [simp]: "A - UNIV = {}" wenzelm@12897: by blast wenzelm@12897: paulson@24286: lemma Diff_insert0 [simp,noatp]: "x \ A ==> A - insert x B = A - B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_insert: "A - insert a B = A - B - {a}" wenzelm@12897: -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"} *} wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_insert2: "A - insert a B = A - {a} - B" wenzelm@12897: -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"} *} wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma insert_Diff_if: "insert x A - B = (if x \ B then A - B else insert x (A - B))" wenzelm@12897: by auto wenzelm@12897: wenzelm@12897: lemma insert_Diff1 [simp]: "x \ B ==> insert x A - B = A - B" wenzelm@12897: by blast wenzelm@12897: nipkow@14302: lemma insert_Diff_single[simp]: "insert a (A - {a}) = insert a A" nipkow@14302: by blast nipkow@14302: wenzelm@12897: lemma insert_Diff: "a \ A ==> insert a (A - {a}) = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_insert_absorb: "x \ A ==> (insert x A) - {x} = A" wenzelm@12897: by auto wenzelm@12897: wenzelm@12897: lemma Diff_disjoint [simp]: "A \ (B - A) = {}" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_partition: "A \ B ==> A \ (B - A) = B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma double_diff: "A \ B ==> B \ C ==> B - (C - A) = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Diff_cancel [simp]: "A \ (B - A) = A \ B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Diff_cancel2 [simp]: "(B - A) \ A = B \ A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_Un: "A - (B \ C) = (A - B) \ (A - C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_Int: "A - (B \ C) = (A - B) \ (A - C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Diff: "(A \ B) - C = (A - C) \ (B - C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_Diff: "(A \ B) - C = A \ (B - C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_Int_distrib: "C \ (A - B) = (C \ A) - (C \ B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_Int_distrib2: "(A - B) \ C = (A \ C) - (B \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_Compl [simp]: "A - (- B) = A \ B" wenzelm@12897: by auto wenzelm@12897: wenzelm@12897: lemma Compl_Diff_eq [simp]: "- (A - B) = -A \ B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip Quantification over type @{typ bool}. *} wenzelm@12897: wenzelm@12897: lemma bool_induct: "P True \ P False \ P x" haftmann@21549: by (cases x) auto haftmann@21549: haftmann@21549: lemma all_bool_eq: "(\b. P b) \ P True \ P False" haftmann@21549: by (auto intro: bool_induct) haftmann@21549: haftmann@21549: lemma bool_contrapos: "P x \ \ P False \ P True" haftmann@21549: by (cases x) auto haftmann@21549: haftmann@21549: lemma ex_bool_eq: "(\b. P b) \ P True \ P False" haftmann@21549: by (auto intro: bool_contrapos) wenzelm@12897: wenzelm@12897: lemma Un_eq_UN: "A \ B = (\b. if b then A else B)" wenzelm@12897: by (auto simp add: split_if_mem2) wenzelm@12897: wenzelm@12897: lemma UN_bool_eq: "(\b::bool. A b) = (A True \ A False)" haftmann@21549: by (auto intro: bool_contrapos) wenzelm@12897: wenzelm@12897: lemma INT_bool_eq: "(\b::bool. A b) = (A True \ A False)" haftmann@21549: by (auto intro: bool_induct) wenzelm@12897: wenzelm@12897: text {* \medskip @{text Pow} *} wenzelm@12897: wenzelm@12897: lemma Pow_empty [simp]: "Pow {} = {{}}" wenzelm@12897: by (auto simp add: Pow_def) wenzelm@12897: wenzelm@12897: lemma Pow_insert: "Pow (insert a A) = Pow A \ (insert a ` Pow A)" wenzelm@12897: by (blast intro: image_eqI [where ?x = "u - {a}", standard]) wenzelm@12897: wenzelm@12897: lemma Pow_Compl: "Pow (- A) = {-B | B. A \ Pow B}" wenzelm@12897: by (blast intro: exI [where ?x = "- u", standard]) wenzelm@12897: wenzelm@12897: lemma Pow_UNIV [simp]: "Pow UNIV = UNIV" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Pow_subset: "Pow A \ Pow B \ Pow (A \ B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_Pow_subset: "(\x\A. Pow (B x)) \ Pow (\x\A. B x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma subset_Pow_Union: "A \ Pow (\A)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Union_Pow_eq [simp]: "\(Pow A) = A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Pow_Int_eq [simp]: "Pow (A \ B) = Pow A \ Pow B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: wenzelm@12897: text {* \medskip Miscellany. *} wenzelm@12897: wenzelm@12897: lemma set_eq_subset: "(A = B) = (A \ B & B \ A)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma subset_iff: "(A \ B) = (\t. t \ A --> t \ B)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma subset_iff_psubset_eq: "(A \ B) = ((A \ B) | (A = B))" berghofe@26800: by (unfold less_le) blast wenzelm@12897: paulson@18447: lemma all_not_in_conv [simp]: "(\x. x \ A) = (A = {})" wenzelm@12897: by blast wenzelm@12897: paulson@13831: lemma ex_in_conv: "(\x. x \ A) = (A \ {})" paulson@13831: by blast paulson@13831: wenzelm@12897: lemma distinct_lemma: "f x \ f y ==> x \ y" nipkow@17589: by iprover wenzelm@12897: wenzelm@12897: paulson@13860: text {* \medskip Miniscoping: pushing in quantifiers and big Unions paulson@13860: and Intersections. *} wenzelm@12897: wenzelm@12897: lemma UN_simps [simp]: wenzelm@12897: "!!a B C. (UN x:C. insert a (B x)) = (if C={} then {} else insert a (UN x:C. B x))" wenzelm@12897: "!!A B C. (UN x:C. A x Un B) = ((if C={} then {} else (UN x:C. A x) Un B))" wenzelm@12897: "!!A B C. (UN x:C. A Un B x) = ((if C={} then {} else A Un (UN x:C. B x)))" wenzelm@12897: "!!A B C. (UN x:C. A x Int B) = ((UN x:C. A x) Int B)" wenzelm@12897: "!!A B C. (UN x:C. A Int B x) = (A Int (UN x:C. B x))" wenzelm@12897: "!!A B C. (UN x:C. A x - B) = ((UN x:C. A x) - B)" wenzelm@12897: "!!A B C. (UN x:C. A - B x) = (A - (INT x:C. B x))" wenzelm@12897: "!!A B. (UN x: Union A. B x) = (UN y:A. UN x:y. B x)" wenzelm@12897: "!!A B C. (UN z: UNION A B. C z) = (UN x:A. UN z: B(x). C z)" wenzelm@12897: "!!A B f. (UN x:f`A. B x) = (UN a:A. B (f a))" wenzelm@12897: by auto wenzelm@12897: wenzelm@12897: lemma INT_simps [simp]: wenzelm@12897: "!!A B C. (INT x:C. A x Int B) = (if C={} then UNIV else (INT x:C. A x) Int B)" wenzelm@12897: "!!A B C. (INT x:C. A Int B x) = (if C={} then UNIV else A Int (INT x:C. B x))" wenzelm@12897: "!!A B C. (INT x:C. A x - B) = (if C={} then UNIV else (INT x:C. A x) - B)" wenzelm@12897: "!!A B C. (INT x:C. A - B x) = (if C={} then UNIV else A - (UN x:C. B x))" wenzelm@12897: "!!a B C. (INT x:C. insert a (B x)) = insert a (INT x:C. B x)" wenzelm@12897: "!!A B C. (INT x:C. A x Un B) = ((INT x:C. A x) Un B)" wenzelm@12897: "!!A B C. (INT x:C. A Un B x) = (A Un (INT x:C. B x))" wenzelm@12897: "!!A B. (INT x: Union A. B x) = (INT y:A. INT x:y. B x)" wenzelm@12897: "!!A B C. (INT z: UNION A B. C z) = (INT x:A. INT z: B(x). C z)" wenzelm@12897: "!!A B f. (INT x:f`A. B x) = (INT a:A. B (f a))" wenzelm@12897: by auto wenzelm@12897: paulson@24286: lemma ball_simps [simp,noatp]: wenzelm@12897: "!!A P Q. (ALL x:A. P x | Q) = ((ALL x:A. P x) | Q)" wenzelm@12897: "!!A P Q. (ALL x:A. P | Q x) = (P | (ALL x:A. Q x))" wenzelm@12897: "!!A P Q. (ALL x:A. P --> Q x) = (P --> (ALL x:A. Q x))" wenzelm@12897: "!!A P Q. (ALL x:A. P x --> Q) = ((EX x:A. P x) --> Q)" wenzelm@12897: "!!P. (ALL x:{}. P x) = True" wenzelm@12897: "!!P. (ALL x:UNIV. P x) = (ALL x. P x)" wenzelm@12897: "!!a B P. (ALL x:insert a B. P x) = (P a & (ALL x:B. P x))" wenzelm@12897: "!!A P. (ALL x:Union A. P x) = (ALL y:A. ALL x:y. P x)" wenzelm@12897: "!!A B P. (ALL x: UNION A B. P x) = (ALL a:A. ALL x: B a. P x)" wenzelm@12897: "!!P Q. (ALL x:Collect Q. P x) = (ALL x. Q x --> P x)" wenzelm@12897: "!!A P f. (ALL x:f`A. P x) = (ALL x:A. P (f x))" wenzelm@12897: "!!A P. (~(ALL x:A. P x)) = (EX x:A. ~P x)" wenzelm@12897: by auto wenzelm@12897: paulson@24286: lemma bex_simps [simp,noatp]: wenzelm@12897: "!!A P Q. (EX x:A. P x & Q) = ((EX x:A. P x) & Q)" wenzelm@12897: "!!A P Q. (EX x:A. P & Q x) = (P & (EX x:A. Q x))" wenzelm@12897: "!!P. (EX x:{}. P x) = False" wenzelm@12897: "!!P. (EX x:UNIV. P x) = (EX x. P x)" wenzelm@12897: "!!a B P. (EX x:insert a B. P x) = (P(a) | (EX x:B. P x))" wenzelm@12897: "!!A P. (EX x:Union A. P x) = (EX y:A. EX x:y. P x)" wenzelm@12897: "!!A B P. (EX x: UNION A B. P x) = (EX a:A. EX x:B a. P x)" wenzelm@12897: "!!P Q. (EX x:Collect Q. P x) = (EX x. Q x & P x)" wenzelm@12897: "!!A P f. (EX x:f`A. P x) = (EX x:A. P (f x))" wenzelm@12897: "!!A P. (~(EX x:A. P x)) = (ALL x:A. ~P x)" wenzelm@12897: by auto wenzelm@12897: wenzelm@12897: lemma ball_conj_distrib: wenzelm@12897: "(ALL x:A. P x & Q x) = ((ALL x:A. P x) & (ALL x:A. Q x))" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma bex_disj_distrib: wenzelm@12897: "(EX x:A. P x | Q x) = ((EX x:A. P x) | (EX x:A. Q x))" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: paulson@13860: text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *} paulson@13860: paulson@13860: lemma UN_extend_simps: paulson@13860: "!!a B C. insert a (UN x:C. B x) = (if C={} then {a} else (UN x:C. insert a (B x)))" paulson@13860: "!!A B C. (UN x:C. A x) Un B = (if C={} then B else (UN x:C. A x Un B))" paulson@13860: "!!A B C. A Un (UN x:C. B x) = (if C={} then A else (UN x:C. A Un B x))" paulson@13860: "!!A B C. ((UN x:C. A x) Int B) = (UN x:C. A x Int B)" paulson@13860: "!!A B C. (A Int (UN x:C. B x)) = (UN x:C. A Int B x)" paulson@13860: "!!A B C. ((UN x:C. A x) - B) = (UN x:C. A x - B)" paulson@13860: "!!A B C. (A - (INT x:C. B x)) = (UN x:C. A - B x)" paulson@13860: "!!A B. (UN y:A. UN x:y. B x) = (UN x: Union A. B x)" paulson@13860: "!!A B C. (UN x:A. UN z: B(x). C z) = (UN z: UNION A B. C z)" paulson@13860: "!!A B f. (UN a:A. B (f a)) = (UN x:f`A. B x)" paulson@13860: by auto paulson@13860: paulson@13860: lemma INT_extend_simps: paulson@13860: "!!A B C. (INT x:C. A x) Int B = (if C={} then B else (INT x:C. A x Int B))" paulson@13860: "!!A B C. A Int (INT x:C. B x) = (if C={} then A else (INT x:C. A Int B x))" paulson@13860: "!!A B C. (INT x:C. A x) - B = (if C={} then UNIV-B else (INT x:C. A x - B))" paulson@13860: "!!A B C. A - (UN x:C. B x) = (if C={} then A else (INT x:C. A - B x))" paulson@13860: "!!a B C. insert a (INT x:C. B x) = (INT x:C. insert a (B x))" paulson@13860: "!!A B C. ((INT x:C. A x) Un B) = (INT x:C. A x Un B)" paulson@13860: "!!A B C. A Un (INT x:C. B x) = (INT x:C. A Un B x)" paulson@13860: "!!A B. (INT y:A. INT x:y. B x) = (INT x: Union A. B x)" paulson@13860: "!!A B C. (INT x:A. INT z: B(x). C z) = (INT z: UNION A B. C z)" paulson@13860: "!!A B f. (INT a:A. B (f a)) = (INT x:f`A. B x)" paulson@13860: by auto paulson@13860: paulson@13860: wenzelm@12897: subsubsection {* Monotonicity of various operations *} wenzelm@12897: wenzelm@12897: lemma image_mono: "A \ B ==> f`A \ f`B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Pow_mono: "A \ B ==> Pow A \ Pow B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Union_mono: "A \ B ==> \A \ \B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Inter_anti_mono: "B \ A ==> \A \ \B" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_mono: wenzelm@12897: "A \ B ==> (!!x. x \ A ==> f x \ g x) ==> wenzelm@12897: (\x\A. f x) \ (\x\B. g x)" wenzelm@12897: by (blast dest: subsetD) wenzelm@12897: wenzelm@12897: lemma INT_anti_mono: wenzelm@12897: "B \ A ==> (!!x. x \ A ==> f x \ g x) ==> wenzelm@12897: (\x\A. f x) \ (\x\A. g x)" wenzelm@12897: -- {* The last inclusion is POSITIVE! *} wenzelm@12897: by (blast dest: subsetD) wenzelm@12897: wenzelm@12897: lemma insert_mono: "C \ D ==> insert a C \ insert a D" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_mono: "A \ C ==> B \ D ==> A \ B \ C \ D" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_mono: "A \ C ==> B \ D ==> A \ B \ C \ D" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Diff_mono: "A \ C ==> D \ B ==> A - B \ C - D" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Compl_anti_mono: "A \ B ==> -B \ -A" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: text {* \medskip Monotonicity of implications. *} wenzelm@12897: wenzelm@12897: lemma in_mono: "A \ B ==> x \ A --> x \ B" wenzelm@12897: apply (rule impI) paulson@14208: apply (erule subsetD, assumption) wenzelm@12897: done wenzelm@12897: wenzelm@12897: lemma conj_mono: "P1 --> Q1 ==> P2 --> Q2 ==> (P1 & P2) --> (Q1 & Q2)" nipkow@17589: by iprover wenzelm@12897: wenzelm@12897: lemma disj_mono: "P1 --> Q1 ==> P2 --> Q2 ==> (P1 | P2) --> (Q1 | Q2)" nipkow@17589: by iprover wenzelm@12897: wenzelm@12897: lemma imp_mono: "Q1 --> P1 ==> P2 --> Q2 ==> (P1 --> P2) --> (Q1 --> Q2)" nipkow@17589: by iprover wenzelm@12897: wenzelm@12897: lemma imp_refl: "P --> P" .. wenzelm@12897: wenzelm@12897: lemma ex_mono: "(!!x. P x --> Q x) ==> (EX x. P x) --> (EX x. Q x)" nipkow@17589: by iprover wenzelm@12897: wenzelm@12897: lemma all_mono: "(!!x. P x --> Q x) ==> (ALL x. P x) --> (ALL x. Q x)" nipkow@17589: by iprover wenzelm@12897: wenzelm@12897: lemma Collect_mono: "(!!x. P x --> Q x) ==> Collect P \ Collect Q" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_Collect_mono: wenzelm@12897: "A \ B ==> (!!x. x \ A ==> P x --> Q x) ==> A \ Collect P \ B \ Collect Q" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemmas basic_monos = wenzelm@12897: subset_refl imp_refl disj_mono conj_mono wenzelm@12897: ex_mono Collect_mono in_mono wenzelm@12897: wenzelm@12897: lemma eq_to_mono: "a = b ==> c = d ==> b --> d ==> a --> c" nipkow@17589: by iprover wenzelm@12897: wenzelm@12897: lemma eq_to_mono2: "a = b ==> c = d ==> ~ b --> ~ d ==> ~ a --> ~ c" nipkow@17589: by iprover wenzelm@11979: wenzelm@12020: haftmann@30531: subsection {* Inverse image of a function *} wenzelm@12257: wenzelm@12257: constdefs wenzelm@12257: vimage :: "('a => 'b) => 'b set => 'a set" (infixr "-`" 90) haftmann@28562: [code del]: "f -` B == {x. f x : B}" wenzelm@12257: haftmann@30531: haftmann@30531: subsubsection {* Basic rules *} haftmann@30531: wenzelm@12257: lemma vimage_eq [simp]: "(a : f -` B) = (f a : B)" wenzelm@12257: by (unfold vimage_def) blast wenzelm@12257: wenzelm@12257: lemma vimage_singleton_eq: "(a : f -` {b}) = (f a = b)" wenzelm@12257: by simp wenzelm@12257: wenzelm@12257: lemma vimageI [intro]: "f a = b ==> b:B ==> a : f -` B" wenzelm@12257: by (unfold vimage_def) blast wenzelm@12257: wenzelm@12257: lemma vimageI2: "f a : A ==> a : f -` A" wenzelm@12257: by (unfold vimage_def) fast wenzelm@12257: wenzelm@12257: lemma vimageE [elim!]: "a: f -` B ==> (!!x. f a = x ==> x:B ==> P) ==> P" wenzelm@12257: by (unfold vimage_def) blast wenzelm@12257: wenzelm@12257: lemma vimageD: "a : f -` A ==> f a : A" wenzelm@12257: by (unfold vimage_def) fast wenzelm@12257: haftmann@30531: haftmann@30531: subsubsection {* Equations *} haftmann@30531: wenzelm@12257: lemma vimage_empty [simp]: "f -` {} = {}" wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_Compl: "f -` (-A) = -(f -` A)" wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_Un [simp]: "f -` (A Un B) = (f -` A) Un (f -` B)" wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_Int [simp]: "f -` (A Int B) = (f -` A) Int (f -` B)" wenzelm@12257: by fast wenzelm@12257: wenzelm@12257: lemma vimage_Union: "f -` (Union A) = (UN X:A. f -` X)" wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_UN: "f-`(UN x:A. B x) = (UN x:A. f -` B x)" wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_INT: "f-`(INT x:A. B x) = (INT x:A. f -` B x)" wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_Collect_eq [simp]: "f -` Collect P = {y. P (f y)}" wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_Collect: "(!!x. P (f x) = Q x) ==> f -` (Collect P) = Collect Q" wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_insert: "f-`(insert a B) = (f-`{a}) Un (f-`B)" wenzelm@12257: -- {* NOT suitable for rewriting because of the recurrence of @{term "{a}"}. *} wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_Diff: "f -` (A - B) = (f -` A) - (f -` B)" wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_UNIV [simp]: "f -` UNIV = UNIV" wenzelm@12257: by blast wenzelm@12257: wenzelm@12257: lemma vimage_eq_UN: "f-`B = (UN y: B. f-`{y})" wenzelm@12257: -- {* NOT suitable for rewriting *} wenzelm@12257: by blast wenzelm@12257: wenzelm@12897: lemma vimage_mono: "A \ B ==> f -` A \ f -` B" wenzelm@12257: -- {* monotonicity *} wenzelm@12257: by blast wenzelm@12257: haftmann@26150: lemma vimage_image_eq [noatp]: "f -` (f ` A) = {y. EX x:A. f x = f y}" haftmann@26150: by (blast intro: sym) haftmann@26150: haftmann@26150: lemma image_vimage_subset: "f ` (f -` A) <= A" haftmann@26150: by blast haftmann@26150: haftmann@26150: lemma image_vimage_eq [simp]: "f ` (f -` A) = A Int range f" haftmann@26150: by blast haftmann@26150: haftmann@26150: lemma image_Int_subset: "f`(A Int B) <= f`A Int f`B" haftmann@26150: by blast haftmann@26150: haftmann@26150: lemma image_diff_subset: "f`A - f`B <= f`(A - B)" haftmann@26150: by blast haftmann@26150: haftmann@26150: lemma image_UN: "(f ` (UNION A B)) = (UN x:A.(f ` (B x)))" haftmann@26150: by blast haftmann@26150: wenzelm@12257: haftmann@30531: subsection {* Getting the Contents of a Singleton Set *} haftmann@30531: haftmann@30531: definition contents :: "'a set \ 'a" where haftmann@30531: [code del]: "contents X = (THE x. X = {x})" haftmann@30531: haftmann@30531: lemma contents_eq [simp]: "contents {x} = x" haftmann@30531: by (simp add: contents_def) haftmann@30531: haftmann@30531: haftmann@30531: subsection {* Transitivity rules for calculational reasoning *} haftmann@30531: haftmann@30531: lemma set_rev_mp: "x:A ==> A \ B ==> x:B" haftmann@30531: by (rule subsetD) haftmann@30531: haftmann@30531: lemma set_mp: "A \ B ==> x:A ==> x:B" haftmann@30531: by (rule subsetD) haftmann@30531: haftmann@30531: lemmas basic_trans_rules [trans] = haftmann@30531: order_trans_rules set_rev_mp set_mp haftmann@30531: haftmann@30531: haftmann@30531: subsection {* Least value operator *} berghofe@26800: berghofe@26800: lemma Least_mono: berghofe@26800: "mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y berghofe@26800: ==> (LEAST y. y : f ` S) = f (LEAST x. x : S)" berghofe@26800: -- {* Courtesy of Stephan Merz *} berghofe@26800: apply clarify berghofe@26800: apply (erule_tac P = "%x. x : S" in LeastI2_order, fast) berghofe@26800: apply (rule LeastI2_order) berghofe@26800: apply (auto elim: monoD intro!: order_antisym) berghofe@26800: done berghofe@26800: haftmann@24420: haftmann@30531: subsection {* Rudimentary code generation *} haftmann@27824: haftmann@28562: lemma empty_code [code]: "{} x \ False" haftmann@27824: unfolding empty_def Collect_def .. haftmann@27824: haftmann@28562: lemma UNIV_code [code]: "UNIV x \ True" haftmann@27824: unfolding UNIV_def Collect_def .. haftmann@27824: haftmann@28562: lemma insert_code [code]: "insert y A x \ y = x \ A x" haftmann@27824: unfolding insert_def Collect_def mem_def Un_def by auto haftmann@27824: haftmann@28562: lemma inter_code [code]: "(A \ B) x \ A x \ B x" haftmann@27824: unfolding Int_def Collect_def mem_def .. haftmann@27824: haftmann@28562: lemma union_code [code]: "(A \ B) x \ A x \ B x" haftmann@27824: unfolding Un_def Collect_def mem_def .. haftmann@27824: haftmann@28562: lemma vimage_code [code]: "(f -` A) x = A (f x)" haftmann@27824: unfolding vimage_def Collect_def mem_def .. haftmann@27824: haftmann@27824: haftmann@30531: subsection {* Complete lattices *} haftmann@30531: haftmann@30531: notation haftmann@30531: less_eq (infix "\" 50) and haftmann@30531: less (infix "\" 50) and haftmann@30531: inf (infixl "\" 70) and haftmann@30531: sup (infixl "\" 65) haftmann@30531: haftmann@30531: class complete_lattice = lattice + bot + top + haftmann@30531: fixes Inf :: "'a set \ 'a" ("\_" [900] 900) haftmann@30531: and Sup :: "'a set \ 'a" ("\_" [900] 900) haftmann@30531: assumes Inf_lower: "x \ A \ \A \ x" haftmann@30531: and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" haftmann@30531: assumes Sup_upper: "x \ A \ x \ \A" haftmann@30531: and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" haftmann@30531: begin haftmann@30531: haftmann@30531: lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" haftmann@30531: by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) haftmann@30531: haftmann@30531: lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" haftmann@30531: by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) haftmann@30531: haftmann@30531: lemma Inf_Univ: "\UNIV = \{}" haftmann@30531: unfolding Sup_Inf by auto haftmann@30531: haftmann@30531: lemma Sup_Univ: "\UNIV = \{}" haftmann@30531: unfolding Inf_Sup by auto haftmann@30531: haftmann@30531: lemma Inf_insert: "\insert a A = a \ \A" haftmann@30531: by (auto intro: antisym Inf_greatest Inf_lower) haftmann@30531: haftmann@30531: lemma Sup_insert: "\insert a A = a \ \A" haftmann@30531: by (auto intro: antisym Sup_least Sup_upper) haftmann@30531: haftmann@30531: lemma Inf_singleton [simp]: haftmann@30531: "\{a} = a" haftmann@30531: by (auto intro: antisym Inf_lower Inf_greatest) haftmann@30531: haftmann@30531: lemma Sup_singleton [simp]: haftmann@30531: "\{a} = a" haftmann@30531: by (auto intro: antisym Sup_upper Sup_least) haftmann@30531: haftmann@30531: lemma Inf_insert_simp: haftmann@30531: "\insert a A = (if A = {} then a else a \ \A)" haftmann@30531: by (cases "A = {}") (simp_all, simp add: Inf_insert) haftmann@30531: haftmann@30531: lemma Sup_insert_simp: haftmann@30531: "\insert a A = (if A = {} then a else a \ \A)" haftmann@30531: by (cases "A = {}") (simp_all, simp add: Sup_insert) haftmann@30531: haftmann@30531: lemma Inf_binary: haftmann@30531: "\{a, b} = a \ b" haftmann@30531: by (simp add: Inf_insert_simp) haftmann@30531: haftmann@30531: lemma Sup_binary: haftmann@30531: "\{a, b} = a \ b" haftmann@30531: by (simp add: Sup_insert_simp) haftmann@30531: haftmann@30531: lemma bot_def: haftmann@30531: "bot = \{}" haftmann@30531: by (auto intro: antisym Sup_least) haftmann@30531: haftmann@30531: lemma top_def: haftmann@30531: "top = \{}" haftmann@30531: by (auto intro: antisym Inf_greatest) haftmann@30531: haftmann@30531: lemma sup_bot [simp]: haftmann@30531: "x \ bot = x" haftmann@30531: using bot_least [of x] by (simp add: le_iff_sup sup_commute) haftmann@30531: haftmann@30531: lemma inf_top [simp]: haftmann@30531: "x \ top = x" haftmann@30531: using top_greatest [of x] by (simp add: le_iff_inf inf_commute) haftmann@30531: haftmann@30531: definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where haftmann@30531: "SUPR A f == \ (f ` A)" haftmann@30531: haftmann@30531: definition INFI :: "'b set \ ('b \ 'a) \ 'a" where haftmann@30531: "INFI A f == \ (f ` A)" haftmann@30531: haftmann@30531: end haftmann@30531: haftmann@30531: syntax haftmann@30531: "_SUP1" :: "pttrns => 'b => 'b" ("(3SUP _./ _)" [0, 10] 10) haftmann@30531: "_SUP" :: "pttrn => 'a set => 'b => 'b" ("(3SUP _:_./ _)" [0, 10] 10) haftmann@30531: "_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10) haftmann@30531: "_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10) haftmann@30531: haftmann@30531: translations haftmann@30531: "SUP x y. B" == "SUP x. SUP y. B" haftmann@30531: "SUP x. B" == "CONST SUPR CONST UNIV (%x. B)" haftmann@30531: "SUP x. B" == "SUP x:CONST UNIV. B" haftmann@30531: "SUP x:A. B" == "CONST SUPR A (%x. B)" haftmann@30531: "INF x y. B" == "INF x. INF y. B" haftmann@30531: "INF x. B" == "CONST INFI CONST UNIV (%x. B)" haftmann@30531: "INF x. B" == "INF x:CONST UNIV. B" haftmann@30531: "INF x:A. B" == "CONST INFI A (%x. B)" haftmann@30531: haftmann@30531: (* To avoid eta-contraction of body: *) haftmann@30531: print_translation {* haftmann@30531: let haftmann@30531: fun btr' syn (A :: Abs abs :: ts) = haftmann@30531: let val (x,t) = atomic_abs_tr' abs haftmann@30531: in list_comb (Syntax.const syn $ x $ A $ t, ts) end haftmann@30531: val const_syntax_name = Sign.const_syntax_name @{theory} o fst o dest_Const haftmann@30531: in haftmann@30531: [(const_syntax_name @{term SUPR}, btr' "_SUP"),(const_syntax_name @{term "INFI"}, btr' "_INF")] haftmann@30531: end haftmann@30531: *} haftmann@30531: haftmann@30531: context complete_lattice haftmann@30531: begin haftmann@30531: haftmann@30531: lemma le_SUPI: "i : A \ M i \ (SUP i:A. M i)" haftmann@30531: by (auto simp add: SUPR_def intro: Sup_upper) haftmann@30531: haftmann@30531: lemma SUP_leI: "(\i. i : A \ M i \ u) \ (SUP i:A. M i) \ u" haftmann@30531: by (auto simp add: SUPR_def intro: Sup_least) haftmann@30531: haftmann@30531: lemma INF_leI: "i : A \ (INF i:A. M i) \ M i" haftmann@30531: by (auto simp add: INFI_def intro: Inf_lower) haftmann@30531: haftmann@30531: lemma le_INFI: "(\i. i : A \ u \ M i) \ u \ (INF i:A. M i)" haftmann@30531: by (auto simp add: INFI_def intro: Inf_greatest) haftmann@30531: haftmann@30531: lemma SUP_const[simp]: "A \ {} \ (SUP i:A. M) = M" haftmann@30531: by (auto intro: antisym SUP_leI le_SUPI) haftmann@30531: haftmann@30531: lemma INF_const[simp]: "A \ {} \ (INF i:A. M) = M" haftmann@30531: by (auto intro: antisym INF_leI le_INFI) haftmann@30531: haftmann@30531: end haftmann@30531: haftmann@30531: haftmann@30531: subsection {* Bool as complete lattice *} haftmann@30531: haftmann@30531: instantiation bool :: complete_lattice haftmann@30531: begin haftmann@30531: haftmann@30531: definition haftmann@30531: Inf_bool_def: "\A \ (\x\A. x)" haftmann@30531: haftmann@30531: definition haftmann@30531: Sup_bool_def: "\A \ (\x\A. x)" haftmann@30531: haftmann@30531: instance haftmann@30531: by intro_classes (auto simp add: Inf_bool_def Sup_bool_def le_bool_def) haftmann@30531: haftmann@30531: end haftmann@30531: haftmann@30531: lemma Inf_empty_bool [simp]: haftmann@30531: "\{}" haftmann@30531: unfolding Inf_bool_def by auto haftmann@30531: haftmann@30531: lemma not_Sup_empty_bool [simp]: wenzelm@30814: "\ \{}" haftmann@30531: unfolding Sup_bool_def by auto haftmann@30531: haftmann@30531: haftmann@30531: subsection {* Fun as complete lattice *} haftmann@30531: haftmann@30531: instantiation "fun" :: (type, complete_lattice) complete_lattice haftmann@30531: begin haftmann@30531: haftmann@30531: definition haftmann@30531: Inf_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" haftmann@30531: haftmann@30531: definition haftmann@30531: Sup_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" haftmann@30531: haftmann@30531: instance haftmann@30531: by intro_classes haftmann@30531: (auto simp add: Inf_fun_def Sup_fun_def le_fun_def haftmann@30531: intro: Inf_lower Sup_upper Inf_greatest Sup_least) haftmann@30531: haftmann@30531: end haftmann@30531: haftmann@30531: lemma Inf_empty_fun: haftmann@30531: "\{} = (\_. \{})" haftmann@30531: by rule (auto simp add: Inf_fun_def) haftmann@30531: haftmann@30531: lemma Sup_empty_fun: haftmann@30531: "\{} = (\_. \{})" haftmann@30531: by rule (auto simp add: Sup_fun_def) haftmann@30531: haftmann@30531: haftmann@30531: subsection {* Set as lattice *} haftmann@30531: haftmann@30531: lemma inf_set_eq: "A \ B = A \ B" haftmann@30531: apply (rule subset_antisym) haftmann@30531: apply (rule Int_greatest) haftmann@30531: apply (rule inf_le1) haftmann@30531: apply (rule inf_le2) haftmann@30531: apply (rule inf_greatest) haftmann@30531: apply (rule Int_lower1) haftmann@30531: apply (rule Int_lower2) haftmann@30531: done haftmann@30531: haftmann@30531: lemma sup_set_eq: "A \ B = A \ B" haftmann@30531: apply (rule subset_antisym) haftmann@30531: apply (rule sup_least) haftmann@30531: apply (rule Un_upper1) haftmann@30531: apply (rule Un_upper2) haftmann@30531: apply (rule Un_least) haftmann@30531: apply (rule sup_ge1) haftmann@30531: apply (rule sup_ge2) haftmann@30531: done haftmann@30531: haftmann@30531: lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" haftmann@30531: apply (fold inf_set_eq sup_set_eq) haftmann@30531: apply (erule mono_inf) haftmann@30531: done haftmann@30531: haftmann@30531: lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" haftmann@30531: apply (fold inf_set_eq sup_set_eq) haftmann@30531: apply (erule mono_sup) haftmann@30531: done haftmann@30531: haftmann@30531: lemma Inf_set_eq: "\S = \S" haftmann@30531: apply (rule subset_antisym) haftmann@30531: apply (rule Inter_greatest) haftmann@30531: apply (erule Inf_lower) haftmann@30531: apply (rule Inf_greatest) haftmann@30531: apply (erule Inter_lower) haftmann@30531: done haftmann@30531: haftmann@30531: lemma Sup_set_eq: "\S = \S" haftmann@30531: apply (rule subset_antisym) haftmann@30531: apply (rule Sup_least) haftmann@30531: apply (erule Union_upper) haftmann@30531: apply (rule Union_least) haftmann@30531: apply (erule Sup_upper) haftmann@30531: done haftmann@30531: haftmann@30531: lemma top_set_eq: "top = UNIV" haftmann@30531: by (iprover intro!: subset_antisym subset_UNIV top_greatest) haftmann@30531: haftmann@30531: lemma bot_set_eq: "bot = {}" haftmann@30531: by (iprover intro!: subset_antisym empty_subsetI bot_least) haftmann@30531: haftmann@30531: no_notation haftmann@30531: less_eq (infix "\" 50) and haftmann@30531: less (infix "\" 50) and haftmann@30531: inf (infixl "\" 70) and haftmann@30531: sup (infixl "\" 65) and haftmann@30531: Inf ("\_" [900] 900) and haftmann@30531: Sup ("\_" [900] 900) haftmann@30531: haftmann@30531: haftmann@30596: subsection {* Misc theorem and ML bindings *} haftmann@30596: haftmann@30596: lemmas equalityI = subset_antisym haftmann@30596: lemmas mem_simps = haftmann@30596: insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff haftmann@30596: mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff haftmann@30596: -- {* Each of these has ALREADY been added @{text "[simp]"} above. *} wenzelm@21669: wenzelm@21669: ML {* wenzelm@22139: val Ball_def = @{thm Ball_def} wenzelm@22139: val Bex_def = @{thm Bex_def} wenzelm@22139: val CollectD = @{thm CollectD} wenzelm@22139: val CollectE = @{thm CollectE} wenzelm@22139: val CollectI = @{thm CollectI} wenzelm@22139: val Collect_conj_eq = @{thm Collect_conj_eq} wenzelm@22139: val Collect_mem_eq = @{thm Collect_mem_eq} wenzelm@22139: val IntD1 = @{thm IntD1} wenzelm@22139: val IntD2 = @{thm IntD2} wenzelm@22139: val IntE = @{thm IntE} wenzelm@22139: val IntI = @{thm IntI} wenzelm@22139: val Int_Collect = @{thm Int_Collect} wenzelm@22139: val UNIV_I = @{thm UNIV_I} wenzelm@22139: val UNIV_witness = @{thm UNIV_witness} wenzelm@22139: val UnE = @{thm UnE} wenzelm@22139: val UnI1 = @{thm UnI1} wenzelm@22139: val UnI2 = @{thm UnI2} wenzelm@22139: val ballE = @{thm ballE} wenzelm@22139: val ballI = @{thm ballI} wenzelm@22139: val bexCI = @{thm bexCI} wenzelm@22139: val bexE = @{thm bexE} wenzelm@22139: val bexI = @{thm bexI} wenzelm@22139: val bex_triv = @{thm bex_triv} wenzelm@22139: val bspec = @{thm bspec} wenzelm@22139: val contra_subsetD = @{thm contra_subsetD} wenzelm@22139: val distinct_lemma = @{thm distinct_lemma} wenzelm@22139: val eq_to_mono = @{thm eq_to_mono} wenzelm@22139: val eq_to_mono2 = @{thm eq_to_mono2} wenzelm@22139: val equalityCE = @{thm equalityCE} wenzelm@22139: val equalityD1 = @{thm equalityD1} wenzelm@22139: val equalityD2 = @{thm equalityD2} wenzelm@22139: val equalityE = @{thm equalityE} wenzelm@22139: val equalityI = @{thm equalityI} wenzelm@22139: val imageE = @{thm imageE} wenzelm@22139: val imageI = @{thm imageI} wenzelm@22139: val image_Un = @{thm image_Un} wenzelm@22139: val image_insert = @{thm image_insert} wenzelm@22139: val insert_commute = @{thm insert_commute} wenzelm@22139: val insert_iff = @{thm insert_iff} wenzelm@22139: val mem_Collect_eq = @{thm mem_Collect_eq} wenzelm@22139: val rangeE = @{thm rangeE} wenzelm@22139: val rangeI = @{thm rangeI} wenzelm@22139: val range_eqI = @{thm range_eqI} wenzelm@22139: val subsetCE = @{thm subsetCE} wenzelm@22139: val subsetD = @{thm subsetD} wenzelm@22139: val subsetI = @{thm subsetI} wenzelm@22139: val subset_refl = @{thm subset_refl} wenzelm@22139: val subset_trans = @{thm subset_trans} wenzelm@22139: val vimageD = @{thm vimageD} wenzelm@22139: val vimageE = @{thm vimageE} wenzelm@22139: val vimageI = @{thm vimageI} wenzelm@22139: val vimageI2 = @{thm vimageI2} wenzelm@22139: val vimage_Collect = @{thm vimage_Collect} wenzelm@22139: val vimage_Int = @{thm vimage_Int} wenzelm@22139: val vimage_Un = @{thm vimage_Un} wenzelm@21669: *} wenzelm@21669: wenzelm@11979: end