haftmann@32139: (* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel *) 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: haftmann@32081: subsection {* Sets as predicates *} haftmann@30531: haftmann@37677: types 'a set = "'a \ bool" wenzelm@3820: haftmann@37677: definition Collect :: "('a \ bool) \ 'a set" where -- "comprehension" haftmann@37677: "Collect P = P" haftmann@30304: haftmann@37677: definition member :: "'a \ 'a set \ bool" where -- "membership" haftmann@37677: mem_def: "member x A = A x" wenzelm@19656: wenzelm@21210: notation haftmann@37677: member ("op :") and haftmann@37677: member ("(_/ : _)" [50, 51] 50) wenzelm@11979: haftmann@37677: abbreviation not_member where haftmann@37677: "not_member x A \ ~ (x : A)" -- "non-membership" wenzelm@19656: wenzelm@21210: notation haftmann@37677: not_member ("op ~:") and haftmann@37677: not_member ("(_/ ~: _)" [50, 51] 50) wenzelm@19656: wenzelm@21210: notation (xsymbols) haftmann@37677: member ("op \") and haftmann@37677: member ("(_/ \ _)" [50, 51] 50) and haftmann@37677: not_member ("op \") and haftmann@37677: not_member ("(_/ \ _)" [50, 51] 50) wenzelm@19656: wenzelm@21210: notation (HTML output) haftmann@37677: member ("op \") and haftmann@37677: member ("(_/ \ _)" [50, 51] 50) and haftmann@37677: not_member ("op \") and haftmann@37677: not_member ("(_/ \ _)" [50, 51] 50) wenzelm@19656: haftmann@32081: text {* Set comprehensions *} haftmann@32081: haftmann@30531: syntax wenzelm@35115: "_Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})") haftmann@30531: translations wenzelm@35115: "{x. P}" == "CONST Collect (%x. P)" haftmann@30531: haftmann@32081: syntax wenzelm@35115: "_Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ :/ _./ _})") haftmann@32081: syntax (xsymbols) wenzelm@35115: "_Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ \/ _./ _})") haftmann@32081: translations wenzelm@35115: "{x:A. P}" => "{x. x:A & P}" haftmann@32081: haftmann@32081: lemma mem_Collect_eq [iff]: "(a : {x. P(x)}) = P(a)" haftmann@32081: by (simp add: Collect_def mem_def) haftmann@32081: haftmann@32081: lemma Collect_mem_eq [simp]: "{x. x:A} = A" haftmann@32081: by (simp add: Collect_def mem_def) haftmann@32081: haftmann@32081: lemma CollectI: "P(a) ==> a : {x. P(x)}" haftmann@32081: by simp haftmann@32081: haftmann@32081: lemma CollectD: "a : {x. P(x)} ==> P(a)" haftmann@32081: by simp haftmann@32081: haftmann@32081: lemma Collect_cong: "(!!x. P x = Q x) ==> {x. P(x)} = {x. Q(x)}" haftmann@32081: by simp haftmann@32081: haftmann@32117: text {* haftmann@32117: Simproc for pulling @{text "x=t"} in @{text "{x. \ & x=t & \}"} haftmann@32117: to the front (and similarly for @{text "t=x"}): haftmann@32117: *} haftmann@32117: haftmann@32117: setup {* haftmann@32117: let haftmann@32117: val Coll_perm_tac = rtac @{thm Collect_cong} 1 THEN rtac @{thm iffI} 1 THEN haftmann@32117: ALLGOALS(EVERY'[REPEAT_DETERM o (etac @{thm conjE}), haftmann@32117: DEPTH_SOLVE_1 o (ares_tac [@{thm conjI}])]) haftmann@32117: val defColl_regroup = Simplifier.simproc @{theory} haftmann@32117: "defined Collect" ["{x. P x & Q x}"] haftmann@32117: (Quantifier1.rearrange_Coll Coll_perm_tac) haftmann@32117: in haftmann@32117: Simplifier.map_simpset (fn ss => ss addsimprocs [defColl_regroup]) haftmann@32117: end haftmann@32117: *} haftmann@32117: haftmann@32081: lemmas CollectE = CollectD [elim_format] haftmann@32081: haftmann@32081: text {* Set enumerations *} haftmann@30531: haftmann@32264: abbreviation empty :: "'a set" ("{}") where haftmann@32264: "{} \ bot" haftmann@31456: haftmann@31456: definition insert :: "'a \ 'a set \ 'a set" where haftmann@32081: insert_compr: "insert a B = {x. x = a \ x \ B}" haftmann@31456: haftmann@31456: syntax wenzelm@35115: "_Finset" :: "args => 'a set" ("{(_)}") haftmann@31456: translations wenzelm@35115: "{x, xs}" == "CONST insert x {xs}" wenzelm@35115: "{x}" == "CONST insert x {}" haftmann@31456: haftmann@32081: haftmann@32081: subsection {* Subsets and bounded quantifiers *} haftmann@32081: haftmann@32081: abbreviation haftmann@32081: subset :: "'a set \ 'a set \ bool" where haftmann@32081: "subset \ less" haftmann@32081: haftmann@32081: abbreviation haftmann@32081: subset_eq :: "'a set \ 'a set \ bool" where haftmann@32081: "subset_eq \ less_eq" haftmann@32081: haftmann@32081: notation (output) haftmann@32081: subset ("op <") and haftmann@32081: subset ("(_/ < _)" [50, 51] 50) and haftmann@32081: subset_eq ("op <=") and haftmann@32081: subset_eq ("(_/ <= _)" [50, 51] 50) haftmann@32081: haftmann@32081: notation (xsymbols) haftmann@32081: subset ("op \") and haftmann@32081: subset ("(_/ \ _)" [50, 51] 50) and haftmann@32081: subset_eq ("op \") and haftmann@32081: subset_eq ("(_/ \ _)" [50, 51] 50) haftmann@32081: haftmann@32081: notation (HTML output) haftmann@32081: subset ("op \") and haftmann@32081: subset ("(_/ \ _)" [50, 51] 50) and haftmann@32081: subset_eq ("op \") and haftmann@32081: subset_eq ("(_/ \ _)" [50, 51] 50) haftmann@32081: haftmann@32081: abbreviation (input) haftmann@32081: supset :: "'a set \ 'a set \ bool" where haftmann@32081: "supset \ greater" haftmann@32081: haftmann@32081: abbreviation (input) haftmann@32081: supset_eq :: "'a set \ 'a set \ bool" where haftmann@32081: "supset_eq \ greater_eq" haftmann@32081: haftmann@32081: notation (xsymbols) haftmann@32081: supset ("op \") and haftmann@32081: supset ("(_/ \ _)" [50, 51] 50) and haftmann@32081: supset_eq ("op \") and haftmann@32081: supset_eq ("(_/ \ _)" [50, 51] 50) haftmann@32081: haftmann@37387: definition Ball :: "'a set \ ('a \ bool) \ bool" where haftmann@37387: "Ball A P \ (\x. x \ A \ P x)" -- "bounded universal quantifiers" haftmann@32077: haftmann@37387: definition Bex :: "'a set \ ('a \ bool) \ bool" where haftmann@37387: "Bex A P \ (\x. x \ A \ P x)" -- "bounded existential quantifiers" haftmann@32077: 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 wenzelm@35115: "ALL x:A. P" == "CONST Ball A (%x. P)" wenzelm@35115: "EX x:A. P" == "CONST Bex A (%x. P)" wenzelm@35115: "EX! x:A. P" => "EX! x. x:A & P" haftmann@30531: "LEAST x:A. P" => "LEAST x. x:A & P" haftmann@30531: 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@35115: val Type (set_type, _) = @{typ "'a set"}; (* FIXME 'a => bool (!?!) *) wenzelm@35115: val All_binder = Syntax.binder_name @{const_syntax All}; wenzelm@35115: val Ex_binder = Syntax.binder_name @{const_syntax Ex}; wenzelm@22377: val impl = @{const_syntax "op -->"}; wenzelm@22377: val conj = @{const_syntax "op &"}; wenzelm@35115: val sbset = @{const_syntax subset}; wenzelm@35115: val sbset_eq = @{const_syntax subset_eq}; haftmann@21819: haftmann@21819: val trans = wenzelm@35115: [((All_binder, impl, sbset), @{syntax_const "_setlessAll"}), wenzelm@35115: ((All_binder, impl, sbset_eq), @{syntax_const "_setleAll"}), wenzelm@35115: ((Ex_binder, conj, sbset), @{syntax_const "_setlessEx"}), wenzelm@35115: ((Ex_binder, conj, sbset_eq), @{syntax_const "_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, wenzelm@35115: fn [Const (@{syntax_const "_bound"}, _) $ Free (v, Type (T, _)), wenzelm@35115: Const (c, _) $ wenzelm@35115: (Const (d, _) $ (Const (@{syntax_const "_bound"}, _) $ Free (v', _)) $ n) $ P] => wenzelm@35115: if T = set_type then wenzelm@35115: (case AList.lookup (op =) trans (q, c, d) of wenzelm@35115: NONE => raise Match wenzelm@35115: | SOME l => mk v v' l n P) wenzelm@35115: else raise Match wenzelm@35115: | _ => 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@35115: syntax wenzelm@35115: "_Setcompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})") wenzelm@35115: wenzelm@11979: parse_translation {* wenzelm@11979: let wenzelm@35115: val ex_tr = snd (mk_binder_tr ("EX ", @{const_syntax Ex})); wenzelm@3947: wenzelm@35115: fun nvars (Const (@{syntax_const "_idts"}, _) $ _ $ idts) = nvars idts + 1 wenzelm@11979: | nvars _ = 1; wenzelm@11979: wenzelm@11979: fun setcompr_tr [e, idts, b] = wenzelm@11979: let wenzelm@35115: val eq = Syntax.const @{const_syntax "op ="} $ Bound (nvars idts) $ e; wenzelm@35115: val P = Syntax.const @{const_syntax "op &"} $ eq $ b; wenzelm@11979: val exP = ex_tr [idts, P]; wenzelm@35115: in Syntax.const @{const_syntax Collect} $ Term.absdummy (dummyT, exP) end; wenzelm@11979: wenzelm@35115: in [(@{syntax_const "_Setcompr"}, setcompr_tr)] end; wenzelm@11979: *} clasohm@923: wenzelm@35115: print_translation {* wenzelm@35115: [Syntax.preserve_binder_abs2_tr' @{const_syntax Ball} @{syntax_const "_Ball"}, wenzelm@35115: Syntax.preserve_binder_abs2_tr' @{const_syntax Bex} @{syntax_const "_Bex"}] wenzelm@35115: *} -- {* to avoid eta-contraction of body *} haftmann@30531: nipkow@13763: print_translation {* nipkow@13763: let wenzelm@35115: val ex_tr' = snd (mk_binder_tr' (@{const_syntax Ex}, "DUMMY")); nipkow@13763: nipkow@13763: fun setcompr_tr' [Abs (abs as (_, _, P))] = nipkow@13763: let wenzelm@35115: fun check (Const (@{const_syntax Ex}, _) $ Abs (_, _, P), n) = check (P, n + 1) wenzelm@35115: | check (Const (@{const_syntax "op &"}, _) $ wenzelm@35115: (Const (@{const_syntax "op ="}, _) $ Bound m $ e) $ P, n) = nipkow@13763: n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso haftmann@33038: subset (op =) (0 upto (n - 1), add_loose_bnos (e, 0, [])) wenzelm@35115: | check _ = false; clasohm@923: wenzelm@11979: fun tr' (_ $ abs) = wenzelm@11979: let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr' [abs] wenzelm@35115: in Syntax.const @{syntax_const "_Setcompr"} $ e $ idts $ Q end; wenzelm@35115: in wenzelm@35115: if check (P, 0) then tr' P wenzelm@35115: else wenzelm@35115: let wenzelm@35115: val (x as _ $ Free(xN, _), t) = atomic_abs_tr' abs; wenzelm@35115: val M = Syntax.const @{syntax_const "_Coll"} $ x $ t; wenzelm@35115: in wenzelm@35115: case t of wenzelm@35115: Const (@{const_syntax "op &"}, _) $ haftmann@37677: (Const (@{const_syntax Set.member}, _) $ wenzelm@35115: (Const (@{syntax_const "_bound"}, _) $ Free (yN, _)) $ A) $ P => wenzelm@35115: if xN = yN then Syntax.const @{syntax_const "_Collect"} $ x $ A $ P else M wenzelm@35115: | _ => M wenzelm@35115: end nipkow@13763: end; wenzelm@35115: in [(@{const_syntax Collect}, setcompr_tr')] end; wenzelm@11979: *} wenzelm@11979: haftmann@32117: setup {* haftmann@32117: let haftmann@32117: val unfold_bex_tac = unfold_tac @{thms "Bex_def"}; haftmann@32117: fun prove_bex_tac ss = unfold_bex_tac ss THEN Quantifier1.prove_one_point_ex_tac; haftmann@32117: val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac; haftmann@32117: val unfold_ball_tac = unfold_tac @{thms "Ball_def"}; haftmann@32117: fun prove_ball_tac ss = unfold_ball_tac ss THEN Quantifier1.prove_one_point_all_tac; haftmann@32117: val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac; haftmann@32117: val defBEX_regroup = Simplifier.simproc @{theory} haftmann@32117: "defined BEX" ["EX x:A. P x & Q x"] rearrange_bex; haftmann@32117: val defBALL_regroup = Simplifier.simproc @{theory} haftmann@32117: "defined BALL" ["ALL x:A. P x --> Q x"] rearrange_ball; haftmann@32117: in haftmann@32117: Simplifier.map_simpset (fn ss => ss addsimprocs [defBALL_regroup, defBEX_regroup]) haftmann@32117: end haftmann@32117: *} haftmann@32117: 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: 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: haftmann@32117: ML {* haftmann@32117: structure Simpdata = haftmann@32117: struct haftmann@32117: haftmann@32117: open Simpdata; haftmann@32117: haftmann@32117: val mksimps_pairs = [(@{const_name Ball}, @{thms bspec})] @ mksimps_pairs; haftmann@32117: haftmann@32117: end; haftmann@32117: haftmann@32117: open Simpdata; haftmann@32117: *} haftmann@32117: haftmann@32117: declaration {* fn _ => haftmann@32117: Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) haftmann@32117: *} haftmann@32117: haftmann@32117: lemma ballE [elim]: "ALL x:A. P x ==> (P x ==> Q) ==> (x ~: A ==> Q) ==> Q" haftmann@32117: by (unfold Ball_def) blast haftmann@32117: 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@11979: haftmann@32081: text {* 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@32081: subsection {* Basic operations *} haftmann@32081: haftmann@30531: subsubsection {* Subsets *} haftmann@30531: paulson@33022: lemma subsetI [intro!]: "(\x. x \ A \ x \ B) \ A \ B" haftmann@32888: unfolding mem_def by (rule le_funI, rule le_boolI) 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@32888: unfolding mem_def by (erule le_funE, erule le_boolE) haftmann@30531: -- {* Rule in Modus Ponens style. *} haftmann@30531: blanchet@35828: lemma rev_subsetD [no_atp,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: blanchet@35828: lemma subsetCE [no_atp,elim]: "A \ B ==> (c \ A ==> P) ==> (c \ B ==> P) ==> P" haftmann@30531: -- {* Classical elimination rule. *} haftmann@32888: unfolding mem_def by (blast dest: le_funE le_boolE) haftmann@30531: blanchet@35828: lemma subset_eq [no_atp]: "A \ B = (\x\A. x \ B)" by blast wenzelm@2388: blanchet@35828: lemma contra_subsetD [no_atp]: "A \ B ==> c \ B ==> c \ A" haftmann@30531: by blast haftmann@30531: paulson@33022: lemma subset_refl [simp]: "A \ A" haftmann@32081: by (fact order_refl) haftmann@30531: haftmann@30531: lemma subset_trans: "A \ B ==> B \ C ==> A \ C" haftmann@32081: by (fact order_trans) haftmann@32081: haftmann@32081: lemma set_rev_mp: "x:A ==> A \ B ==> x:B" haftmann@32081: by (rule subsetD) haftmann@32081: haftmann@32081: lemma set_mp: "A \ B ==> x:A ==> x:B" haftmann@32081: by (rule subsetD) haftmann@32081: paulson@33044: lemma eq_mem_trans: "a=b ==> b \ A ==> a \ A" paulson@33044: by simp paulson@33044: haftmann@32081: lemmas basic_trans_rules [trans] = paulson@33044: order_trans_rules set_rev_mp set_mp eq_mem_trans 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: 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" krauss@34209: by simp haftmann@30531: haftmann@30531: lemma equalityD2: "A = B ==> B \ A" krauss@34209: by simp 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" krauss@34209: by simp 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@32264: abbreviation UNIV :: "'a set" where haftmann@32264: "UNIV \ top" haftmann@32135: haftmann@32135: lemma UNIV_def: haftmann@32117: "UNIV = {x. True}" haftmann@32264: by (simp add: top_fun_eq top_bool_eq Collect_def) haftmann@32081: 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@32135: lemma empty_def: haftmann@32135: "{} = {x. False}" haftmann@32264: by (simp add: bot_fun_eq bot_bool_eq Collect_def) haftmann@32135: 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@32082: -- {* Use for reasoning about disjointness: @{text "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@32077: definition Pow :: "'a set => 'a set set" where haftmann@32077: Pow_def: "Pow A = {B. B \ A}" haftmann@32077: 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" krauss@34209: by simp 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@32683: abbreviation union :: "'a set \ 'a set \ 'a set" (infixl "Un" 65) where haftmann@32683: "op Un \ sup" haftmann@32081: haftmann@32081: notation (xsymbols) haftmann@32135: union (infixl "\" 65) haftmann@32081: haftmann@32081: notation (HTML output) haftmann@32135: union (infixl "\" 65) haftmann@32135: haftmann@32135: lemma Un_def: haftmann@32135: "A \ B = {x. x \ A \ x \ B}" haftmann@32683: by (simp add: sup_fun_eq sup_bool_eq Collect_def mem_def) haftmann@32081: 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@32117: lemma insert_def: "insert a B = {x. x = a} \ B" haftmann@32081: by (simp add: Collect_def mem_def insert_compr Un_def) haftmann@32081: haftmann@32081: lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" haftmann@32683: by (fact mono_sup) haftmann@32081: haftmann@30531: haftmann@30531: subsubsection {* Binary intersection -- Int *} haftmann@30531: haftmann@32683: abbreviation inter :: "'a set \ 'a set \ 'a set" (infixl "Int" 70) where haftmann@32683: "op Int \ inf" haftmann@32081: haftmann@32081: notation (xsymbols) haftmann@32135: inter (infixl "\" 70) haftmann@32081: haftmann@32081: notation (HTML output) haftmann@32135: inter (infixl "\" 70) haftmann@32135: haftmann@32135: lemma Int_def: haftmann@32135: "A \ B = {x. x \ A \ x \ B}" haftmann@32683: by (simp add: inf_fun_eq inf_bool_eq Collect_def mem_def) haftmann@32081: 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@32081: lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" haftmann@32683: by (fact mono_inf) haftmann@32081: 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@31456: subsubsection {* Augmenting a set -- @{const 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: blanchet@35828: lemma singletonI [intro!,no_atp]: "a : {a}" haftmann@30531: -- {* Redundant? But unlike @{text insertCI}, it proves the subgoal immediately! *} haftmann@30531: by (rule insertI1) haftmann@30531: blanchet@35828: lemma singletonD [dest!,no_atp]: "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: blanchet@35828: lemma singleton_insert_inj_eq [iff,no_atp]: haftmann@30531: "({b} = insert a A) = (a = b & A \ {b})" haftmann@30531: by blast haftmann@30531: blanchet@35828: lemma singleton_insert_inj_eq' [iff,no_atp]: 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: haftmann@32077: subsubsection {* Image of a set under a function *} haftmann@32077: haftmann@32077: text {* haftmann@32077: Frequently @{term b} does not have the syntactic form of @{term "f x"}. haftmann@32077: *} haftmann@32077: haftmann@32077: definition image :: "('a => 'b) => 'a set => 'b set" (infixr "`" 90) where blanchet@35828: image_def [no_atp]: "f ` A = {y. EX x:A. y = f(x)}" haftmann@32077: haftmann@32077: abbreviation haftmann@32077: range :: "('a => 'b) => 'b set" where -- "of function" haftmann@32077: "range f == f ` UNIV" haftmann@32077: haftmann@32077: lemma image_eqI [simp, intro]: "b = f x ==> x:A ==> b : f`A" haftmann@32077: by (unfold image_def) blast haftmann@32077: haftmann@32077: lemma imageI: "x : A ==> f x : f ` A" haftmann@32077: by (rule image_eqI) (rule refl) haftmann@32077: haftmann@32077: lemma rev_image_eqI: "x:A ==> b = f x ==> b : f`A" haftmann@32077: -- {* This version's more effective when we already have the haftmann@32077: required @{term x}. *} haftmann@32077: by (unfold image_def) blast haftmann@32077: haftmann@32077: lemma imageE [elim!]: haftmann@32077: "b : (%x. f x)`A ==> (!!x. b = f x ==> x:A ==> P) ==> P" haftmann@32077: -- {* The eta-expansion gives variable-name preservation. *} haftmann@32077: by (unfold image_def) blast haftmann@32077: haftmann@32077: lemma image_Un: "f`(A Un B) = f`A Un f`B" haftmann@32077: by blast haftmann@32077: haftmann@32077: lemma image_iff: "(z : f`A) = (EX x:A. z = f x)" haftmann@32077: by blast haftmann@32077: haftmann@32077: lemma image_subset_iff: "(f`A \ B) = (\x\A. f x \ B)" haftmann@32077: -- {* This rewrite rule would confuse users if made default. *} haftmann@32077: by blast haftmann@32077: haftmann@32077: lemma subset_image_iff: "(B \ f`A) = (EX AA. AA \ A & B = f`AA)" haftmann@32077: apply safe haftmann@32077: prefer 2 apply fast haftmann@32077: apply (rule_tac x = "{a. a : A & f a : B}" in exI, fast) haftmann@32077: done haftmann@32077: haftmann@32077: lemma image_subsetI: "(!!x. x \ A ==> f x \ B) ==> f`A \ B" haftmann@32077: -- {* Replaces the three steps @{text subsetI}, @{text imageE}, haftmann@32077: @{text hypsubst}, but breaks too many existing proofs. *} haftmann@32077: by blast wenzelm@11979: wenzelm@11979: text {* haftmann@32077: \medskip Range of a function -- just a translation for image! haftmann@32077: *} haftmann@32077: haftmann@32077: lemma range_eqI: "b = f x ==> b \ range f" haftmann@32077: by simp haftmann@32077: haftmann@32077: lemma rangeI: "f x \ range f" haftmann@32077: by simp haftmann@32077: haftmann@32077: lemma rangeE [elim?]: "b \ range (\x. f x) ==> (!!x. b = f x ==> P) ==> P" haftmann@32077: by blast haftmann@32077: haftmann@32077: haftmann@32117: subsubsection {* Some rules with @{text "if"} *} haftmann@32081: haftmann@32081: text{* Elimination of @{text"{x. \ & x=t & \}"}. *} haftmann@32081: haftmann@32081: lemma Collect_conv_if: "{x. x=a & P x} = (if P a then {a} else {})" haftmann@32117: by auto haftmann@32081: haftmann@32081: lemma Collect_conv_if2: "{x. a=x & P x} = (if P a then {a} else {})" haftmann@32117: by auto haftmann@32081: haftmann@32081: text {* haftmann@32081: Rewrite rules for boolean case-splitting: faster than @{text haftmann@32081: "split_if [split]"}. haftmann@32081: *} haftmann@32081: haftmann@32081: lemma split_if_eq1: "((if Q then x else y) = b) = ((Q --> x = b) & (~ Q --> y = b))" haftmann@32081: by (rule split_if) haftmann@32081: haftmann@32081: lemma split_if_eq2: "(a = (if Q then x else y)) = ((Q --> a = x) & (~ Q --> a = y))" haftmann@32081: by (rule split_if) haftmann@32081: haftmann@32081: text {* haftmann@32081: Split ifs on either side of the membership relation. Not for @{text haftmann@32081: "[simp]"} -- can cause goals to blow up! haftmann@32081: *} haftmann@32081: haftmann@32081: lemma split_if_mem1: "((if Q then x else y) : b) = ((Q --> x : b) & (~ Q --> y : b))" haftmann@32081: by (rule split_if) haftmann@32081: haftmann@32081: lemma split_if_mem2: "(a : (if Q then x else y)) = ((Q --> a : x) & (~ Q --> a : y))" haftmann@32081: by (rule split_if [where P="%S. a : S"]) haftmann@32081: haftmann@32081: lemmas split_ifs = if_bool_eq_conj split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2 haftmann@32081: haftmann@32081: (*Would like to add these, but the existing code only searches for the haftmann@37677: outer-level constant, which in this case is just Set.member; we instead need haftmann@32081: to use term-nets to associate patterns with rules. Also, if a rule fails to haftmann@32081: apply, then the formula should be kept. haftmann@34974: [("uminus", Compl_iff RS iffD1), ("minus", [Diff_iff RS iffD1]), haftmann@32081: ("Int", [IntD1,IntD2]), haftmann@32081: ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])] haftmann@32081: *) haftmann@32081: haftmann@32081: haftmann@32135: subsection {* Further operations and lemmas *} haftmann@32135: haftmann@32135: subsubsection {* The ``proper subset'' relation *} haftmann@32135: blanchet@35828: lemma psubsetI [intro!,no_atp]: "A \ B ==> A \ B ==> A \ B" haftmann@32135: by (unfold less_le) blast haftmann@32135: blanchet@35828: lemma psubsetE [elim!,no_atp]: haftmann@32135: "[|A \ B; [|A \ B; ~ (B\A)|] ==> R|] ==> R" haftmann@32135: by (unfold less_le) blast haftmann@32135: haftmann@32135: lemma psubset_insert_iff: haftmann@32135: "(A \ insert x B) = (if x \ B then A \ B else if x \ A then A - {x} \ B else A \ B)" haftmann@32135: by (auto simp add: less_le subset_insert_iff) haftmann@32135: haftmann@32135: lemma psubset_eq: "(A \ B) = (A \ B & A \ B)" haftmann@32135: by (simp only: less_le) haftmann@32135: haftmann@32135: lemma psubset_imp_subset: "A \ B ==> A \ B" haftmann@32135: by (simp add: psubset_eq) haftmann@32135: haftmann@32135: lemma psubset_trans: "[| A \ B; B \ C |] ==> A \ C" haftmann@32135: apply (unfold less_le) haftmann@32135: apply (auto dest: subset_antisym) haftmann@32135: done haftmann@32135: haftmann@32135: lemma psubsetD: "[| A \ B; c \ A |] ==> c \ B" haftmann@32135: apply (unfold less_le) haftmann@32135: apply (auto dest: subsetD) haftmann@32135: done haftmann@32135: haftmann@32135: lemma psubset_subset_trans: "A \ B ==> B \ C ==> A \ C" haftmann@32135: by (auto simp add: psubset_eq) haftmann@32135: haftmann@32135: lemma subset_psubset_trans: "A \ B ==> B \ C ==> A \ C" haftmann@32135: by (auto simp add: psubset_eq) haftmann@32135: haftmann@32135: lemma psubset_imp_ex_mem: "A \ B ==> \b. b \ (B - A)" haftmann@32135: by (unfold less_le) blast haftmann@32135: haftmann@32135: lemma atomize_ball: haftmann@32135: "(!!x. x \ A ==> P x) == Trueprop (\x\A. P x)" haftmann@32135: by (simp only: Ball_def atomize_all atomize_imp) haftmann@32135: haftmann@32135: lemmas [symmetric, rulify] = atomize_ball haftmann@32135: and [symmetric, defn] = atomize_ball haftmann@32135: haftmann@32135: subsubsection {* Derived rules involving subsets. *} haftmann@32135: haftmann@32135: text {* @{text insert}. *} haftmann@32135: haftmann@32135: lemma subset_insertI: "B \ insert a B" haftmann@32135: by (rule subsetI) (erule insertI2) haftmann@32135: haftmann@32135: lemma subset_insertI2: "A \ B \ A \ insert b B" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma subset_insert: "x \ A ==> (A \ insert x B) = (A \ B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip Finite Union -- the least upper bound of two sets. *} haftmann@32135: haftmann@32135: lemma Un_upper1: "A \ A \ B" huffman@36009: by (fact sup_ge1) haftmann@32135: haftmann@32135: lemma Un_upper2: "B \ A \ B" huffman@36009: by (fact sup_ge2) haftmann@32135: haftmann@32135: lemma Un_least: "A \ C ==> B \ C ==> A \ B \ C" huffman@36009: by (fact sup_least) haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip Finite Intersection -- the greatest lower bound of two sets. *} haftmann@32135: haftmann@32135: lemma Int_lower1: "A \ B \ A" huffman@36009: by (fact inf_le1) haftmann@32135: haftmann@32135: lemma Int_lower2: "A \ B \ B" huffman@36009: by (fact inf_le2) haftmann@32135: haftmann@32135: lemma Int_greatest: "C \ A ==> C \ B ==> C \ A \ B" huffman@36009: by (fact inf_greatest) haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip Set difference. *} haftmann@32135: haftmann@32135: lemma Diff_subset: "A - B \ A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_subset_conv: "(A - B \ C) = (A \ B \ C)" haftmann@32135: by blast haftmann@32135: haftmann@32135: haftmann@32135: subsubsection {* Equalities involving union, intersection, inclusion, etc. *} haftmann@32135: haftmann@32135: text {* @{text "{}"}. *} haftmann@32135: haftmann@32135: lemma Collect_const [simp]: "{s. P} = (if P then UNIV else {})" haftmann@32135: -- {* supersedes @{text "Collect_False_empty"} *} haftmann@32135: by auto haftmann@32135: haftmann@32135: lemma subset_empty [simp]: "(A \ {}) = (A = {})" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma not_psubset_empty [iff]: "\ (A < {})" haftmann@32135: by (unfold less_le) blast haftmann@32135: haftmann@32135: lemma Collect_empty_eq [simp]: "(Collect P = {}) = (\x. \ P x)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma empty_Collect_eq [simp]: "({} = Collect P) = (\x. \ P x)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Collect_neg_eq: "{x. \ P x} = - {x. P x}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Collect_disj_eq: "{x. P x | Q x} = {x. P x} \ {x. Q x}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Collect_imp_eq: "{x. P x --> Q x} = -{x. P x} \ {x. Q x}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Collect_conj_eq: "{x. P x & Q x} = {x. P x} \ {x. Q x}" haftmann@32135: by blast haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip @{text insert}. *} haftmann@32135: haftmann@32135: lemma insert_is_Un: "insert a A = {a} Un A" haftmann@32135: -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a {}"} *} haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma insert_not_empty [simp]: "insert a A \ {}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemmas empty_not_insert = insert_not_empty [symmetric, standard] haftmann@32135: declare empty_not_insert [simp] haftmann@32135: haftmann@32135: lemma insert_absorb: "a \ A ==> insert a A = A" haftmann@32135: -- {* @{text "[simp]"} causes recursive calls when there are nested inserts *} haftmann@32135: -- {* with \emph{quadratic} running time *} haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma insert_absorb2 [simp]: "insert x (insert x A) = insert x A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma insert_commute: "insert x (insert y A) = insert y (insert x A)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma insert_subset [simp]: "(insert x A \ B) = (x \ B & A \ B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma mk_disjoint_insert: "a \ A ==> \B. A = insert a B & a \ B" haftmann@32135: -- {* use new @{text B} rather than @{text "A - {a}"} to avoid infinite unfolding *} haftmann@32135: apply (rule_tac x = "A - {a}" in exI, blast) haftmann@32135: done haftmann@32135: haftmann@32135: lemma insert_Collect: "insert a (Collect P) = {u. u \ a --> P u}" haftmann@32135: by auto haftmann@32135: haftmann@32135: lemma insert_inter_insert[simp]: "insert a A \ insert a B = insert a (A \ B)" haftmann@32135: by blast haftmann@32135: blanchet@35828: lemma insert_disjoint [simp,no_atp]: haftmann@32135: "(insert a A \ B = {}) = (a \ B \ A \ B = {})" haftmann@32135: "({} = insert a A \ B) = (a \ B \ {} = A \ B)" haftmann@32135: by auto haftmann@32135: blanchet@35828: lemma disjoint_insert [simp,no_atp]: haftmann@32135: "(B \ insert a A = {}) = (a \ B \ B \ A = {})" haftmann@32135: "({} = A \ insert b B) = (b \ A \ {} = A \ B)" haftmann@32135: by auto haftmann@32135: haftmann@32135: text {* \medskip @{text image}. *} haftmann@32135: haftmann@32135: lemma image_empty [simp]: "f`{} = {}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma image_insert [simp]: "f ` insert a B = insert (f a) (f`B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma image_constant: "x \ A ==> (\x. c) ` A = {c}" haftmann@32135: by auto haftmann@32135: haftmann@32135: lemma image_constant_conv: "(%x. c) ` A = (if A = {} then {} else {c})" haftmann@32135: by auto haftmann@32135: haftmann@32135: lemma image_image: "f ` (g ` A) = (\x. f (g x)) ` A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma insert_image [simp]: "x \ A ==> insert (f x) (f`A) = f`A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma image_is_empty [iff]: "(f`A = {}) = (A = {})" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma empty_is_image[iff]: "({} = f ` A) = (A = {})" haftmann@32135: by blast haftmann@32135: haftmann@32135: blanchet@35828: lemma image_Collect [no_atp]: "f ` {x. P x} = {f x | x. P x}" haftmann@32135: -- {* NOT suitable as a default simprule: the RHS isn't simpler than the LHS, haftmann@32135: with its implicit quantifier and conjunction. Also image enjoys better haftmann@32135: equational properties than does the RHS. *} haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma if_image_distrib [simp]: haftmann@32135: "(\x. if P x then f x else g x) ` S haftmann@32135: = (f ` (S \ {x. P x})) \ (g ` (S \ {x. \ P x}))" haftmann@32135: by (auto simp add: image_def) haftmann@32135: haftmann@32135: lemma image_cong: "M = N ==> (!!x. x \ N ==> f x = g x) ==> f`M = g`N" haftmann@32135: by (simp add: image_def) haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip @{text range}. *} haftmann@32135: blanchet@35828: lemma full_SetCompr_eq [no_atp]: "{u. \x. u = f x} = range f" haftmann@32135: by auto haftmann@32135: haftmann@32135: lemma range_composition: "range (\x. f (g x)) = f`range g" haftmann@32135: by (subst image_image, simp) haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip @{text Int} *} haftmann@32135: haftmann@32135: lemma Int_absorb [simp]: "A \ A = A" huffman@36009: by (fact inf_idem) haftmann@32135: haftmann@32135: lemma Int_left_absorb: "A \ (A \ B) = A \ B" huffman@36009: by (fact inf_left_idem) haftmann@32135: haftmann@32135: lemma Int_commute: "A \ B = B \ A" huffman@36009: by (fact inf_commute) haftmann@32135: haftmann@32135: lemma Int_left_commute: "A \ (B \ C) = B \ (A \ C)" huffman@36009: by (fact inf_left_commute) haftmann@32135: haftmann@32135: lemma Int_assoc: "(A \ B) \ C = A \ (B \ C)" huffman@36009: by (fact inf_assoc) haftmann@32135: haftmann@32135: lemmas Int_ac = Int_assoc Int_left_absorb Int_commute Int_left_commute haftmann@32135: -- {* Intersection is an AC-operator *} haftmann@32135: haftmann@32135: lemma Int_absorb1: "B \ A ==> A \ B = B" huffman@36009: by (fact inf_absorb2) haftmann@32135: haftmann@32135: lemma Int_absorb2: "A \ B ==> A \ B = A" huffman@36009: by (fact inf_absorb1) haftmann@32135: haftmann@32135: lemma Int_empty_left [simp]: "{} \ B = {}" huffman@36009: by (fact inf_bot_left) haftmann@32135: haftmann@32135: lemma Int_empty_right [simp]: "A \ {} = {}" huffman@36009: by (fact inf_bot_right) haftmann@32135: haftmann@32135: lemma disjoint_eq_subset_Compl: "(A \ B = {}) = (A \ -B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma disjoint_iff_not_equal: "(A \ B = {}) = (\x\A. \y\B. x \ y)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Int_UNIV_left [simp]: "UNIV \ B = B" huffman@36009: by (fact inf_top_left) haftmann@32135: haftmann@32135: lemma Int_UNIV_right [simp]: "A \ UNIV = A" huffman@36009: by (fact inf_top_right) haftmann@32135: haftmann@32135: lemma Int_Un_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" huffman@36009: by (fact inf_sup_distrib1) haftmann@32135: haftmann@32135: lemma Int_Un_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" huffman@36009: by (fact inf_sup_distrib2) haftmann@32135: blanchet@35828: lemma Int_UNIV [simp,no_atp]: "(A \ B = UNIV) = (A = UNIV & B = UNIV)" huffman@36009: by (fact inf_eq_top_iff) haftmann@32135: haftmann@32135: lemma Int_subset_iff [simp]: "(C \ A \ B) = (C \ A & C \ B)" huffman@36009: by (fact le_inf_iff) haftmann@32135: haftmann@32135: lemma Int_Collect: "(x \ A \ {x. P x}) = (x \ A & P x)" haftmann@32135: by blast haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip @{text Un}. *} haftmann@32135: haftmann@32135: lemma Un_absorb [simp]: "A \ A = A" huffman@36009: by (fact sup_idem) haftmann@32135: haftmann@32135: lemma Un_left_absorb: "A \ (A \ B) = A \ B" huffman@36009: by (fact sup_left_idem) haftmann@32135: haftmann@32135: lemma Un_commute: "A \ B = B \ A" huffman@36009: by (fact sup_commute) haftmann@32135: haftmann@32135: lemma Un_left_commute: "A \ (B \ C) = B \ (A \ C)" huffman@36009: by (fact sup_left_commute) haftmann@32135: haftmann@32135: lemma Un_assoc: "(A \ B) \ C = A \ (B \ C)" huffman@36009: by (fact sup_assoc) haftmann@32135: haftmann@32135: lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute haftmann@32135: -- {* Union is an AC-operator *} haftmann@32135: haftmann@32135: lemma Un_absorb1: "A \ B ==> A \ B = B" huffman@36009: by (fact sup_absorb2) haftmann@32135: haftmann@32135: lemma Un_absorb2: "B \ A ==> A \ B = A" huffman@36009: by (fact sup_absorb1) haftmann@32135: haftmann@32135: lemma Un_empty_left [simp]: "{} \ B = B" huffman@36009: by (fact sup_bot_left) haftmann@32135: haftmann@32135: lemma Un_empty_right [simp]: "A \ {} = A" huffman@36009: by (fact sup_bot_right) haftmann@32135: haftmann@32135: lemma Un_UNIV_left [simp]: "UNIV \ B = UNIV" huffman@36009: by (fact sup_top_left) haftmann@32135: haftmann@32135: lemma Un_UNIV_right [simp]: "A \ UNIV = UNIV" huffman@36009: by (fact sup_top_right) haftmann@32135: haftmann@32135: lemma Un_insert_left [simp]: "(insert a B) \ C = insert a (B \ C)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Un_insert_right [simp]: "A \ (insert a B) = insert a (A \ B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Int_insert_left: haftmann@32135: "(insert a B) Int C = (if a \ C then insert a (B \ C) else B \ C)" haftmann@32135: by auto haftmann@32135: nipkow@32456: lemma Int_insert_left_if0[simp]: nipkow@32456: "a \ C \ (insert a B) Int C = B \ C" nipkow@32456: by auto nipkow@32456: nipkow@32456: lemma Int_insert_left_if1[simp]: nipkow@32456: "a \ C \ (insert a B) Int C = insert a (B Int C)" nipkow@32456: by auto nipkow@32456: haftmann@32135: lemma Int_insert_right: haftmann@32135: "A \ (insert a B) = (if a \ A then insert a (A \ B) else A \ B)" haftmann@32135: by auto haftmann@32135: nipkow@32456: lemma Int_insert_right_if0[simp]: nipkow@32456: "a \ A \ A Int (insert a B) = A Int B" nipkow@32456: by auto nipkow@32456: nipkow@32456: lemma Int_insert_right_if1[simp]: nipkow@32456: "a \ A \ A Int (insert a B) = insert a (A Int B)" nipkow@32456: by auto nipkow@32456: haftmann@32135: lemma Un_Int_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" huffman@36009: by (fact sup_inf_distrib1) haftmann@32135: haftmann@32135: lemma Un_Int_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" huffman@36009: by (fact sup_inf_distrib2) haftmann@32135: haftmann@32135: lemma Un_Int_crazy: haftmann@32135: "(A \ B) \ (B \ C) \ (C \ A) = (A \ B) \ (B \ C) \ (C \ A)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma subset_Un_eq: "(A \ B) = (A \ B = B)" huffman@36009: by (fact le_iff_sup) haftmann@32135: haftmann@32135: lemma Un_empty [iff]: "(A \ B = {}) = (A = {} & B = {})" huffman@36009: by (fact sup_eq_bot_iff) haftmann@32135: haftmann@32135: lemma Un_subset_iff [simp]: "(A \ B \ C) = (A \ C & B \ C)" huffman@36009: by (fact le_sup_iff) haftmann@32135: haftmann@32135: lemma Un_Diff_Int: "(A - B) \ (A \ B) = A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_Int2: "A \ C - B \ C = A \ C - B" haftmann@32135: by blast haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip Set complement *} haftmann@32135: haftmann@32135: lemma Compl_disjoint [simp]: "A \ -A = {}" huffman@36009: by (fact inf_compl_bot) haftmann@32135: haftmann@32135: lemma Compl_disjoint2 [simp]: "-A \ A = {}" huffman@36009: by (fact compl_inf_bot) haftmann@32135: haftmann@32135: lemma Compl_partition: "A \ -A = UNIV" huffman@36009: by (fact sup_compl_top) haftmann@32135: haftmann@32135: lemma Compl_partition2: "-A \ A = UNIV" huffman@36009: by (fact compl_sup_top) haftmann@32135: haftmann@32135: lemma double_complement [simp]: "- (-A) = (A::'a set)" huffman@36009: by (fact double_compl) haftmann@32135: haftmann@32135: lemma Compl_Un [simp]: "-(A \ B) = (-A) \ (-B)" huffman@36009: by (fact compl_sup) haftmann@32135: haftmann@32135: lemma Compl_Int [simp]: "-(A \ B) = (-A) \ (-B)" huffman@36009: by (fact compl_inf) haftmann@32135: haftmann@32135: lemma subset_Compl_self_eq: "(A \ -A) = (A = {})" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Un_Int_assoc_eq: "((A \ B) \ C = A \ (B \ C)) = (C \ A)" haftmann@32135: -- {* Halmos, Naive Set Theory, page 16. *} haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Compl_UNIV_eq [simp]: "-UNIV = {}" huffman@36009: by (fact compl_top_eq) haftmann@32135: haftmann@32135: lemma Compl_empty_eq [simp]: "-{} = UNIV" huffman@36009: by (fact compl_bot_eq) haftmann@32135: haftmann@32135: lemma Compl_subset_Compl_iff [iff]: "(-A \ -B) = (B \ A)" huffman@36009: by (fact compl_le_compl_iff) haftmann@32135: haftmann@32135: lemma Compl_eq_Compl_iff [iff]: "(-A = -B) = (A = (B::'a set))" huffman@36009: by (fact compl_eq_compl_iff) haftmann@32135: haftmann@32135: text {* \medskip Bounded quantifiers. haftmann@32135: haftmann@32135: The following are not added to the default simpset because haftmann@32135: (a) they duplicate the body and (b) there are no similar rules for @{text Int}. *} haftmann@32135: haftmann@32135: lemma ball_Un: "(\x \ A \ B. P x) = ((\x\A. P x) & (\x\B. P x))" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma bex_Un: "(\x \ A \ B. P x) = ((\x\A. P x) | (\x\B. P x))" haftmann@32135: by blast haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip Set difference. *} haftmann@32135: haftmann@32135: lemma Diff_eq: "A - B = A \ (-B)" haftmann@32135: by blast haftmann@32135: blanchet@35828: lemma Diff_eq_empty_iff [simp,no_atp]: "(A - B = {}) = (A \ B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_cancel [simp]: "A - A = {}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_idemp [simp]: "(A - B) - B = A - (B::'a set)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_triv: "A \ B = {} ==> A - B = A" haftmann@32135: by (blast elim: equalityE) haftmann@32135: haftmann@32135: lemma empty_Diff [simp]: "{} - A = {}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_empty [simp]: "A - {} = A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_UNIV [simp]: "A - UNIV = {}" haftmann@32135: by blast haftmann@32135: blanchet@35828: lemma Diff_insert0 [simp,no_atp]: "x \ A ==> A - insert x B = A - B" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_insert: "A - insert a B = A - B - {a}" haftmann@32135: -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"} *} haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_insert2: "A - insert a B = A - {a} - B" haftmann@32135: -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"} *} haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma insert_Diff_if: "insert x A - B = (if x \ B then A - B else insert x (A - B))" haftmann@32135: by auto haftmann@32135: haftmann@32135: lemma insert_Diff1 [simp]: "x \ B ==> insert x A - B = A - B" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma insert_Diff_single[simp]: "insert a (A - {a}) = insert a A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma insert_Diff: "a \ A ==> insert a (A - {a}) = A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_insert_absorb: "x \ A ==> (insert x A) - {x} = A" haftmann@32135: by auto haftmann@32135: haftmann@32135: lemma Diff_disjoint [simp]: "A \ (B - A) = {}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_partition: "A \ B ==> A \ (B - A) = B" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma double_diff: "A \ B ==> B \ C ==> B - (C - A) = A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Un_Diff_cancel [simp]: "A \ (B - A) = A \ B" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Un_Diff_cancel2 [simp]: "(B - A) \ A = B \ A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_Un: "A - (B \ C) = (A - B) \ (A - C)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_Int: "A - (B \ C) = (A - B) \ (A - C)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Un_Diff: "(A \ B) - C = (A - C) \ (B - C)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Int_Diff: "(A \ B) - C = A \ (B - C)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_Int_distrib: "C \ (A - B) = (C \ A) - (C \ B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_Int_distrib2: "(A - B) \ C = (A \ C) - (B \ C)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Diff_Compl [simp]: "A - (- B) = A \ B" haftmann@32135: by auto haftmann@32135: haftmann@32135: lemma Compl_Diff_eq [simp]: "- (A - B) = -A \ B" haftmann@32135: by blast haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip Quantification over type @{typ bool}. *} haftmann@32135: haftmann@32135: lemma bool_induct: "P True \ P False \ P x" haftmann@32135: by (cases x) auto haftmann@32135: haftmann@32135: lemma all_bool_eq: "(\b. P b) \ P True \ P False" haftmann@32135: by (auto intro: bool_induct) haftmann@32135: haftmann@32135: lemma bool_contrapos: "P x \ \ P False \ P True" haftmann@32135: by (cases x) auto haftmann@32135: haftmann@32135: lemma ex_bool_eq: "(\b. P b) \ P True \ P False" haftmann@32135: by (auto intro: bool_contrapos) haftmann@32135: haftmann@32135: text {* \medskip @{text Pow} *} haftmann@32135: haftmann@32135: lemma Pow_empty [simp]: "Pow {} = {{}}" haftmann@32135: by (auto simp add: Pow_def) haftmann@32135: haftmann@32135: lemma Pow_insert: "Pow (insert a A) = Pow A \ (insert a ` Pow A)" haftmann@32135: by (blast intro: image_eqI [where ?x = "u - {a}", standard]) haftmann@32135: haftmann@32135: lemma Pow_Compl: "Pow (- A) = {-B | B. A \ Pow B}" haftmann@32135: by (blast intro: exI [where ?x = "- u", standard]) haftmann@32135: haftmann@32135: lemma Pow_UNIV [simp]: "Pow UNIV = UNIV" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Un_Pow_subset: "Pow A \ Pow B \ Pow (A \ B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Pow_Int_eq [simp]: "Pow (A \ B) = Pow A \ Pow B" haftmann@32135: by blast haftmann@32135: haftmann@32135: haftmann@32135: text {* \medskip Miscellany. *} haftmann@32135: haftmann@32135: lemma set_eq_subset: "(A = B) = (A \ B & B \ A)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma subset_iff: "(A \ B) = (\t. t \ A --> t \ B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma subset_iff_psubset_eq: "(A \ B) = ((A \ B) | (A = B))" haftmann@32135: by (unfold less_le) blast haftmann@32135: haftmann@32135: lemma all_not_in_conv [simp]: "(\x. x \ A) = (A = {})" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma ex_in_conv: "(\x. x \ A) = (A \ {})" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma distinct_lemma: "f x \ f y ==> x \ y" haftmann@32135: by iprover haftmann@32135: haftmann@32135: haftmann@32135: subsubsection {* Monotonicity of various operations *} haftmann@32135: haftmann@32135: lemma image_mono: "A \ B ==> f`A \ f`B" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Pow_mono: "A \ B ==> Pow A \ Pow B" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma insert_mono: "C \ D ==> insert a C \ insert a D" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Un_mono: "A \ C ==> B \ D ==> A \ B \ C \ D" huffman@36009: by (fact sup_mono) haftmann@32135: haftmann@32135: lemma Int_mono: "A \ C ==> B \ D ==> A \ B \ C \ D" huffman@36009: by (fact inf_mono) haftmann@32135: haftmann@32135: lemma Diff_mono: "A \ C ==> D \ B ==> A - B \ C - D" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Compl_anti_mono: "A \ B ==> -B \ -A" huffman@36009: by (fact compl_mono) haftmann@32135: haftmann@32135: text {* \medskip Monotonicity of implications. *} haftmann@32135: haftmann@32135: lemma in_mono: "A \ B ==> x \ A --> x \ B" haftmann@32135: apply (rule impI) haftmann@32135: apply (erule subsetD, assumption) haftmann@32135: done haftmann@32135: haftmann@32135: lemma conj_mono: "P1 --> Q1 ==> P2 --> Q2 ==> (P1 & P2) --> (Q1 & Q2)" haftmann@32135: by iprover haftmann@32135: haftmann@32135: lemma disj_mono: "P1 --> Q1 ==> P2 --> Q2 ==> (P1 | P2) --> (Q1 | Q2)" haftmann@32135: by iprover haftmann@32135: haftmann@32135: lemma imp_mono: "Q1 --> P1 ==> P2 --> Q2 ==> (P1 --> P2) --> (Q1 --> Q2)" haftmann@32135: by iprover haftmann@32135: haftmann@32135: lemma imp_refl: "P --> P" .. haftmann@32135: berghofe@33935: lemma not_mono: "Q --> P ==> ~ P --> ~ Q" berghofe@33935: by iprover berghofe@33935: haftmann@32135: lemma ex_mono: "(!!x. P x --> Q x) ==> (EX x. P x) --> (EX x. Q x)" haftmann@32135: by iprover haftmann@32135: haftmann@32135: lemma all_mono: "(!!x. P x --> Q x) ==> (ALL x. P x) --> (ALL x. Q x)" haftmann@32135: by iprover haftmann@32135: haftmann@32135: lemma Collect_mono: "(!!x. P x --> Q x) ==> Collect P \ Collect Q" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Int_Collect_mono: haftmann@32135: "A \ B ==> (!!x. x \ A ==> P x --> Q x) ==> A \ Collect P \ B \ Collect Q" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemmas basic_monos = haftmann@32135: subset_refl imp_refl disj_mono conj_mono haftmann@32135: ex_mono Collect_mono in_mono haftmann@32135: haftmann@32135: lemma eq_to_mono: "a = b ==> c = d ==> b --> d ==> a --> c" haftmann@32135: by iprover haftmann@32135: haftmann@32135: haftmann@32135: subsubsection {* Inverse image of a function *} haftmann@32135: haftmann@35416: definition vimage :: "('a => 'b) => 'b set => 'a set" (infixr "-`" 90) where haftmann@37767: "f -` B == {x. f x : B}" haftmann@32135: haftmann@32135: lemma vimage_eq [simp]: "(a : f -` B) = (f a : B)" haftmann@32135: by (unfold vimage_def) blast haftmann@32135: haftmann@32135: lemma vimage_singleton_eq: "(a : f -` {b}) = (f a = b)" haftmann@32135: by simp haftmann@32135: haftmann@32135: lemma vimageI [intro]: "f a = b ==> b:B ==> a : f -` B" haftmann@32135: by (unfold vimage_def) blast haftmann@32135: haftmann@32135: lemma vimageI2: "f a : A ==> a : f -` A" haftmann@32135: by (unfold vimage_def) fast haftmann@32135: haftmann@32135: lemma vimageE [elim!]: "a: f -` B ==> (!!x. f a = x ==> x:B ==> P) ==> P" haftmann@32135: by (unfold vimage_def) blast haftmann@32135: haftmann@32135: lemma vimageD: "a : f -` A ==> f a : A" haftmann@32135: by (unfold vimage_def) fast haftmann@32135: haftmann@32135: lemma vimage_empty [simp]: "f -` {} = {}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma vimage_Compl: "f -` (-A) = -(f -` A)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma vimage_Un [simp]: "f -` (A Un B) = (f -` A) Un (f -` B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma vimage_Int [simp]: "f -` (A Int B) = (f -` A) Int (f -` B)" haftmann@32135: by fast haftmann@32135: haftmann@32135: lemma vimage_Collect_eq [simp]: "f -` Collect P = {y. P (f y)}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma vimage_Collect: "(!!x. P (f x) = Q x) ==> f -` (Collect P) = Collect Q" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma vimage_insert: "f-`(insert a B) = (f-`{a}) Un (f-`B)" haftmann@32135: -- {* NOT suitable for rewriting because of the recurrence of @{term "{a}"}. *} haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma vimage_Diff: "f -` (A - B) = (f -` A) - (f -` B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma vimage_UNIV [simp]: "f -` UNIV = UNIV" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma vimage_mono: "A \ B ==> f -` A \ f -` B" haftmann@32135: -- {* monotonicity *} haftmann@32135: by blast haftmann@32135: blanchet@35828: lemma vimage_image_eq [no_atp]: "f -` (f ` A) = {y. EX x:A. f x = f y}" haftmann@32135: by (blast intro: sym) haftmann@32135: haftmann@32135: lemma image_vimage_subset: "f ` (f -` A) <= A" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma image_vimage_eq [simp]: "f ` (f -` A) = A Int range f" haftmann@32135: by blast haftmann@32135: paulson@33533: lemma vimage_const [simp]: "((\x. c) -` A) = (if c \ A then UNIV else {})" paulson@33533: by auto paulson@33533: paulson@33533: lemma vimage_if [simp]: "((\x. if x \ B then c else d) -` A) = paulson@33533: (if c \ A then (if d \ A then UNIV else B) paulson@33533: else if d \ A then -B else {})" paulson@33533: by (auto simp add: vimage_def) paulson@33533: hoelzl@35576: lemma vimage_inter_cong: hoelzl@35576: "(\ w. w \ S \ f w = g w) \ f -` y \ S = g -` y \ S" hoelzl@35576: by auto hoelzl@35576: haftmann@32135: lemma image_Int_subset: "f`(A Int B) <= f`A Int f`B" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma image_diff_subset: "f`A - f`B <= f`(A - B)" haftmann@32135: by blast haftmann@32135: haftmann@32135: haftmann@32135: subsubsection {* Getting the Contents of a Singleton Set *} haftmann@32135: haftmann@32135: definition contents :: "'a set \ 'a" where haftmann@37767: "contents X = (THE x. X = {x})" haftmann@32135: haftmann@32135: lemma contents_eq [simp]: "contents {x} = x" haftmann@32135: by (simp add: contents_def) haftmann@32135: haftmann@32135: haftmann@32135: subsubsection {* Least value operator *} haftmann@32135: haftmann@32135: lemma Least_mono: haftmann@32135: "mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y haftmann@32135: ==> (LEAST y. y : f ` S) = f (LEAST x. x : S)" haftmann@32135: -- {* Courtesy of Stephan Merz *} haftmann@32135: apply clarify haftmann@32135: apply (erule_tac P = "%x. x : S" in LeastI2_order, fast) haftmann@32135: apply (rule LeastI2_order) haftmann@32135: apply (auto elim: monoD intro!: order_antisym) haftmann@32135: done haftmann@32135: haftmann@32135: subsection {* Misc *} haftmann@32135: haftmann@32135: text {* Rudimentary code generation *} haftmann@32135: haftmann@32135: lemma insert_code [code]: "insert y A x \ y = x \ A x" haftmann@32135: by (auto simp add: insert_compr Collect_def mem_def) haftmann@32135: haftmann@32135: lemma vimage_code [code]: "(f -` A) x = A (f x)" haftmann@32135: by (simp add: vimage_def Collect_def mem_def) haftmann@32135: haftmann@37677: hide_const (open) member haftmann@32135: haftmann@32135: text {* Misc theorem and ML bindings *} haftmann@32135: haftmann@32135: lemmas equalityI = subset_antisym haftmann@32135: haftmann@32135: ML {* haftmann@32135: val Ball_def = @{thm Ball_def} haftmann@32135: val Bex_def = @{thm Bex_def} haftmann@32135: val CollectD = @{thm CollectD} haftmann@32135: val CollectE = @{thm CollectE} haftmann@32135: val CollectI = @{thm CollectI} haftmann@32135: val Collect_conj_eq = @{thm Collect_conj_eq} haftmann@32135: val Collect_mem_eq = @{thm Collect_mem_eq} haftmann@32135: val IntD1 = @{thm IntD1} haftmann@32135: val IntD2 = @{thm IntD2} haftmann@32135: val IntE = @{thm IntE} haftmann@32135: val IntI = @{thm IntI} haftmann@32135: val Int_Collect = @{thm Int_Collect} haftmann@32135: val UNIV_I = @{thm UNIV_I} haftmann@32135: val UNIV_witness = @{thm UNIV_witness} haftmann@32135: val UnE = @{thm UnE} haftmann@32135: val UnI1 = @{thm UnI1} haftmann@32135: val UnI2 = @{thm UnI2} haftmann@32135: val ballE = @{thm ballE} haftmann@32135: val ballI = @{thm ballI} haftmann@32135: val bexCI = @{thm bexCI} haftmann@32135: val bexE = @{thm bexE} haftmann@32135: val bexI = @{thm bexI} haftmann@32135: val bex_triv = @{thm bex_triv} haftmann@32135: val bspec = @{thm bspec} haftmann@32135: val contra_subsetD = @{thm contra_subsetD} haftmann@32135: val distinct_lemma = @{thm distinct_lemma} haftmann@32135: val eq_to_mono = @{thm eq_to_mono} haftmann@32135: val equalityCE = @{thm equalityCE} haftmann@32135: val equalityD1 = @{thm equalityD1} haftmann@32135: val equalityD2 = @{thm equalityD2} haftmann@32135: val equalityE = @{thm equalityE} haftmann@32135: val equalityI = @{thm equalityI} haftmann@32135: val imageE = @{thm imageE} haftmann@32135: val imageI = @{thm imageI} haftmann@32135: val image_Un = @{thm image_Un} haftmann@32135: val image_insert = @{thm image_insert} haftmann@32135: val insert_commute = @{thm insert_commute} haftmann@32135: val insert_iff = @{thm insert_iff} haftmann@32135: val mem_Collect_eq = @{thm mem_Collect_eq} haftmann@32135: val rangeE = @{thm rangeE} haftmann@32135: val rangeI = @{thm rangeI} haftmann@32135: val range_eqI = @{thm range_eqI} haftmann@32135: val subsetCE = @{thm subsetCE} haftmann@32135: val subsetD = @{thm subsetD} haftmann@32135: val subsetI = @{thm subsetI} haftmann@32135: val subset_refl = @{thm subset_refl} haftmann@32135: val subset_trans = @{thm subset_trans} haftmann@32135: val vimageD = @{thm vimageD} haftmann@32135: val vimageE = @{thm vimageE} haftmann@32135: val vimageI = @{thm vimageI} haftmann@32135: val vimageI2 = @{thm vimageI2} haftmann@32135: val vimage_Collect = @{thm vimage_Collect} haftmann@32135: val vimage_Int = @{thm vimage_Int} haftmann@32135: val vimage_Un = @{thm vimage_Un} haftmann@32135: *} haftmann@32135: haftmann@32077: end