# HG changeset patch # User Andreas Lochbihler # Date 1380277599 -7200 # Node ID ea38710e731c15e3792567d16666ea0169890364 # Parent 021a8ef97f5f50c4f28343ff67f61b55dcffebeb# Parent 2a4c156e6d367f71d78b002f460f8f4823290a5d merged diff -r 2a4c156e6d36 -r ea38710e731c src/HOL/Partial_Function.thy --- a/src/HOL/Partial_Function.thy Fri Sep 27 11:56:52 2013 +0200 +++ b/src/HOL/Partial_Function.thy Fri Sep 27 12:26:39 2013 +0200 @@ -248,21 +248,21 @@ abbreviation "mono_tailrec \ monotone (fun_ord tailrec_ord) tailrec_ord" lemma tailrec_admissible: - "ccpo.admissible (fun_lub (flat_lub undefined)) (fun_ord tailrec_ord) - (\a. \x. a x \ undefined \ P x (a x))" + "ccpo.admissible (fun_lub (flat_lub c)) (fun_ord (flat_ord c)) + (\a. \x. a x \ c \ P x (a x))" proof(intro ccpo.admissibleI strip) fix A x - assume chain: "Complete_Partial_Order.chain (fun_ord tailrec_ord) A" - and P [rule_format]: "\f\A. \x. f x \ undefined \ P x (f x)" - and defined: "fun_lub (flat_lub undefined) A x \ undefined" - from defined obtain f where f: "f \ A" "f x \ undefined" + assume chain: "Complete_Partial_Order.chain (fun_ord (flat_ord c)) A" + and P [rule_format]: "\f\A. \x. f x \ c \ P x (f x)" + and defined: "fun_lub (flat_lub c) A x \ c" + from defined obtain f where f: "f \ A" "f x \ c" by(auto simp add: fun_lub_def flat_lub_def split: split_if_asm) hence "P x (f x)" by(rule P) - moreover from chain f have "\f' \ A. f' x = undefined \ f' x = f x" + moreover from chain f have "\f' \ A. f' x = c \ f' x = f x" by(auto 4 4 simp add: Complete_Partial_Order.chain_def flat_ord_def fun_ord_def) - hence "fun_lub (flat_lub undefined) A x = f x" + hence "fun_lub (flat_lub c) A x = f x" using f by(auto simp add: fun_lub_def flat_lub_def) - ultimately show "P x (fun_lub (flat_lub undefined) A x)" by simp + ultimately show "P x (fun_lub (flat_lub c) A x)" by simp qed lemma fixp_induct_tailrec: @@ -271,16 +271,17 @@ C :: "('b \ 'a) \ 'c" and P :: "'b \ 'a \ bool" and x :: "'b" - assumes mono: "\x. mono_tailrec (\f. U (F (C f)) x)" - assumes eq: "f \ C (ccpo.fixp (fun_lub (flat_lub undefined)) (fun_ord tailrec_ord) (\f. U (F (C f))))" + assumes mono: "\x. monotone (fun_ord (flat_ord c)) (flat_ord c) (\f. U (F (C f)) x)" + assumes eq: "f \ C (ccpo.fixp (fun_lub (flat_lub c)) (fun_ord (flat_ord c)) (\f. U (F (C f))))" assumes inverse2: "\f. U (C f) = f" - assumes step: "\f x y. (\x y. U f x = y \ y \ undefined \ P x y) \ U (F f) x = y \ y \ undefined \ P x y" + assumes step: "\f x y. (\x y. U f x = y \ y \ c \ P x y) \ U (F f) x = y \ y \ c \ P x y" assumes result: "U f x = y" - assumes defined: "y \ undefined" + assumes defined: "y \ c" shows "P x y" proof - - have "\x y. U f x = y \ y \ undefined \ P x y" - by(rule tailrec.fixp_induct_uc[of U F C, OF mono eq inverse2])(auto intro: step tailrec_admissible) + have "\x y. U f x = y \ y \ c \ P x y" + by(rule partial_function_definitions.fixp_induct_uc[OF flat_interpretation, of _ U F C, OF mono eq inverse2]) + (auto intro: step tailrec_admissible) thus ?thesis using result defined by blast qed @@ -391,7 +392,7 @@ declaration {* Partial_Function.init "tailrec" @{term tailrec.fixp_fun} @{term tailrec.mono_body} @{thm tailrec.fixp_rule_uc} @{thm tailrec.fixp_induct_uc} - (SOME @{thm fixp_induct_tailrec}) *} + (SOME @{thm fixp_induct_tailrec[where c=undefined]}) *} declaration {* Partial_Function.init "option" @{term option.fixp_fun} @{term option.mono_body} @{thm option.fixp_rule_uc} @{thm option.fixp_induct_uc}