# HG changeset patch # User berghofe # Date 1170865539 -3600 # Node ID 476604be7d88ef45a9b422782065d19f1c5d080e # Parent 0967b03844b5cca78f92bec50c8941d9f970689a New theory for converting between predicates and sets. diff -r 0967b03844b5 -r 476604be7d88 src/HOL/Predicate.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Predicate.thy Wed Feb 07 17:25:39 2007 +0100 @@ -0,0 +1,482 @@ +(* Title: HOL/Predicate.thy + ID: $Id$ + Author: Stefan Berghofer, TU Muenchen +*) + +header {* Predicates *} + +theory Predicate +imports Inductive +begin + +subsection {* Converting between predicates and sets *} + +definition + member :: "'a set => 'a => bool" where + "member == %S x. x : S" + +lemma memberI[intro!, Pure.intro!]: "x : S ==> member S x" + by (simp add: member_def) + +lemma memberD[dest!, Pure.dest!]: "member S x ==> x : S" + by (simp add: member_def) + +lemma member_eq[simp]: "member S x = (x : S)" + by (simp add: member_def) + +lemma member_Collect_eq[simp]: "member (Collect P) = P" + by (simp add: member_def) + +lemma Collect_member_eq[simp]: "Collect (member S) = S" + by (simp add: member_def) + +lemma split_set: "(!!S. PROP P S) == (!!S. PROP P (Collect S))" +proof + fix S + assume "!!S. PROP P S" + show "PROP P (Collect S)" . +next + fix S + assume "!!S. PROP P (Collect S)" + have "PROP P {x. x : S}" . + thus "PROP P S" by simp +qed + +lemma split_predicate: "(!!S. PROP P S) == (!!S. PROP P (member S))" +proof + fix S + assume "!!S. PROP P S" + show "PROP P (member S)" . +next + fix S + assume "!!S. PROP P (member S)" + have "PROP P (member {x. S x})" . + thus "PROP P S" by simp +qed + +lemma member_right_eq: "(x == member y) == (Collect x == y)" + by (rule equal_intr_rule, simp, drule symmetric, simp) + +definition + member2 :: "('a * 'b) set => 'a => 'b \ bool" where + "member2 == %S x y. (x, y) : S" + +definition + Collect2 :: "('a => 'b => bool) => ('a * 'b) set" where + "Collect2 == %P. {(x, y). P x y}" + +lemma member2I[intro!, Pure.intro!]: "(x, y) : S ==> member2 S x y" + by (simp add: member2_def) + +lemma member2D[dest!, Pure.dest!]: "member2 S x y ==> (x, y) : S" + by (simp add: member2_def) + +lemma member2_eq[simp]: "member2 S x y = ((x, y) : S)" + by (simp add: member2_def) + +lemma Collect2I: "P x y ==> (x, y) : Collect2 P" + by (simp add: Collect2_def) + +lemma Collect2D: "(x, y) : Collect2 P ==> P x y" + by (simp add: Collect2_def) + +lemma member2_Collect2_eq[simp]: "member2 (Collect2 P) = P" + by (simp add: member2_def Collect2_def) + +lemma Collect2_member2_eq[simp]: "Collect2 (member2 S) = S" + by (auto simp add: member2_def Collect2_def) + +lemma mem_Collect2_eq[iff]: "((x, y) : Collect2 P) = P x y" + by (iprover intro: Collect2I dest: Collect2D) + +lemma member2_Collect_split_eq [simp]: "member2 (Collect (split P)) = P" + by (simp add: expand_fun_eq) + +lemma split_set2: "(!!S. PROP P S) == (!!S. PROP P (Collect2 S))" +proof + fix S + assume "!!S. PROP P S" + show "PROP P (Collect2 S)" . +next + fix S + assume "!!S. PROP P (Collect2 S)" + have "PROP P (Collect2 (member2 S))" . + thus "PROP P S" by simp +qed + +lemma split_predicate2: "(!!S. PROP P S) == (!!S. PROP P (member2 S))" +proof + fix S + assume "!!S. PROP P S" + show "PROP P (member2 S)" . +next + fix S + assume "!!S. PROP P (member2 S)" + have "PROP P (member2 (Collect2 S))" . + thus "PROP P S" by simp +qed + +lemma member2_right_eq: "(x == member2 y) == (Collect2 x == y)" + by (rule equal_intr_rule, simp, drule symmetric, simp) + +ML_setup {* +local + +fun vars_of b (v as Var _) = if b then [] else [v] + | vars_of b (t $ u) = vars_of true t union vars_of false u + | vars_of b (Abs (_, _, t)) = vars_of false t + | vars_of _ _ = []; + +fun rew ths1 ths2 th = Drule.forall_elim_vars 0 + (rewrite_rule ths2 (rewrite_rule ths1 (Drule.forall_intr_list + (map (cterm_of (theory_of_thm th)) (vars_of false (prop_of th))) th))); + +val get_eq = Simpdata.mk_eq o thm; + +val split_predicate = get_eq "split_predicate"; +val split_predicate2 = get_eq "split_predicate2"; +val split_set = get_eq "split_set"; +val split_set2 = get_eq "split_set2"; +val member_eq = get_eq "member_eq"; +val member2_eq = get_eq "member2_eq"; +val member_Collect_eq = get_eq "member_Collect_eq"; +val member2_Collect2_eq = get_eq "member2_Collect2_eq"; +val mem_Collect2_eq = get_eq "mem_Collect2_eq"; +val member_right_eq = thm "member_right_eq"; +val member2_right_eq = thm "member2_right_eq"; + +val rew' = Thm.symmetric o rew [split_set2] [split_set, + member_right_eq, member2_right_eq, member_Collect_eq, member2_Collect2_eq]; + +val rules1 = [split_predicate, split_predicate2, member_eq, member2_eq]; +val rules2 = [split_set, mk_meta_eq mem_Collect_eq, mem_Collect2_eq]; + +structure PredSetConvData = GenericDataFun +(struct + val name = "HOL/pred_set_conv"; + type T = thm list; + val empty = []; + val extend = I; + fun merge _ = Drule.merge_rules; + fun print _ _ = () +end) + +fun mk_attr ths1 ths2 f = Attrib.syntax (Attrib.thms >> (fn ths => + Thm.rule_attribute (fn ctxt => rew ths1 (map (f o Simpdata.mk_eq) + (ths @ PredSetConvData.get ctxt) @ ths2)))); + +val pred_set_conv_att = Attrib.no_args (Thm.declaration_attribute + (Drule.add_rule #> PredSetConvData.map)); + +in + +val _ = ML_Context.>> (PredSetConvData.init #> + Attrib.add_attributes + [("pred_set_conv", pred_set_conv_att, + "declare rules for converting between predicate and set notation"), + ("to_set", mk_attr [] rules1 I, "convert rule to set notation"), + ("to_pred", mk_attr [split_set2] rules2 rew', + "convert rule to predicate notation")]) + +end; +*} + +lemma member_inject [pred_set_conv]: "(member R = member S) = (R = S)" + by (auto simp add: expand_fun_eq) + +lemma member2_inject [pred_set_conv]: "(member2 R = member2 S) = (R = S)" + by (auto simp add: expand_fun_eq) + +lemma member_mono [pred_set_conv]: "(member R <= member S) = (R <= S)" + by fast + +lemma member2_mono [pred_set_conv]: "(member2 R <= member2 S) = (R <= S)" + by fast + +lemma member_empty [pred_set_conv]: "(%x. False) = member {}" + by (simp add: expand_fun_eq) + +lemma member2_empty [pred_set_conv]: "(%x y. False) = member2 {}" + by (simp add: expand_fun_eq) + + +subsubsection {* Binary union *} + +lemma member_Un [pred_set_conv]: "join (member R) (member S) = member (R Un S)" + by (simp add: expand_fun_eq join_fun_eq join_bool_eq) + +lemma member2_Un [pred_set_conv]: "join (member2 R) (member2 S) = member2 (R Un S)" + by (simp add: expand_fun_eq join_fun_eq join_bool_eq) + +lemma join1_iff [simp]: "(join A B x) = (A x | B x)" + by (simp add: join_fun_eq join_bool_eq) + +lemma join2_iff [simp]: "(join A B x y) = (A x y | B x y)" + by (simp add: join_fun_eq join_bool_eq) + +lemma join1I1 [elim?]: "A x ==> join A B x" + by simp + +lemma join2I1 [elim?]: "A x y ==> join A B x y" + by simp + +lemma join1I2 [elim?]: "B x ==> join A B x" + by simp + +lemma join2I2 [elim?]: "B x y ==> join A B x y" + by simp + +text {* + \medskip Classical introduction rule: no commitment to @{text A} vs + @{text B}. +*} + +lemma join1CI [intro!]: "(~ B x ==> A x) ==> join A B x" + by auto + +lemma join2CI [intro!]: "(~ B x y ==> A x y) ==> join A B x y" + by auto + +lemma join1E [elim!]: "join A B x ==> (A x ==> P) ==> (B x ==> P) ==> P" + by simp iprover + +lemma join2E [elim!]: "join A B x y ==> (A x y ==> P) ==> (B x y ==> P) ==> P" + by simp iprover + + +subsubsection {* Binary intersection *} + +lemma member_Int [pred_set_conv]: "meet (member R) (member S) = member (R Int S)" + by (simp add: expand_fun_eq meet_fun_eq meet_bool_eq) + +lemma member2_Int [pred_set_conv]: "meet (member2 R) (member2 S) = member2 (R Int S)" + by (simp add: expand_fun_eq meet_fun_eq meet_bool_eq) + +lemma meet1_iff [simp]: "(meet A B x) = (A x & B x)" + by (simp add: meet_fun_eq meet_bool_eq) + +lemma meet2_iff [simp]: "(meet A B x y) = (A x y & B x y)" + by (simp add: meet_fun_eq meet_bool_eq) + +lemma meet1I [intro!]: "A x ==> B x ==> meet A B x" + by simp + +lemma meet2I [intro!]: "A x y ==> B x y ==> meet A B x y" + by simp + +lemma meet1D1: "meet A B x ==> A x" + by simp + +lemma meet2D1: "meet A B x y ==> A x y" + by simp + +lemma meet1D2: "meet A B x ==> B x" + by simp + +lemma meet2D2: "meet A B x y ==> B x y" + by simp + +lemma meet1E [elim!]: "meet A B x ==> (A x ==> B x ==> P) ==> P" + by simp + +lemma meet2E [elim!]: "meet A B x y ==> (A x y ==> B x y ==> P) ==> P" + by simp + + +subsubsection {* Unions of families *} + +lemma member_SUP: "(SUP i. member (r i)) = member (UN i. r i)" + by (simp add: SUP_def Sup_fun_eq Sup_bool_eq expand_fun_eq) blast + +lemma member2_SUP: "(SUP i. member2 (r i)) = member2 (UN i. r i)" + by (simp add: SUP_def Sup_fun_eq Sup_bool_eq expand_fun_eq) blast + +lemma SUP1_iff [simp]: "((SUP x. B x) b) = (EX x. B x b)" + by (simp add: SUP_def Sup_fun_eq Sup_bool_eq) blast + +lemma SUP2_iff [simp]: "((SUP x. B x) b c) = (EX x. B x b c)" + by (simp add: SUP_def Sup_fun_eq Sup_bool_eq) blast + +lemma SUP1_I [intro]: "B a b ==> (SUP x. B x) b" + by auto + +lemma SUP2_I [intro]: "B a b c ==> (SUP x. B x) b c" + by auto + +lemma SUP1_E [elim!]: "(SUP x. B x) b ==> (!!x. B x b ==> R) ==> R" + by simp iprover + +lemma SUP2_E [elim!]: "(SUP x. B x) b c ==> (!!x. B x b c ==> R) ==> R" + by simp iprover + + +subsection {* Composition of two relations *} + +inductive2 + pred_comp :: "['b => 'c => bool, 'a => 'b => bool] => 'a => 'c => bool" + (infixr "OO" 75) + for r :: "'b => 'c => bool" and s :: "'a => 'b => bool" +where + pred_compI [intro]: "s a b ==> r b c ==> (r OO s) a c" + +inductive_cases2 pred_compE [elim!]: "(r OO s) a c" + +lemma pred_comp_rel_comp_eq [pred_set_conv]: + "(member2 r OO member2 s) = member2 (r O s)" + by (auto simp add: expand_fun_eq elim: pred_compE) + + +subsection {* Converse *} + +inductive2 + conversep :: "('a => 'b => bool) => 'b => 'a => bool" + ("(_^--1)" [1000] 1000) + for r :: "'a => 'b => bool" +where + conversepI: "r a b ==> r^--1 b a" + +notation (xsymbols) + conversep ("(_\\)" [1000] 1000) + +lemma conversepD: + assumes ab: "r^--1 a b" + shows "r b a" using ab + by cases simp + +lemma conversep_iff [iff]: "r^--1 a b = r b a" + by (iprover intro: conversepI dest: conversepD) + +lemma conversep_converse_eq [pred_set_conv]: + "(member2 r)^--1 = member2 (r^-1)" + by (auto simp add: expand_fun_eq) + +lemma conversep_conversep [simp]: "(r^--1)^--1 = r" + by (iprover intro: order_antisym conversepI dest: conversepD) + +lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1" + by (iprover intro: order_antisym conversepI pred_compI + elim: pred_compE dest: conversepD) + +lemma converse_meet: "(meet r s)^--1 = meet r^--1 s^--1" + by (simp add: meet_fun_eq meet_bool_eq) + (iprover intro: conversepI ext dest: conversepD) + +lemma converse_join: "(join r s)^--1 = join r^--1 s^--1" + by (simp add: join_fun_eq join_bool_eq) + (iprover intro: conversepI ext dest: conversepD) + +lemma conversep_noteq [simp]: "(op ~=)^--1 = op ~=" + by (auto simp add: expand_fun_eq) + +lemma conversep_eq [simp]: "(op =)^--1 = op =" + by (auto simp add: expand_fun_eq) + + +subsection {* Domain *} + +inductive2 + DomainP :: "('a => 'b => bool) => 'a => bool" + for r :: "'a => 'b => bool" +where + DomainPI [intro]: "r a b ==> DomainP r a" + +inductive_cases2 DomainPE [elim!]: "DomainP r a" + +lemma member2_DomainP [pred_set_conv]: "DomainP (member2 r) = member (Domain r)" + by (blast intro!: Orderings.order_antisym) + + +subsection {* Range *} + +inductive2 + RangeP :: "('a => 'b => bool) => 'b => bool" + for r :: "'a => 'b => bool" +where + RangePI [intro]: "r a b ==> RangeP r b" + +inductive_cases2 RangePE [elim!]: "RangeP r b" + +lemma member2_RangeP [pred_set_conv]: "RangeP (member2 r) = member (Range r)" + by (blast intro!: Orderings.order_antisym) + + +subsection {* Inverse image *} + +definition + inv_imagep :: "('b => 'b => bool) => ('a => 'b) => 'a => 'a => bool" where + "inv_imagep r f == %x y. r (f x) (f y)" + +lemma [pred_set_conv]: "inv_imagep (member2 r) f = member2 (inv_image r f)" + by (simp add: inv_image_def inv_imagep_def) + +lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)" + by (simp add: inv_imagep_def) + + +subsection {* Properties of relations - predicate versions *} + +abbreviation antisymP :: "('a => 'a => bool) => bool" where + "antisymP r == antisym (Collect2 r)" + +abbreviation transP :: "('a => 'a => bool) => bool" where + "transP r == trans (Collect2 r)" + +abbreviation single_valuedP :: "('a => 'b => bool) => bool" where + "single_valuedP r == single_valued (Collect2 r)" + + +subsection {* Bounded quantifiers for predicates *} + +text {* +Bounded existential quantifier for predicates (executable). +*} + +inductive2 bexp :: "('a \ bool) \ ('a \ bool) \ bool" + for P :: "'a \ bool" and Q :: "'a \ bool" +where + bexpI [intro]: "P x \ Q x \ bexp P Q" + +lemmas bexpE [elim!] = bexp.cases + +syntax + Bexp :: "'a \ ('a \ bool) \ bool \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) + +translations + "\x\P. Q" \ "CONST bexp P (\x. Q)" + +constdefs + ballp :: "('a \ bool) \ ('a \ bool) \ bool" + "ballp P Q \ \x. P x \ Q x" + +syntax + Ballp :: "'a \ ('a \ bool) \ bool \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) + +translations + "\x\P. Q" \ "CONST ballp P (\x. Q)" + +(* To avoid eta-contraction of body: *) +print_translation {* +let + fun btr' syn [A,Abs abs] = + let val (x,t) = atomic_abs_tr' abs + in Syntax.const syn $ x $ A $ t end +in +[("ballp", btr' "Ballp"),("bexp", btr' "Bexp")] +end +*} + +lemma ballpI [intro!]: "(\x. A x \ P x) \ \x\A. P x" + by (simp add: ballp_def) + +lemma bspecp [dest?]: "\x\A. P x \ A x \ P x" + by (simp add: ballp_def) + +lemma ballpE [elim]: "\x\A. P x \ (P x \ Q) \ (\ A x \ Q) \ Q" + by (unfold ballp_def) blast + +lemma ballp_not_bexp_eq: "(\x\P. Q x) = (\ (\x\P. \ Q x))" + by (blast dest: bspecp) + +declare ballp_not_bexp_eq [THEN eq_reflection, code unfold] + +end