berghofe@22259: (* Title: HOL/Predicate.thy haftmann@46664: Author: Lukas Bulwahn and Florian Haftmann, TU Muenchen berghofe@22259: *) berghofe@22259: haftmann@46664: header {* Predicates as enumerations *} berghofe@22259: berghofe@22259: theory Predicate Andreas@53943: imports String berghofe@22259: begin berghofe@22259: haftmann@46664: subsection {* 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@40616: lemma pred_eqI: haftmann@40616: "(\w. eval P w \ eval Q w) \ P = Q" haftmann@40616: by (cases P, cases Q) (auto simp add: fun_eq_iff) haftmann@30328: haftmann@46038: lemma pred_eq_iff: haftmann@46038: "P = Q \ (\w. eval P w \ eval Q w)" haftmann@46038: by (simp add: pred_eqI) haftmann@46038: haftmann@44033: instantiation pred :: (type) complete_lattice 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@40616: lemma eval_bot [simp]: haftmann@40616: "eval \ = \" haftmann@40616: by (simp add: bot_pred_def) haftmann@40616: haftmann@30328: definition haftmann@30328: "\ = Pred \" haftmann@30328: haftmann@40616: lemma eval_top [simp]: haftmann@40616: "eval \ = \" haftmann@40616: by (simp add: top_pred_def) haftmann@40616: haftmann@30328: definition haftmann@30328: "P \ Q = Pred (eval P \ eval Q)" haftmann@30328: haftmann@40616: lemma eval_inf [simp]: haftmann@40616: "eval (P \ Q) = eval P \ eval Q" haftmann@40616: by (simp add: inf_pred_def) haftmann@40616: haftmann@30328: definition haftmann@30328: "P \ Q = Pred (eval P \ eval Q)" haftmann@30328: haftmann@40616: lemma eval_sup [simp]: haftmann@40616: "eval (P \ Q) = eval P \ eval Q" haftmann@40616: by (simp add: sup_pred_def) haftmann@40616: haftmann@30328: definition haftmann@37767: "\A = Pred (INFI A eval)" haftmann@30328: haftmann@40616: lemma eval_Inf [simp]: haftmann@40616: "eval (\A) = INFI A eval" haftmann@40616: by (simp add: Inf_pred_def) haftmann@40616: haftmann@30328: definition haftmann@37767: "\A = Pred (SUPR A eval)" haftmann@30328: haftmann@40616: lemma eval_Sup [simp]: haftmann@40616: "eval (\A) = SUPR A eval" haftmann@40616: by (simp add: Sup_pred_def) haftmann@40616: haftmann@44033: instance proof haftmann@44415: qed (auto intro!: pred_eqI simp add: less_eq_pred_def less_pred_def le_fun_def less_fun_def) haftmann@44033: haftmann@44033: end haftmann@44033: haftmann@44033: lemma eval_INFI [simp]: haftmann@44033: "eval (INFI A f) = INFI A (eval \ f)" hoelzl@44928: by (simp only: INF_def eval_Inf image_compose) haftmann@44033: haftmann@44033: lemma eval_SUPR [simp]: haftmann@44033: "eval (SUPR A f) = SUPR A (eval \ f)" hoelzl@44928: by (simp only: SUP_def eval_Sup image_compose) haftmann@44033: haftmann@44033: instantiation pred :: (type) complete_boolean_algebra haftmann@44033: begin haftmann@44033: haftmann@32578: definition haftmann@32578: "- P = Pred (- eval P)" haftmann@32578: haftmann@40616: lemma eval_compl [simp]: haftmann@40616: "eval (- P) = - eval P" haftmann@40616: by (simp add: uminus_pred_def) haftmann@40616: haftmann@32578: definition haftmann@32578: "P - Q = Pred (eval P - eval Q)" haftmann@32578: haftmann@40616: lemma eval_minus [simp]: haftmann@40616: "eval (P - Q) = eval P - eval Q" haftmann@40616: by (simp add: minus_pred_def) haftmann@40616: haftmann@32578: instance proof noschinl@46884: qed (auto intro!: pred_eqI) haftmann@30328: berghofe@22259: end haftmann@30328: haftmann@40616: definition single :: "'a \ 'a pred" where haftmann@40616: "single x = Pred ((op =) x)" haftmann@40616: haftmann@40616: lemma eval_single [simp]: haftmann@40616: "eval (single x) = (op =) x" haftmann@40616: by (simp add: single_def) haftmann@40616: haftmann@40616: definition bind :: "'a pred \ ('a \ 'b pred) \ 'b pred" (infixl "\=" 70) where haftmann@41080: "P \= f = (SUPR {x. eval P x} f)" haftmann@40616: haftmann@40616: lemma eval_bind [simp]: haftmann@41080: "eval (P \= f) = eval (SUPR {x. eval P x} f)" haftmann@40616: by (simp add: bind_def) haftmann@40616: haftmann@30328: lemma bind_bind: haftmann@30328: "(P \= Q) \= R = P \= (\x. Q x \= R)" noschinl@46884: by (rule pred_eqI) auto haftmann@30328: haftmann@30328: lemma bind_single: haftmann@30328: "P \= single = P" haftmann@40616: by (rule pred_eqI) auto haftmann@30328: haftmann@30328: lemma single_bind: haftmann@30328: "single x \= P = P x" haftmann@40616: by (rule pred_eqI) auto haftmann@30328: haftmann@30328: lemma bottom_bind: haftmann@30328: "\ \= P = \" haftmann@40674: by (rule pred_eqI) auto haftmann@30328: haftmann@30328: lemma sup_bind: haftmann@30328: "(P \ Q) \= R = P \= R \ Q \= R" haftmann@40674: by (rule pred_eqI) auto haftmann@30328: haftmann@40616: lemma Sup_bind: haftmann@40616: "(\A \= f) = \((\x. x \= f) ` A)" noschinl@46884: by (rule pred_eqI) auto 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@40616: using assms by (auto intro: pred_eqI) haftmann@30328: haftmann@30328: lemma singleI: "eval (single x) x" haftmann@40616: by simp haftmann@30328: haftmann@30328: lemma singleI_unit: "eval (single ()) x" haftmann@40616: by simp haftmann@30328: haftmann@30328: lemma singleE: "eval (single x) y \ (y = x \ P) \ P" haftmann@40616: by simp haftmann@30328: haftmann@30328: lemma singleE': "eval (single x) y \ (x = y \ P) \ P" haftmann@40616: by simp haftmann@30328: haftmann@30328: lemma bindI: "eval P x \ eval (Q x) y \ eval (P \= Q) y" haftmann@40616: by auto haftmann@30328: haftmann@30328: lemma bindE: "eval (R \= Q) y \ (\x. eval R x \ eval (Q x) y \ P) \ P" haftmann@40616: by auto haftmann@30328: haftmann@30328: lemma botE: "eval \ x \ P" haftmann@40616: by auto haftmann@30328: haftmann@30328: lemma supI1: "eval A x \ eval (A \ B) x" haftmann@40616: by auto haftmann@30328: haftmann@30328: lemma supI2: "eval B x \ eval (A \ B) x" haftmann@40616: by auto haftmann@30328: haftmann@30328: lemma supE: "eval (A \ B) x \ (eval A x \ P) \ (eval B x \ P) \ P" haftmann@40616: by auto haftmann@30328: haftmann@32578: lemma single_not_bot [simp]: haftmann@32578: "single x \ \" nipkow@39302: by (auto simp add: single_def bot_pred_def fun_eq_iff) haftmann@32578: haftmann@32578: lemma not_bot: haftmann@32578: assumes "A \ \" haftmann@32578: obtains x where "eval A x" haftmann@45970: using assms by (cases A) (auto simp add: bot_pred_def) haftmann@45970: haftmann@32578: haftmann@46664: subsection {* 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)" nipkow@39302: by (auto simp add: is_empty_def single_def bot_pred_def fun_eq_iff) haftmann@32578: haftmann@32578: lemma is_empty_sup: haftmann@32578: "is_empty (A \ B) \ is_empty A \ is_empty B" huffman@36008: by (auto simp add: is_empty_def) haftmann@32578: haftmann@40616: definition singleton :: "(unit \ 'a) \ 'a pred \ 'a" where bulwahn@33111: "singleton dfault A = (if \!x. eval A x then THE x. eval A x else dfault ())" haftmann@32578: haftmann@32578: lemma singleton_eqI: bulwahn@33110: "\!x. eval A x \ eval A x \ singleton dfault A = x" haftmann@32578: by (auto simp add: singleton_def) haftmann@32578: haftmann@32578: lemma eval_singletonI: bulwahn@33110: "\!x. eval A x \ eval A (singleton dfault A)" haftmann@32578: proof - haftmann@32578: assume assm: "\!x. eval A x" wenzelm@53374: then obtain x where x: "eval A x" .. wenzelm@53374: with assm have "singleton dfault A = x" by (rule singleton_eqI) wenzelm@53374: with x show ?thesis by simp haftmann@32578: qed haftmann@32578: haftmann@32578: lemma single_singleton: bulwahn@33110: "\!x. eval A x \ single (singleton dfault A) = A" haftmann@32578: proof - haftmann@32578: assume assm: "\!x. eval A x" bulwahn@33110: then have "eval A (singleton dfault A)" haftmann@32578: by (rule eval_singletonI) bulwahn@33110: moreover from assm have "\x. eval A x \ singleton dfault A = x" haftmann@32578: by (rule singleton_eqI) bulwahn@33110: ultimately have "eval (single (singleton dfault A)) = eval A" nipkow@39302: by (simp (no_asm_use) add: single_def fun_eq_iff) blast haftmann@40616: then have "\x. eval (single (singleton dfault A)) x = eval A x" haftmann@40616: by simp haftmann@40616: then show ?thesis by (rule pred_eqI) haftmann@32578: qed haftmann@32578: haftmann@32578: lemma singleton_undefinedI: bulwahn@33111: "\ (\!x. eval A x) \ singleton dfault A = dfault ()" haftmann@32578: by (simp add: singleton_def) haftmann@32578: haftmann@32578: lemma singleton_bot: bulwahn@33111: "singleton dfault \ = dfault ()" haftmann@32578: by (auto simp add: bot_pred_def intro: singleton_undefinedI) haftmann@32578: haftmann@32578: lemma singleton_single: bulwahn@33110: "singleton dfault (single x) = x" haftmann@32578: by (auto simp add: intro: singleton_eqI singleI elim: singleE) haftmann@32578: haftmann@32578: lemma singleton_sup_single_single: bulwahn@33111: "singleton dfault (single x \ single y) = (if x = y then x else dfault ())" 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 bulwahn@33111: then have "singleton dfault (single x \ single y) = dfault ()" 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: bulwahn@33110: "singleton dfault (A \ B) = (if A = \ then singleton dfault B bulwahn@33110: else if B = \ then singleton dfault A bulwahn@33110: else singleton dfault bulwahn@33110: (single (singleton dfault A) \ single (singleton dfault 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: bulwahn@33111: "singleton dfault A = dfault () \ singleton dfault B = dfault ()" haftmann@32578: by (auto intro!: singleton_undefinedI) bulwahn@33110: then have rhs: "singleton dfault bulwahn@33111: (single (singleton dfault A) \ single (singleton dfault B)) = dfault ()" 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) bulwahn@33111: then have "singleton dfault (A \ B) = dfault ()" 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: bulwahn@33110: "singleton dfault (A \ B) = (if A = \ then singleton dfault B bulwahn@33110: else if B = \ then singleton dfault A bulwahn@33111: else if singleton dfault A = singleton dfault B then singleton dfault A else dfault ())" bulwahn@33110: using singleton_sup_aux [of dfault A B] by (simp only: singleton_sup_single_single) haftmann@32578: haftmann@30328: haftmann@46664: subsection {* 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: bulwahn@33754: definition holds :: "unit pred \ bool" where bulwahn@33754: holds_eq: "holds P = eval P ()" bulwahn@33754: 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) bulwahn@33754: lemma "f () = False \ f () = True" bulwahn@33754: by simp haftmann@30328: blanchet@37549: lemma closure_of_bool_cases [no_atp]: haftmann@44007: fixes f :: "unit \ bool" haftmann@44007: assumes "f = (\u. False) \ P f" haftmann@44007: assumes "f = (\u. True) \ P f" haftmann@44007: shows "P f" bulwahn@33754: proof - haftmann@44007: have "f = (\u. False) \ f = (\u. True)" bulwahn@33754: apply (cases "f ()") bulwahn@33754: apply (rule disjI2) bulwahn@33754: apply (rule ext) bulwahn@33754: apply (simp add: unit_eq) bulwahn@33754: apply (rule disjI1) bulwahn@33754: apply (rule ext) bulwahn@33754: apply (simp add: unit_eq) bulwahn@33754: done wenzelm@41550: from this assms show ?thesis by blast bulwahn@33754: qed bulwahn@33754: bulwahn@33754: lemma unit_pred_cases: haftmann@44007: assumes "P \" haftmann@44007: assumes "P (single ())" haftmann@44007: shows "P Q" haftmann@44415: using assms unfolding bot_pred_def bot_fun_def bot_bool_def empty_def single_def proof (cases Q) haftmann@44007: fix f haftmann@44007: assume "P (Pred (\u. False))" "P (Pred (\u. () = u))" haftmann@44007: then have "P (Pred f)" haftmann@44007: by (cases _ f rule: closure_of_bool_cases) simp_all haftmann@44007: moreover assume "Q = Pred f" haftmann@44007: ultimately show "P Q" by simp haftmann@44007: qed haftmann@44007: bulwahn@33754: lemma holds_if_pred: bulwahn@33754: "holds (if_pred b) = b" bulwahn@33754: unfolding if_pred_eq holds_eq bulwahn@33754: by (cases b) (auto intro: singleI elim: botE) bulwahn@33754: bulwahn@33754: lemma if_pred_holds: bulwahn@33754: "if_pred (holds P) = P" bulwahn@33754: unfolding if_pred_eq holds_eq bulwahn@33754: by (rule unit_pred_cases) (auto intro: singleI elim: botE) bulwahn@33754: bulwahn@33754: lemma is_empty_holds: bulwahn@33754: "is_empty P \ \ holds P" bulwahn@33754: unfolding is_empty_def holds_eq bulwahn@33754: by (rule unit_pred_cases) (auto elim: botE intro: singleI) haftmann@30328: haftmann@41311: definition map :: "('a \ 'b) \ 'a pred \ 'b pred" where haftmann@41311: "map f P = P \= (single o f)" haftmann@41311: haftmann@41311: lemma eval_map [simp]: haftmann@44363: "eval (map f P) = (\x\{x. eval P x}. (\y. f x = y))" haftmann@44415: by (auto simp add: map_def comp_def) haftmann@41311: haftmann@41505: enriched_type map: map haftmann@44363: by (rule ext, rule pred_eqI, auto)+ haftmann@41311: haftmann@41311: haftmann@46664: subsection {* 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@44414: "pred_of_seq Empty = \" haftmann@44414: | "pred_of_seq (Insert x P) = single x \ P" haftmann@44414: | "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@44414: | "member (Insert y P) x \ x = y \ eval P x" haftmann@44414: | "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 nipkow@39302: by (auto simp add: fun_eq_iff elim: botE) haftmann@30328: next haftmann@30328: case Insert show ?case nipkow@39302: by (auto simp add: fun_eq_iff elim: supE singleE intro: supI1 supI2 singleI) haftmann@30328: next haftmann@30328: case Join then show ?case nipkow@39302: by (auto simp add: fun_eq_iff elim: supE intro: supI1 supI2) haftmann@30328: qed haftmann@30328: haftmann@46038: lemma eval_code [(* FIXME declare simp *)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@41080: primrec "apply" :: "('a \ 'b pred) \ 'a seq \ 'b seq" where haftmann@44415: "apply f Empty = Empty" haftmann@44415: | "apply f (Insert x P) = Join (f x) (Join (P \= f) Empty)" haftmann@44415: | "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@44415: "adjunct P Empty = Join P Empty" haftmann@44415: | "adjunct P (Insert x Q) = Insert x (Q \ P)" haftmann@44415: | "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@34007: unfolding Seq_def by (simp add: sup_commute [of "\"]) 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@46664: lemma [code]: haftmann@46664: "size (P :: 'a Predicate.pred) = 0" by (cases P) simp haftmann@46664: haftmann@46664: lemma [code]: haftmann@46664: "pred_size f P = 0" by (cases P) simp haftmann@46664: haftmann@30430: primrec contained :: "'a seq \ 'a pred \ bool" where haftmann@44415: "contained Empty Q \ True" haftmann@44415: | "contained (Insert x P) Q \ eval Q x \ P \ Q" haftmann@44415: | "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@44415: by (auto simp add: less_eq_pred_def le_fun_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@38857: shows "HOL.equal P Q \ P \ Q \ Q \ P" haftmann@38857: by (auto simp add: equal) haftmann@38857: haftmann@38857: lemma [code nbe]: haftmann@38857: "HOL.equal (x :: 'a pred) x \ True" haftmann@38857: by (fact equal_refl) 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@32578: primrec null :: "'a seq \ bool" where haftmann@44415: "null Empty \ True" haftmann@44415: | "null (Insert x P) \ False" haftmann@44415: | "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: bulwahn@33111: primrec the_only :: "(unit \ 'a) \ 'a seq \ 'a" where bulwahn@33111: [code del]: "the_only dfault Empty = dfault ()" haftmann@44415: | "the_only dfault (Insert x P) = (if is_empty P then x else let y = singleton dfault P in if x = y then x else dfault ())" haftmann@44415: | "the_only dfault (Join P xq) = (if is_empty P then the_only dfault xq else if null xq then singleton dfault P bulwahn@33110: else let x = singleton dfault P; y = the_only dfault xq in bulwahn@33111: if x = y then x else dfault ())" haftmann@32578: haftmann@32578: lemma the_only_singleton: bulwahn@33110: "the_only dfault xq = singleton dfault (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]: bulwahn@33110: "singleton dfault (Seq f) = (case f () bulwahn@33111: of Empty \ dfault () haftmann@32578: | Insert x P \ if is_empty P then x bulwahn@33110: else let y = singleton dfault P in bulwahn@33111: if x = y then x else dfault () bulwahn@33110: | Join P xq \ if is_empty P then the_only dfault xq bulwahn@33110: else if null xq then singleton dfault P bulwahn@33110: else let x = singleton dfault P; y = the_only dfault xq in bulwahn@33111: if x = y then x else dfault ())" 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@44414: definition the :: "'a pred \ 'a" where haftmann@37767: "the A = (THE x. eval A x)" bulwahn@33111: haftmann@40674: lemma the_eqI: haftmann@41080: "(THE x. eval P x) = x \ the P = x" haftmann@40674: by (simp add: the_def) haftmann@40674: Andreas@53943: lemma the_eq [code]: "the A = singleton (\x. Code.abort (STR ''not_unique'') (\_. the A)) A" Andreas@53943: by (rule the_eqI) (simp add: singleton_def the_def) bulwahn@33110: haftmann@36531: code_reflect Predicate haftmann@36513: datatypes pred = Seq and seq = Empty | Insert | Join haftmann@36513: haftmann@30948: ML {* haftmann@30948: signature PREDICATE = haftmann@30948: sig haftmann@51126: val anamorph: ('a -> ('b * 'a) option) -> int -> 'a -> 'b list * 'a 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@51126: val map: ('a -> 'b) -> 'a pred -> 'b pred haftmann@30959: val yield: 'a pred -> ('a * 'a pred) option haftmann@30959: val yieldn: int -> 'a pred -> 'a list * 'a pred haftmann@30948: end; haftmann@30948: haftmann@30948: structure Predicate : PREDICATE = haftmann@30948: struct haftmann@30948: haftmann@51126: fun anamorph f k x = haftmann@51126: (if k = 0 then ([], x) haftmann@51126: else case f x haftmann@51126: of NONE => ([], x) haftmann@51126: | SOME (v, y) => let haftmann@51126: val k' = k - 1; haftmann@51126: val (vs, z) = anamorph f k' y haftmann@51126: in (v :: vs, z) end); haftmann@51126: haftmann@36513: datatype pred = datatype Predicate.pred haftmann@36513: datatype seq = datatype Predicate.seq haftmann@36513: haftmann@51126: fun map f = @{code Predicate.map} f; haftmann@30959: haftmann@36513: fun yield (Seq f) = next (f ()) haftmann@36513: and next Empty = NONE haftmann@36513: | next (Insert (x, P)) = SOME (x, P) haftmann@36513: | next (Join (P, xq)) = (case yield P haftmann@30959: of NONE => next xq haftmann@36513: | SOME (x, Q) => SOME (x, Seq (fn _ => Join (Q, xq)))); haftmann@30959: haftmann@51126: fun yieldn k = anamorph yield k; haftmann@30948: haftmann@30948: end; haftmann@30948: *} haftmann@30948: haftmann@46038: text {* Conversion from and to sets *} haftmann@46038: haftmann@46038: definition pred_of_set :: "'a set \ 'a pred" where haftmann@46038: "pred_of_set = Pred \ (\A x. x \ A)" haftmann@46038: haftmann@46038: lemma eval_pred_of_set [simp]: haftmann@46038: "eval (pred_of_set A) x \ x \A" haftmann@46038: by (simp add: pred_of_set_def) haftmann@46038: haftmann@46038: definition set_of_pred :: "'a pred \ 'a set" where haftmann@46038: "set_of_pred = Collect \ eval" haftmann@46038: haftmann@46038: lemma member_set_of_pred [simp]: haftmann@46038: "x \ set_of_pred P \ Predicate.eval P x" haftmann@46038: by (simp add: set_of_pred_def) haftmann@46038: haftmann@46038: definition set_of_seq :: "'a seq \ 'a set" where haftmann@46038: "set_of_seq = set_of_pred \ pred_of_seq" haftmann@46038: haftmann@46038: lemma member_set_of_seq [simp]: haftmann@46038: "x \ set_of_seq xq = Predicate.member xq x" haftmann@46038: by (simp add: set_of_seq_def eval_member) haftmann@46038: haftmann@46038: lemma of_pred_code [code]: haftmann@46038: "set_of_pred (Predicate.Seq f) = (case f () of haftmann@46038: Predicate.Empty \ {} haftmann@46038: | Predicate.Insert x P \ insert x (set_of_pred P) haftmann@46038: | Predicate.Join P xq \ set_of_pred P \ set_of_seq xq)" haftmann@46038: by (auto split: seq.split simp add: eval_code) haftmann@46038: haftmann@46038: lemma of_seq_code [code]: haftmann@46038: "set_of_seq Predicate.Empty = {}" haftmann@46038: "set_of_seq (Predicate.Insert x P) = insert x (set_of_pred P)" haftmann@46038: "set_of_seq (Predicate.Join P xq) = set_of_pred P \ set_of_seq xq" haftmann@46038: by auto haftmann@46038: haftmann@46664: text {* Lazy Evaluation of an indexed function *} haftmann@46664: haftmann@51143: function iterate_upto :: "(natural \ 'a) \ natural \ natural \ 'a Predicate.pred" haftmann@46664: where haftmann@46664: "iterate_upto f n m = haftmann@46664: Predicate.Seq (%u. if n > m then Predicate.Empty haftmann@46664: else Predicate.Insert (f n) (iterate_upto f (n + 1) m))" haftmann@46664: by pat_completeness auto haftmann@46664: haftmann@51143: termination by (relation "measure (%(f, n, m). nat_of_natural (m + 1 - n))") haftmann@51143: (auto simp add: less_natural_def) haftmann@46664: haftmann@46664: text {* Misc *} haftmann@46664: haftmann@47399: declare Inf_set_fold [where 'a = "'a Predicate.pred", code] haftmann@47399: declare Sup_set_fold [where 'a = "'a Predicate.pred", code] haftmann@46664: haftmann@46664: (* FIXME: better implement conversion by bisection *) haftmann@46664: haftmann@46664: lemma pred_of_set_fold_sup: haftmann@46664: assumes "finite A" haftmann@46664: shows "pred_of_set A = Finite_Set.fold sup bot (Predicate.single ` A)" (is "?lhs = ?rhs") haftmann@46664: proof (rule sym) haftmann@46664: interpret comp_fun_idem "sup :: 'a Predicate.pred \ 'a Predicate.pred \ 'a Predicate.pred" haftmann@46664: by (fact comp_fun_idem_sup) haftmann@46664: from `finite A` show "?rhs = ?lhs" by (induct A) (auto intro!: pred_eqI) haftmann@46664: qed haftmann@46664: haftmann@46664: lemma pred_of_set_set_fold_sup: haftmann@46664: "pred_of_set (set xs) = fold sup (List.map Predicate.single xs) bot" haftmann@46664: proof - haftmann@46664: interpret comp_fun_idem "sup :: 'a Predicate.pred \ 'a Predicate.pred \ 'a Predicate.pred" haftmann@46664: by (fact comp_fun_idem_sup) haftmann@46664: show ?thesis by (simp add: pred_of_set_fold_sup fold_set_fold [symmetric]) haftmann@46664: qed haftmann@46664: haftmann@46664: lemma pred_of_set_set_foldr_sup [code]: haftmann@46664: "pred_of_set (set xs) = foldr sup (List.map Predicate.single xs) bot" haftmann@46664: by (simp add: pred_of_set_set_fold_sup ac_simps foldr_fold fun_eq_iff) haftmann@46664: haftmann@30328: no_notation haftmann@30328: bind (infixl "\=" 70) haftmann@30328: wenzelm@36176: hide_type (open) pred seq wenzelm@36176: hide_const (open) Pred eval single bind is_empty singleton if_pred not_pred holds Andreas@53943: Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map the haftmann@46664: iterate_upto haftmann@46664: hide_fact (open) null_def member_def haftmann@30328: haftmann@30328: end haftmann@46664: