berghofe@22259: (* Title: HOL/Predicate.thy berghofe@22259: ID: $Id$ berghofe@22259: Author: Stefan Berghofer, TU Muenchen berghofe@22259: *) berghofe@22259: berghofe@22259: header {* Predicates *} berghofe@22259: berghofe@22259: theory Predicate haftmann@23708: imports Inductive Relation berghofe@22259: begin berghofe@22259: berghofe@22259: subsection {* Converting between predicates and sets *} berghofe@22259: berghofe@22259: definition berghofe@22259: member :: "'a set => 'a => bool" where berghofe@22259: "member == %S x. x : S" berghofe@22259: berghofe@22259: lemma memberI[intro!, Pure.intro!]: "x : S ==> member S x" berghofe@22259: by (simp add: member_def) berghofe@22259: berghofe@22259: lemma memberD[dest!, Pure.dest!]: "member S x ==> x : S" berghofe@22259: by (simp add: member_def) berghofe@22259: berghofe@22259: lemma member_eq[simp]: "member S x = (x : S)" berghofe@22259: by (simp add: member_def) berghofe@22259: berghofe@22259: lemma member_Collect_eq[simp]: "member (Collect P) = P" berghofe@22259: by (simp add: member_def) berghofe@22259: berghofe@22259: lemma Collect_member_eq[simp]: "Collect (member S) = S" berghofe@22259: by (simp add: member_def) berghofe@22259: berghofe@22259: lemma split_set: "(!!S. PROP P S) == (!!S. PROP P (Collect S))" berghofe@22259: proof berghofe@22259: fix S berghofe@22259: assume "!!S. PROP P S" wenzelm@23389: then show "PROP P (Collect S)" . berghofe@22259: next berghofe@22259: fix S berghofe@22259: assume "!!S. PROP P (Collect S)" wenzelm@23389: then have "PROP P {x. x : S}" . berghofe@22259: thus "PROP P S" by simp berghofe@22259: qed berghofe@22259: berghofe@22259: lemma split_predicate: "(!!S. PROP P S) == (!!S. PROP P (member S))" berghofe@22259: proof berghofe@22259: fix S berghofe@22259: assume "!!S. PROP P S" wenzelm@23389: then show "PROP P (member S)" . berghofe@22259: next berghofe@22259: fix S berghofe@22259: assume "!!S. PROP P (member S)" wenzelm@23389: then have "PROP P (member {x. S x})" . berghofe@22259: thus "PROP P S" by simp berghofe@22259: qed berghofe@22259: berghofe@22259: lemma member_right_eq: "(x == member y) == (Collect x == y)" berghofe@22259: by (rule equal_intr_rule, simp, drule symmetric, simp) berghofe@22259: berghofe@22259: definition berghofe@22259: member2 :: "('a * 'b) set => 'a => 'b \ bool" where berghofe@22259: "member2 == %S x y. (x, y) : S" berghofe@22259: berghofe@22259: definition berghofe@22259: Collect2 :: "('a => 'b => bool) => ('a * 'b) set" where berghofe@22259: "Collect2 == %P. {(x, y). P x y}" berghofe@22259: berghofe@22259: lemma member2I[intro!, Pure.intro!]: "(x, y) : S ==> member2 S x y" berghofe@22259: by (simp add: member2_def) berghofe@22259: berghofe@22259: lemma member2D[dest!, Pure.dest!]: "member2 S x y ==> (x, y) : S" berghofe@22259: by (simp add: member2_def) berghofe@22259: berghofe@22259: lemma member2_eq[simp]: "member2 S x y = ((x, y) : S)" berghofe@22259: by (simp add: member2_def) berghofe@22259: berghofe@22259: lemma Collect2I: "P x y ==> (x, y) : Collect2 P" berghofe@22259: by (simp add: Collect2_def) berghofe@22259: berghofe@22259: lemma Collect2D: "(x, y) : Collect2 P ==> P x y" berghofe@22259: by (simp add: Collect2_def) berghofe@22259: berghofe@22259: lemma member2_Collect2_eq[simp]: "member2 (Collect2 P) = P" berghofe@22259: by (simp add: member2_def Collect2_def) berghofe@22259: berghofe@22259: lemma Collect2_member2_eq[simp]: "Collect2 (member2 S) = S" berghofe@22259: by (auto simp add: member2_def Collect2_def) berghofe@22259: berghofe@22259: lemma mem_Collect2_eq[iff]: "((x, y) : Collect2 P) = P x y" berghofe@22259: by (iprover intro: Collect2I dest: Collect2D) berghofe@22259: berghofe@22259: lemma member2_Collect_split_eq [simp]: "member2 (Collect (split P)) = P" berghofe@22259: by (simp add: expand_fun_eq) berghofe@22259: berghofe@22259: lemma split_set2: "(!!S. PROP P S) == (!!S. PROP P (Collect2 S))" berghofe@22259: proof berghofe@22259: fix S berghofe@22259: assume "!!S. PROP P S" wenzelm@23389: then show "PROP P (Collect2 S)" . berghofe@22259: next berghofe@22259: fix S berghofe@22259: assume "!!S. PROP P (Collect2 S)" wenzelm@23389: then have "PROP P (Collect2 (member2 S))" . berghofe@22259: thus "PROP P S" by simp berghofe@22259: qed berghofe@22259: berghofe@22259: lemma split_predicate2: "(!!S. PROP P S) == (!!S. PROP P (member2 S))" berghofe@22259: proof berghofe@22259: fix S berghofe@22259: assume "!!S. PROP P S" wenzelm@23389: then show "PROP P (member2 S)" . berghofe@22259: next berghofe@22259: fix S berghofe@22259: assume "!!S. PROP P (member2 S)" wenzelm@23389: then have "PROP P (member2 (Collect2 S))" . berghofe@22259: thus "PROP P S" by simp berghofe@22259: qed berghofe@22259: berghofe@22259: lemma member2_right_eq: "(x == member2 y) == (Collect2 x == y)" berghofe@22259: by (rule equal_intr_rule, simp, drule symmetric, simp) berghofe@22259: berghofe@22259: ML_setup {* berghofe@22259: local berghofe@22259: berghofe@22259: fun vars_of b (v as Var _) = if b then [] else [v] berghofe@22259: | vars_of b (t $ u) = vars_of true t union vars_of false u berghofe@22259: | vars_of b (Abs (_, _, t)) = vars_of false t berghofe@22259: | vars_of _ _ = []; berghofe@22259: berghofe@22259: fun rew ths1 ths2 th = Drule.forall_elim_vars 0 berghofe@22259: (rewrite_rule ths2 (rewrite_rule ths1 (Drule.forall_intr_list berghofe@22259: (map (cterm_of (theory_of_thm th)) (vars_of false (prop_of th))) th))); berghofe@22259: berghofe@22259: val get_eq = Simpdata.mk_eq o thm; berghofe@22259: berghofe@22259: val split_predicate = get_eq "split_predicate"; berghofe@22259: val split_predicate2 = get_eq "split_predicate2"; berghofe@22259: val split_set = get_eq "split_set"; berghofe@22259: val split_set2 = get_eq "split_set2"; berghofe@22259: val member_eq = get_eq "member_eq"; berghofe@22259: val member2_eq = get_eq "member2_eq"; berghofe@22259: val member_Collect_eq = get_eq "member_Collect_eq"; berghofe@22259: val member2_Collect2_eq = get_eq "member2_Collect2_eq"; berghofe@22259: val mem_Collect2_eq = get_eq "mem_Collect2_eq"; berghofe@22259: val member_right_eq = thm "member_right_eq"; berghofe@22259: val member2_right_eq = thm "member2_right_eq"; berghofe@22259: berghofe@22259: val rew' = Thm.symmetric o rew [split_set2] [split_set, berghofe@22259: member_right_eq, member2_right_eq, member_Collect_eq, member2_Collect2_eq]; berghofe@22259: berghofe@22259: val rules1 = [split_predicate, split_predicate2, member_eq, member2_eq]; berghofe@22259: val rules2 = [split_set, mk_meta_eq mem_Collect_eq, mem_Collect2_eq]; berghofe@22259: berghofe@22259: structure PredSetConvData = GenericDataFun wenzelm@22846: ( berghofe@22259: type T = thm list; berghofe@22259: val empty = []; berghofe@22259: val extend = I; berghofe@22259: fun merge _ = Drule.merge_rules; wenzelm@22846: ); berghofe@22259: berghofe@22259: fun mk_attr ths1 ths2 f = Attrib.syntax (Attrib.thms >> (fn ths => berghofe@22259: Thm.rule_attribute (fn ctxt => rew ths1 (map (f o Simpdata.mk_eq) berghofe@22259: (ths @ PredSetConvData.get ctxt) @ ths2)))); berghofe@22259: berghofe@22259: val pred_set_conv_att = Attrib.no_args (Thm.declaration_attribute berghofe@22259: (Drule.add_rule #> PredSetConvData.map)); berghofe@22259: berghofe@22259: in berghofe@22259: wenzelm@22846: val _ = ML_Context.>> ( berghofe@22259: Attrib.add_attributes berghofe@22259: [("pred_set_conv", pred_set_conv_att, berghofe@22259: "declare rules for converting between predicate and set notation"), berghofe@22259: ("to_set", mk_attr [] rules1 I, "convert rule to set notation"), berghofe@22259: ("to_pred", mk_attr [split_set2] rules2 rew', berghofe@22259: "convert rule to predicate notation")]) berghofe@22259: berghofe@22259: end; berghofe@22259: *} berghofe@22259: berghofe@22259: lemma member_inject [pred_set_conv]: "(member R = member S) = (R = S)" berghofe@22259: by (auto simp add: expand_fun_eq) berghofe@22259: berghofe@22259: lemma member2_inject [pred_set_conv]: "(member2 R = member2 S) = (R = S)" berghofe@22259: by (auto simp add: expand_fun_eq) berghofe@22259: berghofe@22259: lemma member_mono [pred_set_conv]: "(member R <= member S) = (R <= S)" berghofe@22259: by fast berghofe@22259: berghofe@22259: lemma member2_mono [pred_set_conv]: "(member2 R <= member2 S) = (R <= S)" berghofe@22259: by fast berghofe@22259: berghofe@22259: lemma member_empty [pred_set_conv]: "(%x. False) = member {}" berghofe@22259: by (simp add: expand_fun_eq) berghofe@22259: berghofe@22259: lemma member2_empty [pred_set_conv]: "(%x y. False) = member2 {}" berghofe@22259: by (simp add: expand_fun_eq) berghofe@22259: berghofe@22259: berghofe@22259: subsubsection {* Binary union *} berghofe@22259: haftmann@22422: lemma member_Un [pred_set_conv]: "sup (member R) (member S) = member (R Un S)" haftmann@22422: by (simp add: expand_fun_eq sup_fun_eq sup_bool_eq) berghofe@22259: haftmann@22422: lemma member2_Un [pred_set_conv]: "sup (member2 R) (member2 S) = member2 (R Un S)" haftmann@22422: by (simp add: expand_fun_eq sup_fun_eq sup_bool_eq) berghofe@22259: haftmann@22422: lemma sup1_iff [simp]: "sup A B x \ A x | B x" haftmann@22422: by (simp add: sup_fun_eq sup_bool_eq) berghofe@22259: haftmann@22422: lemma sup2_iff [simp]: "sup A B x y \ A x y | B x y" haftmann@22422: by (simp add: sup_fun_eq sup_bool_eq) berghofe@22259: haftmann@22422: lemma sup1I1 [elim?]: "A x \ sup A B x" berghofe@22259: by simp berghofe@22259: haftmann@22422: lemma sup2I1 [elim?]: "A x y \ sup A B x y" berghofe@22259: by simp berghofe@22259: haftmann@22422: lemma join1I2 [elim?]: "B x \ sup A B x" berghofe@22259: by simp berghofe@22259: haftmann@22422: lemma sup1I2 [elim?]: "B x y \ sup A B x y" berghofe@22259: by simp berghofe@22259: berghofe@22259: text {* berghofe@22259: \medskip Classical introduction rule: no commitment to @{text A} vs berghofe@22259: @{text B}. berghofe@22259: *} berghofe@22259: haftmann@22422: lemma sup1CI [intro!]: "(~ B x ==> A x) ==> sup A B x" berghofe@22259: by auto berghofe@22259: haftmann@22422: lemma sup2CI [intro!]: "(~ B x y ==> A x y) ==> sup A B x y" berghofe@22259: by auto berghofe@22259: haftmann@22422: lemma sup1E [elim!]: "sup A B x ==> (A x ==> P) ==> (B x ==> P) ==> P" berghofe@22259: by simp iprover berghofe@22259: haftmann@22422: lemma sup2E [elim!]: "sup A B x y ==> (A x y ==> P) ==> (B x y ==> P) ==> P" berghofe@22259: by simp iprover berghofe@22259: berghofe@22259: berghofe@22259: subsubsection {* Binary intersection *} berghofe@22259: haftmann@22422: lemma member_Int [pred_set_conv]: "inf (member R) (member S) = member (R Int S)" haftmann@22422: by (simp add: expand_fun_eq inf_fun_eq inf_bool_eq) berghofe@22259: haftmann@22422: lemma member2_Int [pred_set_conv]: "inf (member2 R) (member2 S) = member2 (R Int S)" haftmann@22422: by (simp add: expand_fun_eq inf_fun_eq inf_bool_eq) berghofe@22259: haftmann@22422: lemma inf1_iff [simp]: "inf A B x \ A x \ B x" haftmann@22422: by (simp add: inf_fun_eq inf_bool_eq) berghofe@22259: haftmann@22422: lemma inf2_iff [simp]: "inf A B x y \ A x y \ B x y" haftmann@22422: by (simp add: inf_fun_eq inf_bool_eq) berghofe@22259: haftmann@22422: lemma inf1I [intro!]: "A x ==> B x ==> inf A B x" berghofe@22259: by simp berghofe@22259: haftmann@22422: lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y" berghofe@22259: by simp berghofe@22259: haftmann@22422: lemma inf1D1: "inf A B x ==> A x" berghofe@22259: by simp berghofe@22259: haftmann@22422: lemma inf2D1: "inf A B x y ==> A x y" berghofe@22259: by simp berghofe@22259: haftmann@22422: lemma inf1D2: "inf A B x ==> B x" berghofe@22259: by simp berghofe@22259: haftmann@22422: lemma inf2D2: "inf A B x y ==> B x y" berghofe@22259: by simp berghofe@22259: haftmann@22422: lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P" berghofe@22259: by simp berghofe@22259: haftmann@22422: lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P" berghofe@22259: by simp berghofe@22259: berghofe@22259: berghofe@22259: subsubsection {* Unions of families *} berghofe@22259: berghofe@22259: lemma member_SUP: "(SUP i. member (r i)) = member (UN i. r i)" berghofe@22430: by (simp add: SUPR_def Sup_fun_eq Sup_bool_eq expand_fun_eq) blast berghofe@22259: berghofe@22259: lemma member2_SUP: "(SUP i. member2 (r i)) = member2 (UN i. r i)" berghofe@22430: by (simp add: SUPR_def Sup_fun_eq Sup_bool_eq expand_fun_eq) blast berghofe@22259: berghofe@22430: lemma SUP1_iff [simp]: "(SUP x:A. B x) b = (EX x:A. B x b)" berghofe@22430: by (simp add: SUPR_def Sup_fun_eq Sup_bool_eq) blast berghofe@22430: berghofe@22430: lemma SUP2_iff [simp]: "(SUP x:A. B x) b c = (EX x:A. B x b c)" berghofe@22430: by (simp add: SUPR_def Sup_fun_eq Sup_bool_eq) blast berghofe@22259: berghofe@22430: lemma SUP1_I [intro]: "a : A ==> B a b ==> (SUP x:A. B x) b" berghofe@22430: by auto berghofe@22259: berghofe@22430: lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c" berghofe@22430: by auto berghofe@22430: berghofe@22430: lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R" berghofe@22430: by auto berghofe@22430: berghofe@22430: lemma SUP2_E [elim!]: "(SUP x:A. B x) b c ==> (!!x. x : A ==> B x b c ==> R) ==> R" berghofe@22259: by auto berghofe@22259: berghofe@22430: berghofe@22430: subsubsection {* Intersections of families *} berghofe@22430: berghofe@22430: lemma member_INF: "(INF i. member (r i)) = member (INT i. r i)" berghofe@22430: by (simp add: INFI_def Inf_fun_def Inf_bool_def expand_fun_eq) blast berghofe@22430: berghofe@22430: lemma member2_INF: "(INF i. member2 (r i)) = member2 (INT i. r i)" berghofe@22430: by (simp add: INFI_def Inf_fun_def Inf_bool_def expand_fun_eq) blast berghofe@22430: berghofe@22430: lemma INF1_iff [simp]: "(INF x:A. B x) b = (ALL x:A. B x b)" berghofe@22430: by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast berghofe@22430: berghofe@22430: lemma INF2_iff [simp]: "(INF x:A. B x) b c = (ALL x:A. B x b c)" berghofe@22430: by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast berghofe@22430: berghofe@22430: lemma INF1_I [intro!]: "(!!x. x : A ==> B x b) ==> (INF x:A. B x) b" berghofe@22259: by auto berghofe@22259: berghofe@22430: lemma INF2_I [intro!]: "(!!x. x : A ==> B x b c) ==> (INF x:A. B x) b c" berghofe@22430: by auto berghofe@22430: berghofe@22430: lemma INF1_D [elim]: "(INF x:A. B x) b ==> a : A ==> B a b" berghofe@22430: by auto berghofe@22259: berghofe@22430: lemma INF2_D [elim]: "(INF x:A. B x) b c ==> a : A ==> B a b c" berghofe@22430: by auto berghofe@22430: berghofe@22430: lemma INF1_E [elim]: "(INF x:A. B x) b ==> (B a b ==> R) ==> (a ~: A ==> R) ==> R" berghofe@22430: by auto berghofe@22430: berghofe@22430: lemma INF2_E [elim]: "(INF x:A. B x) b c ==> (B a b c ==> R) ==> (a ~: A ==> R) ==> R" berghofe@22430: by auto berghofe@22259: berghofe@22259: berghofe@22259: subsection {* Composition of two relations *} berghofe@22259: berghofe@22259: inductive2 berghofe@22259: pred_comp :: "['b => 'c => bool, 'a => 'b => bool] => 'a => 'c => bool" berghofe@22259: (infixr "OO" 75) berghofe@22259: for r :: "'b => 'c => bool" and s :: "'a => 'b => bool" berghofe@22259: where berghofe@22259: pred_compI [intro]: "s a b ==> r b c ==> (r OO s) a c" berghofe@22259: berghofe@22259: inductive_cases2 pred_compE [elim!]: "(r OO s) a c" berghofe@22259: berghofe@22259: lemma pred_comp_rel_comp_eq [pred_set_conv]: berghofe@22259: "(member2 r OO member2 s) = member2 (r O s)" berghofe@22259: by (auto simp add: expand_fun_eq elim: pred_compE) berghofe@22259: berghofe@22259: berghofe@22259: subsection {* Converse *} berghofe@22259: berghofe@22259: inductive2 berghofe@22259: conversep :: "('a => 'b => bool) => 'b => 'a => bool" berghofe@22259: ("(_^--1)" [1000] 1000) berghofe@22259: for r :: "'a => 'b => bool" berghofe@22259: where berghofe@22259: conversepI: "r a b ==> r^--1 b a" berghofe@22259: berghofe@22259: notation (xsymbols) berghofe@22259: conversep ("(_\\)" [1000] 1000) berghofe@22259: berghofe@22259: lemma conversepD: berghofe@22259: assumes ab: "r^--1 a b" berghofe@22259: shows "r b a" using ab berghofe@22259: by cases simp berghofe@22259: berghofe@22259: lemma conversep_iff [iff]: "r^--1 a b = r b a" berghofe@22259: by (iprover intro: conversepI dest: conversepD) berghofe@22259: berghofe@22259: lemma conversep_converse_eq [pred_set_conv]: berghofe@22259: "(member2 r)^--1 = member2 (r^-1)" berghofe@22259: by (auto simp add: expand_fun_eq) berghofe@22259: berghofe@22259: lemma conversep_conversep [simp]: "(r^--1)^--1 = r" berghofe@22259: by (iprover intro: order_antisym conversepI dest: conversepD) berghofe@22259: berghofe@22259: lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1" berghofe@22259: by (iprover intro: order_antisym conversepI pred_compI berghofe@22259: elim: pred_compE dest: conversepD) berghofe@22259: haftmann@22422: lemma converse_meet: "(inf r s)^--1 = inf r^--1 s^--1" haftmann@22422: by (simp add: inf_fun_eq inf_bool_eq) berghofe@22259: (iprover intro: conversepI ext dest: conversepD) berghofe@22259: haftmann@22422: lemma converse_join: "(sup r s)^--1 = sup r^--1 s^--1" haftmann@22422: by (simp add: sup_fun_eq sup_bool_eq) berghofe@22259: (iprover intro: conversepI ext dest: conversepD) berghofe@22259: berghofe@22259: lemma conversep_noteq [simp]: "(op ~=)^--1 = op ~=" berghofe@22259: by (auto simp add: expand_fun_eq) berghofe@22259: berghofe@22259: lemma conversep_eq [simp]: "(op =)^--1 = op =" berghofe@22259: by (auto simp add: expand_fun_eq) berghofe@22259: berghofe@22259: berghofe@22259: subsection {* Domain *} berghofe@22259: berghofe@22259: inductive2 berghofe@22259: DomainP :: "('a => 'b => bool) => 'a => bool" berghofe@22259: for r :: "'a => 'b => bool" berghofe@22259: where berghofe@22259: DomainPI [intro]: "r a b ==> DomainP r a" berghofe@22259: berghofe@22259: inductive_cases2 DomainPE [elim!]: "DomainP r a" berghofe@22259: berghofe@22259: lemma member2_DomainP [pred_set_conv]: "DomainP (member2 r) = member (Domain r)" berghofe@22259: by (blast intro!: Orderings.order_antisym) berghofe@22259: berghofe@22259: berghofe@22259: subsection {* Range *} berghofe@22259: berghofe@22259: inductive2 berghofe@22259: RangeP :: "('a => 'b => bool) => 'b => bool" berghofe@22259: for r :: "'a => 'b => bool" berghofe@22259: where berghofe@22259: RangePI [intro]: "r a b ==> RangeP r b" berghofe@22259: berghofe@22259: inductive_cases2 RangePE [elim!]: "RangeP r b" berghofe@22259: berghofe@22259: lemma member2_RangeP [pred_set_conv]: "RangeP (member2 r) = member (Range r)" berghofe@22259: by (blast intro!: Orderings.order_antisym) berghofe@22259: berghofe@22259: berghofe@22259: subsection {* Inverse image *} berghofe@22259: berghofe@22259: definition berghofe@22259: inv_imagep :: "('b => 'b => bool) => ('a => 'b) => 'a => 'a => bool" where berghofe@22259: "inv_imagep r f == %x y. r (f x) (f y)" berghofe@22259: berghofe@22259: lemma [pred_set_conv]: "inv_imagep (member2 r) f = member2 (inv_image r f)" berghofe@22259: by (simp add: inv_image_def inv_imagep_def) berghofe@22259: berghofe@22259: lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)" berghofe@22259: by (simp add: inv_imagep_def) berghofe@22259: berghofe@22259: berghofe@22259: subsection {* Properties of relations - predicate versions *} berghofe@22259: berghofe@22259: abbreviation antisymP :: "('a => 'a => bool) => bool" where berghofe@22259: "antisymP r == antisym (Collect2 r)" berghofe@22259: berghofe@22259: abbreviation transP :: "('a => 'a => bool) => bool" where berghofe@22259: "transP r == trans (Collect2 r)" berghofe@22259: berghofe@22259: abbreviation single_valuedP :: "('a => 'b => bool) => bool" where berghofe@22259: "single_valuedP r == single_valued (Collect2 r)" berghofe@22259: berghofe@22259: berghofe@22259: subsection {* Bounded quantifiers for predicates *} berghofe@22259: berghofe@22259: text {* berghofe@22259: Bounded existential quantifier for predicates (executable). berghofe@22259: *} berghofe@22259: berghofe@22259: inductive2 bexp :: "('a \ bool) \ ('a \ bool) \ bool" berghofe@22259: for P :: "'a \ bool" and Q :: "'a \ bool" berghofe@22259: where berghofe@22259: bexpI [intro]: "P x \ Q x \ bexp P Q" berghofe@22259: berghofe@22259: lemmas bexpE [elim!] = bexp.cases berghofe@22259: berghofe@22259: syntax berghofe@22259: Bexp :: "'a \ ('a \ bool) \ bool \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) berghofe@22259: berghofe@22259: translations berghofe@22259: "\x\P. Q" \ "CONST bexp P (\x. Q)" berghofe@22259: berghofe@22259: constdefs berghofe@22259: ballp :: "('a \ bool) \ ('a \ bool) \ bool" berghofe@22259: "ballp P Q \ \x. P x \ Q x" berghofe@22259: berghofe@22259: syntax berghofe@22259: Ballp :: "'a \ ('a \ bool) \ bool \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) berghofe@22259: berghofe@22259: translations berghofe@22259: "\x\P. Q" \ "CONST ballp P (\x. Q)" berghofe@22259: berghofe@22259: (* To avoid eta-contraction of body: *) berghofe@22259: print_translation {* berghofe@22259: let berghofe@22259: fun btr' syn [A,Abs abs] = berghofe@22259: let val (x,t) = atomic_abs_tr' abs berghofe@22259: in Syntax.const syn $ x $ A $ t end berghofe@22259: in berghofe@22259: [("ballp", btr' "Ballp"),("bexp", btr' "Bexp")] berghofe@22259: end berghofe@22259: *} berghofe@22259: berghofe@22259: lemma ballpI [intro!]: "(\x. A x \ P x) \ \x\A. P x" berghofe@22259: by (simp add: ballp_def) berghofe@22259: berghofe@22259: lemma bspecp [dest?]: "\x\A. P x \ A x \ P x" berghofe@22259: by (simp add: ballp_def) berghofe@22259: berghofe@22259: lemma ballpE [elim]: "\x\A. P x \ (P x \ Q) \ (\ A x \ Q) \ Q" berghofe@22259: by (unfold ballp_def) blast berghofe@22259: berghofe@22259: lemma ballp_not_bexp_eq: "(\x\P. Q x) = (\ (\x\P. \ Q x))" berghofe@22259: by (blast dest: bspecp) berghofe@22259: berghofe@22259: declare ballp_not_bexp_eq [THEN eq_reflection, code unfold] berghofe@22259: berghofe@22259: end