# HG changeset patch # User huffman # Date 1214857007 -7200 # Node ID 253a06dfadcec43d22ae8a5ba16773b8471c37e9 # Parent 4edc81f93e35bb0929eef9485fe7d0efa2e353e1 reuse proofs from Deflation.thy; clean up proof of finite_range_cfun_lemma diff -r 4edc81f93e35 -r 253a06dfadce src/HOLCF/Bifinite.thy --- a/src/HOLCF/Bifinite.thy Mon Jun 30 21:52:17 2008 +0200 +++ b/src/HOLCF/Bifinite.thy Mon Jun 30 22:16:47 2008 +0200 @@ -6,7 +6,7 @@ header {* Bifinite domains and approximation *} theory Bifinite -imports Cfun +imports Deflation begin subsection {* Omega-profinite and bifinite domains *} @@ -20,40 +20,45 @@ class bifinite = profinite + pcpo -lemma finite_range_imp_finite_fixes: - "finite {x. \y. x = f y} \ finite {x. f x = x}" -apply (subgoal_tac "{x. f x = x} \ {x. \y. x = f y}") -apply (erule (1) finite_subset) -apply (clarify, erule subst, rule exI, rule refl) -done +lemma approx_less: "approx i\x \ x" +proof - + have "chain (\i. approx i\x)" by simp + hence "approx i\x \ (\i. approx i\x)" by (rule is_ub_thelub) + thus "approx i\x \ x" by simp +qed + +lemma finite_deflation_approx: "finite_deflation (approx i)" +proof + fix x :: 'a + show "approx i\(approx i\x) = approx i\x" + by (rule approx_idem) + show "approx i\x \ x" + by (rule approx_less) + show "finite {x. approx i\x = x}" + by (rule finite_fixes_approx) +qed + +interpretation approx: finite_deflation ["approx i"] +by (rule finite_deflation_approx) + +lemma deflation_approx: "deflation (approx i)" +by (rule approx.deflation_axioms) lemma lub_approx [simp]: "(\i. approx i) = (\ x. x)" by (rule ext_cfun, simp add: contlub_cfun_fun) -lemma approx_less: "approx i\x \ x" -apply (subgoal_tac "approx i\x \ (\i. approx i\x)", simp) -apply (rule is_ub_thelub, simp) -done - lemma approx_strict [simp]: "approx i\\ = \" by (rule UU_I, rule approx_less) lemma approx_approx1: "i \ j \ approx i\(approx j\x) = approx i\x" -apply (rule antisym_less) -apply (rule monofun_cfun_arg [OF approx_less]) -apply (rule sq_ord_eq_less_trans [OF approx_idem [symmetric]]) -apply (rule monofun_cfun_arg) -apply (rule monofun_cfun_fun) +apply (rule deflation_less_comp1 [OF deflation_approx deflation_approx]) apply (erule chain_mono [OF chain_approx]) done lemma approx_approx2: "j \ i \ approx i\(approx j\x) = approx j\x" -apply (rule antisym_less) -apply (rule approx_less) -apply (rule sq_ord_eq_less_trans [OF approx_idem [symmetric]]) -apply (rule monofun_cfun_fun) +apply (rule deflation_less_comp2 [OF deflation_approx deflation_approx]) apply (erule chain_mono [OF chain_approx]) done @@ -64,50 +69,17 @@ apply (simp add: approx_approx2 min_def) done -lemma idem_fixes_eq_range: - "\x. f (f x) = f x \ {x. f x = x} = {y. \x. y = f x}" -by (auto simp add: eq_sym_conv) +lemma finite_image_approx: "finite ((\x. approx n\x) ` A)" +by (rule approx.finite_image) -lemma finite_approx: "finite {y. \x. y = approx n\x}" -using finite_fixes_approx by (simp add: idem_fixes_eq_range) - -lemma finite_image_approx: "finite ((\x. approx n\x) ` A)" -by (rule finite_subset [OF _ finite_fixes_approx [where i=n]]) auto - -lemma finite_range_approx: "finite (range (\x. approx n\x))" -by (rule finite_image_approx) +lemma finite_range_approx: "finite (range (\x. approx i\x))" +by (rule approx.finite_range) lemma compact_approx [simp]: "compact (approx n\x)" -proof (rule compactI2) - fix Y::"nat \ 'a" - assume Y: "chain Y" - have "finite_chain (\i. approx n\(Y i))" - proof (rule finite_range_imp_finch) - show "chain (\i. approx n\(Y i))" - using Y by simp - have "range (\i. approx n\(Y i)) \ {x. approx n\x = x}" - by clarsimp - thus "finite (range (\i. approx n\(Y i)))" - using finite_fixes_approx by (rule finite_subset) - qed - hence "\j. (\i. approx n\(Y i)) = approx n\(Y j)" - by (simp add: finite_chain_def maxinch_is_thelub Y) - then obtain j where j: "(\i. approx n\(Y i)) = approx n\(Y j)" .. - - assume "approx n\x \ (\i. Y i)" - hence "approx n\(approx n\x) \ approx n\(\i. Y i)" - by (rule monofun_cfun_arg) - hence "approx n\x \ (\i. approx n\(Y i))" - by (simp add: contlub_cfun_arg Y) - hence "approx n\x \ approx n\(Y j)" - using j by simp - hence "approx n\x \ Y j" - using approx_less by (rule trans_less) - thus "\j. approx n\x \ Y j" .. -qed +by (rule approx.compact) lemma profinite_compact_eq_approx: "compact x \ \i. approx i\x = x" -by (rule admD2) simp_all +by (rule admD2, simp_all) lemma profinite_compact_iff: "compact x \ (\n. approx n\x = x)" apply (rule iffI) @@ -133,25 +105,34 @@ subsection {* Instance for continuous function space *} -lemma finite_range_lemma: - fixes h :: "'a::cpo \ 'b::cpo" - fixes k :: "'c::cpo \ 'd::cpo" - shows "\finite {y. \x. y = h\x}; finite {y. \x. y = k\x}\ - \ finite {g. \f. g = (\ x. k\(f\(h\x)))}" - apply (rule_tac f="\g. {(h\x, y) |x y. y = g\x}" in finite_imageD) - apply (rule_tac B="Pow ({y. \x. y = h\x} \ {y. \x. y = k\x})" - in finite_subset) - apply (rule image_subsetI) - apply (clarsimp, fast) - apply simp - apply (rule inj_onI) - apply (clarsimp simp add: expand_set_eq) - apply (rule ext_cfun, simp) - apply (drule_tac x="h\x" in spec) - apply (drule_tac x="k\(f\(h\x))" in spec) - apply (drule iffD1, fast) - apply clarsimp -done +lemma finite_range_cfun_lemma: + assumes a: "finite (range (\x. a\x))" + assumes b: "finite (range (\y. b\y))" + shows "finite (range (\f. \ x. b\(f\(a\x))))" (is "finite (range ?h)") +proof (rule finite_imageD) + let ?f = "\g. range (\x. (a\x, g\x))" + show "finite (?f ` range ?h)" + proof (rule finite_subset) + let ?B = "Pow (range (\x. a\x) \ range (\y. b\y))" + show "?f ` range ?h \ ?B" + by clarsimp + show "finite ?B" + by (simp add: a b) + qed + show "inj_on ?f (range ?h)" + proof (rule inj_onI, rule ext_cfun, clarsimp) + fix x f g + assume "range (\x. (a\x, b\(f\(a\x)))) = range (\x. (a\x, b\(g\(a\x))))" + hence "range (\x. (a\x, b\(f\(a\x)))) \ range (\x. (a\x, b\(g\(a\x))))" + by (rule equalityD1) + hence "(a\x, b\(f\(a\x))) \ range (\x. (a\x, b\(g\(a\x))))" + by (simp add: subset_eq) + then obtain y where "(a\x, b\(f\(a\x))) = (a\y, b\(g\(a\y)))" + by (rule rangeE) + thus "b\(f\(a\x)) = b\(g\(a\x))" + by clarsimp + qed +qed instantiation "->" :: (profinite, profinite) profinite begin @@ -160,15 +141,26 @@ approx_cfun_def: "approx = (\n. \ f x. approx n\(f\(approx n\x)))" -instance - apply (intro_classes, unfold approx_cfun_def) - apply simp - apply (simp add: lub_distribs eta_cfun) - apply simp - apply simp - apply (rule finite_range_imp_finite_fixes) - apply (intro finite_range_lemma finite_approx) -done +instance proof + show "chain (approx :: nat \ ('a \ 'b) \ ('a \ 'b))" + unfolding approx_cfun_def by simp +next + fix x :: "'a \ 'b" + show "(\i. approx i\x) = x" + unfolding approx_cfun_def + by (simp add: lub_distribs eta_cfun) +next + fix i :: nat and x :: "'a \ 'b" + show "approx i\(approx i\x) = approx i\x" + unfolding approx_cfun_def by simp +next + fix i :: nat + show "finite {x::'a \ 'b. approx i\x = x}" + apply (rule finite_range_imp_finite_fixes) + apply (simp add: approx_cfun_def) + apply (intro finite_range_cfun_lemma finite_range_approx) + done +qed end