# HG changeset patch # User krauss # Date 1306833077 -7200 # Node ID 1a39c9898ae6f6189dd40faec6bbd5bda29ff5c8 # Parent 73a1d6a7ef1de05f8b6fd55acf160cff40cf62f6 admissibility on option type diff -r 73a1d6a7ef1d -r 1a39c9898ae6 src/HOL/Partial_Function.thy --- a/src/HOL/Partial_Function.thy Mon May 23 21:34:37 2011 +0200 +++ b/src/HOL/Partial_Function.thy Tue May 31 11:11:17 2011 +0200 @@ -23,6 +23,17 @@ definition "img_ord f ord = (\x y. ord (f x) (f y))" definition "img_lub f g Lub = (\A. g (Lub (f ` A)))" +lemma chain_fun: + assumes A: "chain (fun_ord ord) A" + shows "chain ord {y. \f\A. y = f a}" (is "chain ord ?C") +proof (rule chainI) + fix x y assume "x \ ?C" "y \ ?C" + then obtain f g where fg: "f \ A" "g \ A" + and [simp]: "x = f a" "y = g a" by blast + from chainD[OF A fg] + show "ord x y \ ord y x" unfolding fun_ord_def by auto +qed + lemma call_mono[partial_function_mono]: "monotone (fun_ord ord) ord (\f. f t)" by (rule monotoneI) (auto simp: fun_ord_def) @@ -53,17 +64,6 @@ proof - interpret partial_function_definitions ord lb by fact - { fix A a assume A: "chain ?ordf A" - have "chain ord {y. \f\A. y = f a}" (is "chain ord ?C") - proof (rule chainI) - fix x y assume "x \ ?C" "y \ ?C" - then obtain f g where fg: "f \ A" "g \ A" - and [simp]: "x = f a" "y = g a" by blast - from chainD[OF A fg] - show "ord x y \ ord y x" unfolding fun_ord_def by auto - qed } - note chain_fun = this - show ?thesis proof fix x show "?ordf x x" @@ -75,7 +75,7 @@ next fix x y assume "?ordf x y" "?ordf y x" thus "x = y" unfolding fun_ord_def - by (force intro!: ext dest: leq_antisym) + by (force intro!: dest: leq_antisym) next fix A f assume f: "f \ A" and A: "chain ?ordf A" thus "?ordf f (?lubf A)" @@ -129,6 +129,7 @@ abbreviation "lub_fun \ fun_lub lub" abbreviation "fixp_fun == ccpo.fixp le_fun lub_fun" abbreviation "mono_body \ monotone le_fun leq" +abbreviation "admissible \ ccpo.admissible le_fun lub_fun" text {* Interpret manually, to avoid flooding everything with facts about orders *} @@ -251,6 +252,43 @@ show "option_ord (Option.bind (B f) (\y. C y f)) (Option.bind (B g) (\y'. C y' g))" . qed +lemma flat_lub_in_chain: + assumes ch: "chain (flat_ord b) A " + assumes lub: "flat_lub b A = a" + shows "a = b \ a \ A" +proof (cases "A \ {b}") + case True + then have "flat_lub b A = b" unfolding flat_lub_def by simp + with lub show ?thesis by simp +next + case False + then obtain c where "c \ A" and "c \ b" by auto + { fix z assume "z \ A" + from chainD[OF ch `c \ A` this] have "z = c \ z = b" + unfolding flat_ord_def using `c \ b` by auto } + with False have "A - {b} = {c}" by auto + with False have "flat_lub b A = c" by (auto simp: flat_lub_def) + with `c \ A` lub show ?thesis by simp +qed + +lemma option_admissible: "option.admissible (%(f::'a \ 'b option). + (\x y. f x = Some y \ P x y))" +proof (rule ccpo.admissibleI[OF option.ccpo]) + fix A :: "('a \ 'b option) set" + assume ch: "chain option.le_fun A" + and IH: "\f\A. \x y. f x = Some y \ P x y" + from ch have ch': "\x. chain option_ord {y. \f\A. y = f x}" by (rule chain_fun) + show "\x y. option.lub_fun A x = Some y \ P x y" + proof (intro allI impI) + fix x y assume "option.lub_fun A x = Some y" + from flat_lub_in_chain[OF ch' this[unfolded fun_lub_def]] + have "Some y \ {y. \f\A. y = f x}" by simp + then have "\f\A. f x = Some y" by auto + with IH show "P x y" by auto + qed +qed + + hide_const (open) chain end