berghofe@22259: (* Title: HOL/Predicate.thy haftmann@30328: Author: Stefan Berghofer and Lukas Bulwahn and Florian Haftmann, TU Muenchen berghofe@22259: *) berghofe@22259: haftmann@30328: header {* Predicates as relations and enumerations *} berghofe@22259: berghofe@22259: theory Predicate haftmann@23708: imports Inductive Relation berghofe@22259: begin berghofe@22259: haftmann@30328: notation haftmann@30328: inf (infixl "\" 70) and haftmann@30328: sup (infixl "\" 65) and haftmann@30328: Inf ("\_" [900] 900) and haftmann@30328: Sup ("\_" [900] 900) and haftmann@30328: top ("\") and haftmann@30328: bot ("\") haftmann@30328: haftmann@30328: haftmann@30328: subsection {* Predicates as (complete) lattices *} haftmann@30328: haftmann@30328: subsubsection {* @{const sup} on @{typ bool} *} haftmann@30328: haftmann@30328: lemma sup_boolI1: haftmann@30328: "P \ P \ Q" haftmann@30328: by (simp add: sup_bool_eq) haftmann@30328: haftmann@30328: lemma sup_boolI2: haftmann@30328: "Q \ P \ Q" haftmann@30328: by (simp add: sup_bool_eq) haftmann@30328: haftmann@30328: lemma sup_boolE: haftmann@30328: "P \ Q \ (P \ R) \ (Q \ R) \ R" haftmann@30328: by (auto simp add: sup_bool_eq) haftmann@30328: haftmann@30328: haftmann@30328: subsubsection {* Equality and Subsets *} berghofe@22259: berghofe@26797: lemma pred_equals_eq: "((\x. x \ R) = (\x. x \ S)) = (R = S)" berghofe@26797: by (simp add: mem_def) berghofe@22259: berghofe@23741: lemma pred_equals_eq2 [pred_set_conv]: "((\x y. (x, y) \ R) = (\x y. (x, y) \ S)) = (R = S)" berghofe@26797: by (simp add: expand_fun_eq mem_def) berghofe@22259: berghofe@26797: lemma pred_subset_eq: "((\x. x \ R) <= (\x. x \ S)) = (R <= S)" berghofe@26797: by (simp add: mem_def) berghofe@22259: berghofe@23741: lemma pred_subset_eq2 [pred_set_conv]: "((\x y. (x, y) \ R) <= (\x y. (x, y) \ S)) = (R <= S)" berghofe@22259: by fast berghofe@22259: berghofe@23741: haftmann@30328: subsubsection {* Top and bottom elements *} berghofe@23741: berghofe@23741: lemma top1I [intro!]: "top x" berghofe@23741: by (simp add: top_fun_eq top_bool_eq) berghofe@22259: berghofe@23741: lemma top2I [intro!]: "top x y" berghofe@23741: by (simp add: top_fun_eq top_bool_eq) berghofe@23741: berghofe@23741: lemma bot1E [elim!]: "bot x \ P" berghofe@23741: by (simp add: bot_fun_eq bot_bool_eq) berghofe@23741: berghofe@23741: lemma bot2E [elim!]: "bot x y \ P" berghofe@23741: by (simp add: bot_fun_eq bot_bool_eq) berghofe@22259: berghofe@22259: haftmann@30328: subsubsection {* The empty set *} berghofe@23741: berghofe@23741: lemma bot_empty_eq: "bot = (\x. x \ {})" berghofe@23741: by (auto simp add: expand_fun_eq) berghofe@22259: berghofe@23741: lemma bot_empty_eq2: "bot = (\x y. (x, y) \ {})" berghofe@23741: by (auto simp add: expand_fun_eq) berghofe@22259: berghofe@23741: haftmann@30328: subsubsection {* Binary union *} berghofe@22259: haftmann@32601: lemma sup1_iff: "sup A B x \ A x | B x" haftmann@22422: by (simp add: sup_fun_eq sup_bool_eq) berghofe@22259: haftmann@32601: lemma sup2_iff: "sup A B x y \ A x y | B x y" haftmann@22422: by (simp add: sup_fun_eq sup_bool_eq) berghofe@22259: berghofe@23741: lemma sup_Un_eq [pred_set_conv]: "sup (\x. x \ R) (\x. x \ S) = (\x. x \ R \ S)" haftmann@32601: by (simp add: sup1_iff expand_fun_eq) berghofe@23741: berghofe@23741: lemma sup_Un_eq2 [pred_set_conv]: "sup (\x y. (x, y) \ R) (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" haftmann@32601: by (simp add: sup2_iff expand_fun_eq) berghofe@23741: haftmann@22422: lemma sup1I1 [elim?]: "A x \ sup A B x" haftmann@32601: by (simp add: sup1_iff) berghofe@22259: haftmann@22422: lemma sup2I1 [elim?]: "A x y \ sup A B x y" haftmann@32601: by (simp add: sup2_iff) berghofe@22259: berghofe@23741: lemma sup1I2 [elim?]: "B x \ sup A B x" haftmann@32601: by (simp add: sup1_iff) berghofe@22259: berghofe@23741: lemma sup2I2 [elim?]: "B x y \ sup A B x y" haftmann@32601: by (simp add: sup2_iff) 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" haftmann@32601: by (auto simp add: sup1_iff) berghofe@22259: haftmann@22422: lemma sup2CI [intro!]: "(~ B x y ==> A x y) ==> sup A B x y" haftmann@32601: by (auto simp add: sup2_iff) berghofe@22259: haftmann@22422: lemma sup1E [elim!]: "sup A B x ==> (A x ==> P) ==> (B x ==> P) ==> P" haftmann@32601: by (simp add: sup1_iff) iprover berghofe@22259: haftmann@22422: lemma sup2E [elim!]: "sup A B x y ==> (A x y ==> P) ==> (B x y ==> P) ==> P" haftmann@32601: by (simp add: sup2_iff) iprover berghofe@22259: berghofe@22259: haftmann@30328: subsubsection {* Binary intersection *} berghofe@22259: haftmann@32601: lemma inf1_iff: "inf A B x \ A x \ B x" haftmann@22422: by (simp add: inf_fun_eq inf_bool_eq) berghofe@22259: haftmann@32601: lemma inf2_iff: "inf A B x y \ A x y \ B x y" haftmann@22422: by (simp add: inf_fun_eq inf_bool_eq) berghofe@22259: berghofe@23741: lemma inf_Int_eq [pred_set_conv]: "inf (\x. x \ R) (\x. x \ S) = (\x. x \ R \ S)" haftmann@32601: by (simp add: inf1_iff expand_fun_eq) berghofe@23741: berghofe@23741: lemma inf_Int_eq2 [pred_set_conv]: "inf (\x y. (x, y) \ R) (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" haftmann@32601: by (simp add: inf2_iff expand_fun_eq) berghofe@23741: haftmann@22422: lemma inf1I [intro!]: "A x ==> B x ==> inf A B x" haftmann@32601: by (simp add: inf1_iff) berghofe@22259: haftmann@22422: lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y" haftmann@32601: by (simp add: inf2_iff) berghofe@22259: haftmann@22422: lemma inf1D1: "inf A B x ==> A x" haftmann@32601: by (simp add: inf1_iff) berghofe@22259: haftmann@22422: lemma inf2D1: "inf A B x y ==> A x y" haftmann@32601: by (simp add: inf2_iff) berghofe@22259: haftmann@22422: lemma inf1D2: "inf A B x ==> B x" haftmann@32601: by (simp add: inf1_iff) berghofe@22259: haftmann@22422: lemma inf2D2: "inf A B x y ==> B x y" haftmann@32601: by (simp add: inf2_iff) berghofe@22259: haftmann@22422: lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P" haftmann@32601: by (simp add: inf1_iff) berghofe@22259: haftmann@22422: lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P" haftmann@32601: by (simp add: inf2_iff) berghofe@22259: berghofe@22259: haftmann@30328: subsubsection {* Unions of families *} berghofe@22259: haftmann@32601: lemma SUP1_iff: "(SUP x:A. B x) b = (EX x:A. B x b)" haftmann@24345: by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast berghofe@22430: haftmann@32601: lemma SUP2_iff: "(SUP x:A. B x) b c = (EX x:A. B x b c)" haftmann@24345: by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast berghofe@22259: berghofe@22430: lemma SUP1_I [intro]: "a : A ==> B a b ==> (SUP x:A. B x) b" haftmann@32601: by (auto simp add: SUP1_iff) berghofe@22259: berghofe@22430: lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c" haftmann@32601: by (auto simp add: SUP2_iff) berghofe@22430: berghofe@22430: lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R" haftmann@32601: by (auto simp add: SUP1_iff) berghofe@22430: berghofe@22430: lemma SUP2_E [elim!]: "(SUP x:A. B x) b c ==> (!!x. x : A ==> B x b c ==> R) ==> R" haftmann@32601: by (auto simp add: SUP2_iff) berghofe@22259: berghofe@23741: lemma SUP_UN_eq: "(SUP i. (\x. x \ r i)) = (\x. x \ (UN i. r i))" haftmann@32601: by (simp add: SUP1_iff expand_fun_eq) berghofe@22430: berghofe@23741: lemma SUP_UN_eq2: "(SUP i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (UN i. r i))" haftmann@32601: by (simp add: SUP2_iff expand_fun_eq) berghofe@22430: berghofe@23741: haftmann@30328: subsubsection {* Intersections of families *} berghofe@22430: haftmann@32601: lemma INF1_iff: "(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: haftmann@32601: lemma INF2_iff: "(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" haftmann@32601: by (auto simp add: INF1_iff) berghofe@22259: berghofe@22430: lemma INF2_I [intro!]: "(!!x. x : A ==> B x b c) ==> (INF x:A. B x) b c" haftmann@32601: by (auto simp add: INF2_iff) berghofe@22430: berghofe@22430: lemma INF1_D [elim]: "(INF x:A. B x) b ==> a : A ==> B a b" haftmann@32601: by (auto simp add: INF1_iff) berghofe@22259: berghofe@22430: lemma INF2_D [elim]: "(INF x:A. B x) b c ==> a : A ==> B a b c" haftmann@32601: by (auto simp add: INF2_iff) berghofe@22430: berghofe@22430: lemma INF1_E [elim]: "(INF x:A. B x) b ==> (B a b ==> R) ==> (a ~: A ==> R) ==> R" haftmann@32601: by (auto simp add: INF1_iff) berghofe@22430: berghofe@22430: lemma INF2_E [elim]: "(INF x:A. B x) b c ==> (B a b c ==> R) ==> (a ~: A ==> R) ==> R" haftmann@32601: by (auto simp add: INF2_iff) berghofe@22259: berghofe@23741: lemma INF_INT_eq: "(INF i. (\x. x \ r i)) = (\x. x \ (INT i. r i))" haftmann@32601: by (simp add: INF1_iff expand_fun_eq) berghofe@23741: berghofe@23741: lemma INF_INT_eq2: "(INF i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (INT i. r i))" haftmann@32601: by (simp add: INF2_iff expand_fun_eq) berghofe@23741: berghofe@22259: haftmann@30328: subsection {* Predicates as relations *} haftmann@30328: haftmann@30328: subsubsection {* Composition *} berghofe@22259: berghofe@23741: inductive krauss@32235: pred_comp :: "['a => 'b => bool, 'b => 'c => bool] => 'a => 'c => bool" berghofe@22259: (infixr "OO" 75) krauss@32235: for r :: "'a => 'b => bool" and s :: "'b => 'c => bool" berghofe@22259: where krauss@32235: pred_compI [intro]: "r a b ==> s b c ==> (r OO s) a c" berghofe@22259: berghofe@23741: inductive_cases pred_compE [elim!]: "(r OO s) a c" berghofe@22259: berghofe@22259: lemma pred_comp_rel_comp_eq [pred_set_conv]: berghofe@23741: "((\x y. (x, y) \ r) OO (\x y. (x, y) \ s)) = (\x y. (x, y) \ r O s)" berghofe@22259: by (auto simp add: expand_fun_eq elim: pred_compE) berghofe@22259: berghofe@22259: haftmann@30328: subsubsection {* Converse *} berghofe@22259: berghofe@23741: inductive 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@23741: "(\x y. (x, y) \ r)^--1 = (\x y. (x, y) \ 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: haftmann@30328: subsubsection {* Domain *} berghofe@22259: berghofe@23741: inductive 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@23741: inductive_cases DomainPE [elim!]: "DomainP r a" berghofe@22259: berghofe@23741: lemma DomainP_Domain_eq [pred_set_conv]: "DomainP (\x y. (x, y) \ r) = (\x. x \ Domain r)" berghofe@26797: by (blast intro!: Orderings.order_antisym predicate1I) berghofe@22259: berghofe@22259: haftmann@30328: subsubsection {* Range *} berghofe@22259: berghofe@23741: inductive 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@23741: inductive_cases RangePE [elim!]: "RangeP r b" berghofe@22259: berghofe@23741: lemma RangeP_Range_eq [pred_set_conv]: "RangeP (\x y. (x, y) \ r) = (\x. x \ Range r)" berghofe@26797: by (blast intro!: Orderings.order_antisym predicate1I) berghofe@22259: berghofe@22259: haftmann@30328: subsubsection {* 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@23741: lemma [pred_set_conv]: "inv_imagep (\x y. (x, y) \ r) f = (\x y. (x, y) \ 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: haftmann@30328: subsubsection {* Powerset *} berghofe@23741: berghofe@23741: definition Powp :: "('a \ bool) \ 'a set \ bool" where berghofe@23741: "Powp A == \B. \x \ B. A x" berghofe@23741: berghofe@23741: lemma Powp_Pow_eq [pred_set_conv]: "Powp (\x. x \ A) = (\x. x \ Pow A)" berghofe@23741: by (auto simp add: Powp_def expand_fun_eq) berghofe@23741: berghofe@26797: lemmas Powp_mono [mono] = Pow_mono [to_pred pred_subset_eq] berghofe@26797: berghofe@23741: haftmann@30328: subsubsection {* Properties of relations *} berghofe@22259: berghofe@22259: abbreviation antisymP :: "('a => 'a => bool) => bool" where berghofe@23741: "antisymP r == antisym {(x, y). r x y}" berghofe@22259: berghofe@22259: abbreviation transP :: "('a => 'a => bool) => bool" where berghofe@23741: "transP r == trans {(x, y). r x y}" berghofe@22259: berghofe@22259: abbreviation single_valuedP :: "('a => 'b => bool) => bool" where berghofe@23741: "single_valuedP r == single_valued {(x, y). r x y}" berghofe@22259: haftmann@30328: haftmann@30328: subsection {* Predicates as enumerations *} haftmann@30328: haftmann@30328: subsubsection {* The type of predicate enumerations (a monad) *} haftmann@30328: haftmann@30328: datatype 'a pred = Pred "'a \ bool" haftmann@30328: haftmann@30328: primrec eval :: "'a pred \ 'a \ bool" where haftmann@30328: eval_pred: "eval (Pred f) = f" haftmann@30328: haftmann@30328: lemma Pred_eval [simp]: haftmann@30328: "Pred (eval x) = x" haftmann@30328: by (cases x) simp haftmann@30328: haftmann@30328: lemma eval_inject: "eval x = eval y \ x = y" haftmann@30328: by (cases x) auto haftmann@30328: haftmann@30328: definition single :: "'a \ 'a pred" where haftmann@30328: "single x = Pred ((op =) x)" haftmann@30328: haftmann@30328: definition bind :: "'a pred \ ('a \ 'b pred) \ 'b pred" (infixl "\=" 70) where haftmann@30328: "P \= f = Pred (\x. (\y. eval P y \ eval (f y) x))" haftmann@30328: haftmann@32578: instantiation pred :: (type) "{complete_lattice, boolean_algebra}" haftmann@30328: begin haftmann@30328: haftmann@30328: definition haftmann@30328: "P \ Q \ eval P \ eval Q" haftmann@30328: haftmann@30328: definition haftmann@30328: "P < Q \ eval P < eval Q" haftmann@30328: haftmann@30328: definition haftmann@30328: "\ = Pred \" haftmann@30328: haftmann@30328: definition haftmann@30328: "\ = Pred \" haftmann@30328: haftmann@30328: definition haftmann@30328: "P \ Q = Pred (eval P \ eval Q)" haftmann@30328: haftmann@30328: definition haftmann@30328: "P \ Q = Pred (eval P \ eval Q)" haftmann@30328: haftmann@30328: definition haftmann@31932: [code del]: "\A = Pred (INFI A eval)" haftmann@30328: haftmann@30328: definition haftmann@31932: [code del]: "\A = Pred (SUPR A eval)" haftmann@30328: haftmann@32578: definition haftmann@32578: "- P = Pred (- eval P)" haftmann@32578: haftmann@32578: definition haftmann@32578: "P - Q = Pred (eval P - eval Q)" haftmann@32578: haftmann@32578: instance proof haftmann@32578: qed (auto simp add: less_eq_pred_def less_pred_def haftmann@30328: inf_pred_def sup_pred_def bot_pred_def top_pred_def haftmann@32578: Inf_pred_def Sup_pred_def uminus_pred_def minus_pred_def fun_Compl_def bool_Compl_def, haftmann@30328: auto simp add: le_fun_def less_fun_def le_bool_def less_bool_def haftmann@30328: eval_inject mem_def) haftmann@30328: berghofe@22259: end haftmann@30328: haftmann@30328: lemma bind_bind: haftmann@30328: "(P \= Q) \= R = P \= (\x. Q x \= R)" haftmann@30328: by (auto simp add: bind_def expand_fun_eq) haftmann@30328: haftmann@30328: lemma bind_single: haftmann@30328: "P \= single = P" haftmann@30328: by (simp add: bind_def single_def) haftmann@30328: haftmann@30328: lemma single_bind: haftmann@30328: "single x \= P = P x" haftmann@30328: by (simp add: bind_def single_def) haftmann@30328: haftmann@30328: lemma bottom_bind: haftmann@30328: "\ \= P = \" haftmann@30328: by (auto simp add: bot_pred_def bind_def expand_fun_eq) haftmann@30328: haftmann@30328: lemma sup_bind: haftmann@30328: "(P \ Q) \= R = P \= R \ Q \= R" haftmann@30328: by (auto simp add: bind_def sup_pred_def expand_fun_eq) haftmann@30328: haftmann@30328: lemma Sup_bind: "(\A \= f) = \((\x. x \= f) ` A)" haftmann@32601: by (auto simp add: bind_def Sup_pred_def SUP1_iff expand_fun_eq) haftmann@30328: haftmann@30328: lemma pred_iffI: haftmann@30328: assumes "\x. eval A x \ eval B x" haftmann@30328: and "\x. eval B x \ eval A x" haftmann@30328: shows "A = B" haftmann@30328: proof - haftmann@30328: from assms have "\x. eval A x \ eval B x" by blast haftmann@30328: then show ?thesis by (cases A, cases B) (simp add: expand_fun_eq) haftmann@30328: qed haftmann@30328: haftmann@30328: lemma singleI: "eval (single x) x" haftmann@30328: unfolding single_def by simp haftmann@30328: haftmann@30328: lemma singleI_unit: "eval (single ()) x" haftmann@30328: by simp (rule singleI) haftmann@30328: haftmann@30328: lemma singleE: "eval (single x) y \ (y = x \ P) \ P" haftmann@30328: unfolding single_def by simp haftmann@30328: haftmann@30328: lemma singleE': "eval (single x) y \ (x = y \ P) \ P" haftmann@30328: by (erule singleE) simp haftmann@30328: haftmann@30328: lemma bindI: "eval P x \ eval (Q x) y \ eval (P \= Q) y" haftmann@30328: unfolding bind_def by auto haftmann@30328: haftmann@30328: lemma bindE: "eval (R \= Q) y \ (\x. eval R x \ eval (Q x) y \ P) \ P" haftmann@30328: unfolding bind_def by auto haftmann@30328: haftmann@30328: lemma botE: "eval \ x \ P" haftmann@30328: unfolding bot_pred_def by auto haftmann@30328: haftmann@30328: lemma supI1: "eval A x \ eval (A \ B) x" haftmann@32601: unfolding sup_pred_def by (simp add: sup1_iff) haftmann@30328: haftmann@30328: lemma supI2: "eval B x \ eval (A \ B) x" haftmann@32601: unfolding sup_pred_def by (simp add: sup1_iff) haftmann@30328: haftmann@30328: lemma supE: "eval (A \ B) x \ (eval A x \ P) \ (eval B x \ P) \ P" haftmann@30328: unfolding sup_pred_def by auto haftmann@30328: haftmann@32578: lemma single_not_bot [simp]: haftmann@32578: "single x \ \" haftmann@32578: by (auto simp add: single_def bot_pred_def expand_fun_eq) haftmann@32578: haftmann@32578: lemma not_bot: haftmann@32578: assumes "A \ \" haftmann@32578: obtains x where "eval A x" haftmann@32578: using assms by (cases A) haftmann@32578: (auto simp add: bot_pred_def, auto simp add: mem_def) haftmann@32578: haftmann@32578: haftmann@32578: subsubsection {* Emptiness check and definite choice *} haftmann@32578: haftmann@32578: definition is_empty :: "'a pred \ bool" where haftmann@32578: "is_empty A \ A = \" haftmann@32578: haftmann@32578: lemma is_empty_bot: haftmann@32578: "is_empty \" haftmann@32578: by (simp add: is_empty_def) haftmann@32578: haftmann@32578: lemma not_is_empty_single: haftmann@32578: "\ is_empty (single x)" haftmann@32578: by (auto simp add: is_empty_def single_def bot_pred_def expand_fun_eq) haftmann@32578: haftmann@32578: lemma is_empty_sup: haftmann@32578: "is_empty (A \ B) \ is_empty A \ is_empty B" haftmann@32578: by (auto simp add: is_empty_def intro: sup_eq_bot_eq1 sup_eq_bot_eq2) haftmann@32578: haftmann@32578: definition singleton :: "'a pred \ 'a" where haftmann@32578: "singleton A = (if \!x. eval A x then THE x. eval A x else undefined)" haftmann@32578: haftmann@32578: lemma singleton_eqI: haftmann@32578: "\!x. eval A x \ eval A x \ singleton A = x" haftmann@32578: by (auto simp add: singleton_def) haftmann@32578: haftmann@32578: lemma eval_singletonI: haftmann@32578: "\!x. eval A x \ eval A (singleton A)" haftmann@32578: proof - haftmann@32578: assume assm: "\!x. eval A x" haftmann@32578: then obtain x where "eval A x" .. haftmann@32578: moreover with assm have "singleton A = x" by (rule singleton_eqI) haftmann@32578: ultimately show ?thesis by simp haftmann@32578: qed haftmann@32578: haftmann@32578: lemma single_singleton: haftmann@32578: "\!x. eval A x \ single (singleton A) = A" haftmann@32578: proof - haftmann@32578: assume assm: "\!x. eval A x" haftmann@32578: then have "eval A (singleton A)" haftmann@32578: by (rule eval_singletonI) haftmann@32578: moreover from assm have "\x. eval A x \ singleton A = x" haftmann@32578: by (rule singleton_eqI) haftmann@32578: ultimately have "eval (single (singleton A)) = eval A" haftmann@32578: by (simp (no_asm_use) add: single_def expand_fun_eq) blast haftmann@32578: then show ?thesis by (simp add: eval_inject) haftmann@32578: qed haftmann@32578: haftmann@32578: lemma singleton_undefinedI: haftmann@32578: "\ (\!x. eval A x) \ singleton A = undefined" haftmann@32578: by (simp add: singleton_def) haftmann@32578: haftmann@32578: lemma singleton_bot: haftmann@32578: "singleton \ = undefined" haftmann@32578: by (auto simp add: bot_pred_def intro: singleton_undefinedI) haftmann@32578: haftmann@32578: lemma singleton_single: haftmann@32578: "singleton (single x) = x" haftmann@32578: by (auto simp add: intro: singleton_eqI singleI elim: singleE) haftmann@32578: haftmann@32578: lemma singleton_sup_single_single: haftmann@32578: "singleton (single x \ single y) = (if x = y then x else undefined)" haftmann@32578: proof (cases "x = y") haftmann@32578: case True then show ?thesis by (simp add: singleton_single) haftmann@32578: next haftmann@32578: case False haftmann@32578: have "eval (single x \ single y) x" haftmann@32578: and "eval (single x \ single y) y" haftmann@32578: by (auto intro: supI1 supI2 singleI) haftmann@32578: with False have "\ (\!z. eval (single x \ single y) z)" haftmann@32578: by blast haftmann@32578: then have "singleton (single x \ single y) = undefined" haftmann@32578: by (rule singleton_undefinedI) haftmann@32578: with False show ?thesis by simp haftmann@32578: qed haftmann@32578: haftmann@32578: lemma singleton_sup_aux: haftmann@32578: "singleton (A \ B) = (if A = \ then singleton B haftmann@32578: else if B = \ then singleton A haftmann@32578: else singleton haftmann@32578: (single (singleton A) \ single (singleton B)))" haftmann@32578: proof (cases "(\!x. eval A x) \ (\!y. eval B y)") haftmann@32578: case True then show ?thesis by (simp add: single_singleton) haftmann@32578: next haftmann@32578: case False haftmann@32578: from False have A_or_B: haftmann@32578: "singleton A = undefined \ singleton B = undefined" haftmann@32578: by (auto intro!: singleton_undefinedI) haftmann@32578: then have rhs: "singleton haftmann@32578: (single (singleton A) \ single (singleton B)) = undefined" haftmann@32578: by (auto simp add: singleton_sup_single_single singleton_single) haftmann@32578: from False have not_unique: haftmann@32578: "\ (\!x. eval A x) \ \ (\!y. eval B y)" by simp haftmann@32578: show ?thesis proof (cases "A \ \ \ B \ \") haftmann@32578: case True haftmann@32578: then obtain a b where a: "eval A a" and b: "eval B b" haftmann@32578: by (blast elim: not_bot) haftmann@32578: with True not_unique have "\ (\!x. eval (A \ B) x)" haftmann@32578: by (auto simp add: sup_pred_def bot_pred_def) haftmann@32578: then have "singleton (A \ B) = undefined" by (rule singleton_undefinedI) haftmann@32578: with True rhs show ?thesis by simp haftmann@32578: next haftmann@32578: case False then show ?thesis by auto haftmann@32578: qed haftmann@32578: qed haftmann@32578: haftmann@32578: lemma singleton_sup: haftmann@32578: "singleton (A \ B) = (if A = \ then singleton B haftmann@32578: else if B = \ then singleton A haftmann@32578: else if singleton A = singleton B then singleton A else undefined)" haftmann@32578: using singleton_sup_aux [of A B] by (simp only: singleton_sup_single_single) haftmann@32578: haftmann@30328: haftmann@30328: subsubsection {* Derived operations *} haftmann@30328: haftmann@30328: definition if_pred :: "bool \ unit pred" where haftmann@30328: if_pred_eq: "if_pred b = (if b then single () else \)" haftmann@30328: haftmann@30328: definition not_pred :: "unit pred \ unit pred" where haftmann@30328: not_pred_eq: "not_pred P = (if eval P () then \ else single ())" haftmann@30328: haftmann@30328: lemma if_predI: "P \ eval (if_pred P) ()" haftmann@30328: unfolding if_pred_eq by (auto intro: singleI) haftmann@30328: haftmann@30328: lemma if_predE: "eval (if_pred b) x \ (b \ x = () \ P) \ P" haftmann@30328: unfolding if_pred_eq by (cases b) (auto elim: botE) haftmann@30328: haftmann@30328: lemma not_predI: "\ P \ eval (not_pred (Pred (\u. P))) ()" haftmann@30328: unfolding not_pred_eq eval_pred by (auto intro: singleI) haftmann@30328: haftmann@30328: lemma not_predI': "\ eval P () \ eval (not_pred P) ()" haftmann@30328: unfolding not_pred_eq by (auto intro: singleI) haftmann@30328: haftmann@30328: lemma not_predE: "eval (not_pred (Pred (\u. P))) x \ (\ P \ thesis) \ thesis" haftmann@30328: unfolding not_pred_eq haftmann@30328: by (auto split: split_if_asm elim: botE) haftmann@30328: haftmann@30328: lemma not_predE': "eval (not_pred P) x \ (\ eval P x \ thesis) \ thesis" haftmann@30328: unfolding not_pred_eq haftmann@30328: by (auto split: split_if_asm elim: botE) haftmann@30328: haftmann@30328: haftmann@30328: subsubsection {* Implementation *} haftmann@30328: haftmann@30328: datatype 'a seq = Empty | Insert "'a" "'a pred" | Join "'a pred" "'a seq" haftmann@30328: haftmann@30328: primrec pred_of_seq :: "'a seq \ 'a pred" where haftmann@30328: "pred_of_seq Empty = \" haftmann@30328: | "pred_of_seq (Insert x P) = single x \ P" haftmann@30328: | "pred_of_seq (Join P xq) = P \ pred_of_seq xq" haftmann@30328: haftmann@30328: definition Seq :: "(unit \ 'a seq) \ 'a pred" where haftmann@30328: "Seq f = pred_of_seq (f ())" haftmann@30328: haftmann@30328: code_datatype Seq haftmann@30328: haftmann@30328: primrec member :: "'a seq \ 'a \ bool" where haftmann@30328: "member Empty x \ False" haftmann@30328: | "member (Insert y P) x \ x = y \ eval P x" haftmann@30328: | "member (Join P xq) x \ eval P x \ member xq x" haftmann@30328: haftmann@30328: lemma eval_member: haftmann@30328: "member xq = eval (pred_of_seq xq)" haftmann@30328: proof (induct xq) haftmann@30328: case Empty show ?case haftmann@30328: by (auto simp add: expand_fun_eq elim: botE) haftmann@30328: next haftmann@30328: case Insert show ?case haftmann@30328: by (auto simp add: expand_fun_eq elim: supE singleE intro: supI1 supI2 singleI) haftmann@30328: next haftmann@30328: case Join then show ?case haftmann@30328: by (auto simp add: expand_fun_eq elim: supE intro: supI1 supI2) haftmann@30328: qed haftmann@30328: haftmann@30328: lemma eval_code [code]: "eval (Seq f) = member (f ())" haftmann@30328: unfolding Seq_def by (rule sym, rule eval_member) haftmann@30328: haftmann@30328: lemma single_code [code]: haftmann@30328: "single x = Seq (\u. Insert x \)" haftmann@30328: unfolding Seq_def by simp haftmann@30328: haftmann@30328: primrec "apply" :: "('a \ 'b Predicate.pred) \ 'a seq \ 'b seq" where haftmann@30328: "apply f Empty = Empty" haftmann@30328: | "apply f (Insert x P) = Join (f x) (Join (P \= f) Empty)" haftmann@30328: | "apply f (Join P xq) = Join (P \= f) (apply f xq)" haftmann@30328: haftmann@30328: lemma apply_bind: haftmann@30328: "pred_of_seq (apply f xq) = pred_of_seq xq \= f" haftmann@30328: proof (induct xq) haftmann@30328: case Empty show ?case haftmann@30328: by (simp add: bottom_bind) haftmann@30328: next haftmann@30328: case Insert show ?case haftmann@30328: by (simp add: single_bind sup_bind) haftmann@30328: next haftmann@30328: case Join then show ?case haftmann@30328: by (simp add: sup_bind) haftmann@30328: qed haftmann@30328: haftmann@30328: lemma bind_code [code]: haftmann@30328: "Seq g \= f = Seq (\u. apply f (g ()))" haftmann@30328: unfolding Seq_def by (rule sym, rule apply_bind) haftmann@30328: haftmann@30328: lemma bot_set_code [code]: haftmann@30328: "\ = Seq (\u. Empty)" haftmann@30328: unfolding Seq_def by simp haftmann@30328: haftmann@30376: primrec adjunct :: "'a pred \ 'a seq \ 'a seq" where haftmann@30376: "adjunct P Empty = Join P Empty" haftmann@30376: | "adjunct P (Insert x Q) = Insert x (Q \ P)" haftmann@30376: | "adjunct P (Join Q xq) = Join Q (adjunct P xq)" haftmann@30376: haftmann@30376: lemma adjunct_sup: haftmann@30376: "pred_of_seq (adjunct P xq) = P \ pred_of_seq xq" haftmann@30376: by (induct xq) (simp_all add: sup_assoc sup_commute sup_left_commute) haftmann@30376: haftmann@30328: lemma sup_code [code]: haftmann@30328: "Seq f \ Seq g = Seq (\u. case f () haftmann@30328: of Empty \ g () haftmann@30328: | Insert x P \ Insert x (P \ Seq g) haftmann@30376: | Join P xq \ adjunct (Seq g) (Join P xq))" haftmann@30328: proof (cases "f ()") haftmann@30328: case Empty haftmann@30328: thus ?thesis haftmann@30376: unfolding Seq_def by (simp add: sup_commute [of "\"] sup_bot) haftmann@30328: next haftmann@30328: case Insert haftmann@30328: thus ?thesis haftmann@30328: unfolding Seq_def by (simp add: sup_assoc) haftmann@30328: next haftmann@30328: case Join haftmann@30328: thus ?thesis haftmann@30376: unfolding Seq_def haftmann@30376: by (simp add: adjunct_sup sup_assoc sup_commute sup_left_commute) haftmann@30328: qed haftmann@30328: haftmann@30430: primrec contained :: "'a seq \ 'a pred \ bool" where haftmann@30430: "contained Empty Q \ True" haftmann@30430: | "contained (Insert x P) Q \ eval Q x \ P \ Q" haftmann@30430: | "contained (Join P xq) Q \ P \ Q \ contained xq Q" haftmann@30430: haftmann@30430: lemma single_less_eq_eval: haftmann@30430: "single x \ P \ eval P x" haftmann@30430: by (auto simp add: single_def less_eq_pred_def mem_def) haftmann@30430: haftmann@30430: lemma contained_less_eq: haftmann@30430: "contained xq Q \ pred_of_seq xq \ Q" haftmann@30430: by (induct xq) (simp_all add: single_less_eq_eval) haftmann@30430: haftmann@30430: lemma less_eq_pred_code [code]: haftmann@30430: "Seq f \ Q = (case f () haftmann@30430: of Empty \ True haftmann@30430: | Insert x P \ eval Q x \ P \ Q haftmann@30430: | Join P xq \ P \ Q \ contained xq Q)" haftmann@30430: by (cases "f ()") haftmann@30430: (simp_all add: Seq_def single_less_eq_eval contained_less_eq) haftmann@30430: haftmann@30430: lemma eq_pred_code [code]: haftmann@31133: fixes P Q :: "'a pred" haftmann@30430: shows "eq_class.eq P Q \ P \ Q \ Q \ P" haftmann@30430: unfolding eq by auto haftmann@30430: haftmann@30430: lemma [code]: haftmann@30430: "pred_case f P = f (eval P)" haftmann@30430: by (cases P) simp haftmann@30430: haftmann@30430: lemma [code]: haftmann@30430: "pred_rec f P = f (eval P)" haftmann@30430: by (cases P) simp haftmann@30328: bulwahn@31105: inductive eq :: "'a \ 'a \ bool" where "eq x x" bulwahn@31105: bulwahn@31105: lemma eq_is_eq: "eq x y \ (x = y)" haftmann@31108: by (rule eq_reflection) (auto intro: eq.intros elim: eq.cases) haftmann@30948: haftmann@31216: definition map :: "('a \ 'b) \ 'a pred \ 'b pred" where haftmann@31216: "map f P = P \= (single o f)" haftmann@31216: haftmann@32578: primrec null :: "'a seq \ bool" where haftmann@32578: "null Empty \ True" haftmann@32578: | "null (Insert x P) \ False" haftmann@32578: | "null (Join P xq) \ is_empty P \ null xq" haftmann@32578: haftmann@32578: lemma null_is_empty: haftmann@32578: "null xq \ is_empty (pred_of_seq xq)" haftmann@32578: by (induct xq) (simp_all add: is_empty_bot not_is_empty_single is_empty_sup) haftmann@32578: haftmann@32578: lemma is_empty_code [code]: haftmann@32578: "is_empty (Seq f) \ null (f ())" haftmann@32578: by (simp add: null_is_empty Seq_def) haftmann@32578: haftmann@32578: primrec the_only :: "'a seq \ 'a" where haftmann@32578: [code del]: "the_only Empty = undefined" haftmann@32578: | "the_only (Insert x P) = (if is_empty P then x else let y = singleton P in if x = y then x else undefined)" haftmann@32578: | "the_only (Join P xq) = (if is_empty P then the_only xq else if null xq then singleton P haftmann@32578: else let x = singleton P; y = the_only xq in haftmann@32578: if x = y then x else undefined)" haftmann@32578: haftmann@32578: lemma the_only_singleton: haftmann@32578: "the_only xq = singleton (pred_of_seq xq)" haftmann@32578: by (induct xq) haftmann@32578: (auto simp add: singleton_bot singleton_single is_empty_def haftmann@32578: null_is_empty Let_def singleton_sup) haftmann@32578: haftmann@32578: lemma singleton_code [code]: haftmann@32578: "singleton (Seq f) = (case f () haftmann@32578: of Empty \ undefined haftmann@32578: | Insert x P \ if is_empty P then x haftmann@32578: else let y = singleton P in haftmann@32578: if x = y then x else undefined haftmann@32578: | Join P xq \ if is_empty P then the_only xq haftmann@32578: else if null xq then singleton P haftmann@32578: else let x = singleton P; y = the_only xq in haftmann@32578: if x = y then x else undefined)" haftmann@32578: by (cases "f ()") haftmann@32578: (auto simp add: Seq_def the_only_singleton is_empty_def haftmann@32578: null_is_empty singleton_bot singleton_single singleton_sup Let_def) haftmann@32578: haftmann@30948: ML {* haftmann@30948: signature PREDICATE = haftmann@30948: sig haftmann@30948: datatype 'a pred = Seq of (unit -> 'a seq) haftmann@30948: and 'a seq = Empty | Insert of 'a * 'a pred | Join of 'a pred * 'a seq haftmann@30959: val yield: 'a pred -> ('a * 'a pred) option haftmann@30959: val yieldn: int -> 'a pred -> 'a list * 'a pred haftmann@31222: val map: ('a -> 'b) -> 'a pred -> 'b pred haftmann@30948: end; haftmann@30948: haftmann@30948: structure Predicate : PREDICATE = haftmann@30948: struct haftmann@30948: haftmann@30959: @{code_datatype pred = Seq}; haftmann@30959: @{code_datatype seq = Empty | Insert | Join}; haftmann@30959: haftmann@32372: fun yield (@{code Seq} f) = next (f ()) haftmann@30959: and next @{code Empty} = NONE haftmann@30959: | next (@{code Insert} (x, P)) = SOME (x, P) haftmann@30959: | next (@{code Join} (P, xq)) = (case yield P haftmann@30959: of NONE => next xq haftmann@30959: | SOME (x, Q) => SOME (x, @{code Seq} (fn _ => @{code Join} (Q, xq)))) haftmann@30959: haftmann@30959: fun anamorph f k x = (if k = 0 then ([], x) haftmann@30959: else case f x haftmann@30959: of NONE => ([], x) haftmann@30959: | SOME (v, y) => let haftmann@30959: val (vs, z) = anamorph f (k - 1) y haftmann@30959: in (v :: vs, z) end) haftmann@30959: haftmann@30959: fun yieldn P = anamorph yield P; haftmann@30948: haftmann@31222: fun map f = @{code map} f; haftmann@31222: haftmann@30948: end; haftmann@30948: *} haftmann@30948: haftmann@30948: code_reserved Eval Predicate haftmann@30948: haftmann@30948: code_type pred and seq haftmann@30948: (Eval "_/ Predicate.pred" and "_/ Predicate.seq") haftmann@30948: haftmann@30948: code_const Seq and Empty and Insert and Join haftmann@30948: (Eval "Predicate.Seq" and "Predicate.Empty" and "Predicate.Insert/ (_,/ _)" and "Predicate.Join/ (_,/ _)") haftmann@30948: haftmann@31122: text {* dummy setup for @{text code_pred} and @{text values} keywords *} haftmann@31108: haftmann@31108: ML {* haftmann@31122: local haftmann@31122: haftmann@31122: structure P = OuterParse; haftmann@31122: haftmann@31122: val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) []; haftmann@31122: haftmann@31122: in haftmann@31122: haftmann@31122: val _ = OuterSyntax.local_theory_to_proof "code_pred" "sets up goal for cases rule from given introduction rules and compiles predicate" haftmann@31122: OuterKeyword.thy_goal (P.term_group >> (K (Proof.theorem_i NONE (K I) [[]]))); haftmann@31122: haftmann@31216: val _ = OuterSyntax.improper_command "values" "enumerate and print comprehensions" haftmann@31122: OuterKeyword.diag ((opt_modes -- P.term) haftmann@31122: >> (fn (modes, t) => Toplevel.no_timing o Toplevel.keep haftmann@31122: (K ()))); haftmann@31122: haftmann@31122: end haftmann@31108: *} haftmann@30959: haftmann@30328: no_notation haftmann@30328: inf (infixl "\" 70) and haftmann@30328: sup (infixl "\" 65) and haftmann@30328: Inf ("\_" [900] 900) and haftmann@30328: Sup ("\_" [900] 900) and haftmann@30328: top ("\") and haftmann@30328: bot ("\") and haftmann@30328: bind (infixl "\=" 70) haftmann@30328: haftmann@30328: hide (open) type pred seq haftmann@32582: hide (open) const Pred eval single bind is_empty singleton if_pred not_pred haftmann@32582: Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map haftmann@30328: haftmann@30328: end