# HG changeset patch # User huffman # Date 1213908658 -7200 # Node ID c49d427867aab967fd5b6c7069716a115223d576 # Parent 274b806912598c346ee0958882988156b1ed462c move lemmas into locales; restructure some proofs diff -r 274b80691259 -r c49d427867aa src/HOLCF/CompactBasis.thy --- a/src/HOLCF/CompactBasis.thy Thu Jun 19 22:43:59 2008 +0200 +++ b/src/HOLCF/CompactBasis.thy Thu Jun 19 22:50:58 2008 +0200 @@ -240,7 +240,7 @@ lemma principal_less_iff_mem_rep: "principal a \ x \ a \ rep x" by (simp add: rep_eq) -lemma principal_less_iff: "principal a \ principal b \ a \ b" +lemma principal_less_iff [simp]: "principal a \ principal b \ a \ b" by (simp add: principal_less_iff_mem_rep rep_principal) lemma principal_eq_iff: "principal a = principal b \ a \ b \ b \ a" @@ -250,7 +250,7 @@ by (simp add: rep_eq) lemma principal_mono: "a \ b \ principal a \ principal b" -by (simp add: principal_less_iff) +by (simp only: principal_less_iff) lemma lessI: "(\a. principal a \ x \ principal a \ u) \ x \ u" unfolding principal_less_iff_mem_rep @@ -316,7 +316,7 @@ apply (erule imageI) done -lemma compact_principal: "compact (principal a)" +lemma compact_principal [simp]: "compact (principal a)" by (rule compactI2, simp add: principal_less_iff_mem_rep rep_contlub) definition @@ -392,6 +392,16 @@ apply (clarify, simp add: P) done +lemma principal_induct2: + "\\y. adm (\x. P x y); \x. adm (\y. P x y); + \a b. P (principal a) (principal b)\ \ P x y" +apply (rule_tac x=y in spec) +apply (rule_tac x=x in principal_induct, simp) +apply (rule allI, rename_tac y) +apply (rule_tac x=y in principal_induct, simp) +apply simp +done + lemma compact_imp_principal: "compact x \ \a. x = principal a" apply (drule adm_compact_neq [OF _ cont_id]) apply (drule admD2 [where Y="\n. completion_approx n\x"]) @@ -412,17 +422,9 @@ typedef (open) 'a compact_basis = "{x::'a::profinite. compact x}" by (fast intro: compact_approx) -lemma compact_Rep_compact_basis [simp]: "compact (Rep_compact_basis a)" +lemma compact_Rep_compact_basis: "compact (Rep_compact_basis a)" by (rule Rep_compact_basis [unfolded mem_Collect_eq]) -lemma Rep_Abs_compact_basis_approx [simp]: - "Rep_compact_basis (Abs_compact_basis (approx n\x)) = approx n\x" -by (rule Abs_compact_basis_inverse, simp) - -lemma compact_imp_Rep_compact_basis: - "compact x \ \y. x = Rep_compact_basis y" -by (rule exI, rule Abs_compact_basis_inverse [symmetric], simp) - instantiation compact_basis :: (profinite) sq_ord begin @@ -438,81 +440,7 @@ by (rule typedef_po [OF type_definition_compact_basis compact_le_def]) -text {* minimal compact element *} - -definition - compact_bot :: "'a::bifinite compact_basis" where - "compact_bot = Abs_compact_basis \" - -lemma Rep_compact_bot: "Rep_compact_basis compact_bot = \" -unfolding compact_bot_def by (simp add: Abs_compact_basis_inverse) - -lemma compact_minimal [simp]: "compact_bot \ a" -unfolding compact_le_def Rep_compact_bot by simp - -text {* compacts *} - -definition - compacts :: "'a \ 'a compact_basis set" where - "compacts = (\x. {a. Rep_compact_basis a \ x})" - -lemma ideal_compacts: "sq_le.ideal (compacts w)" -unfolding compacts_def - apply (rule sq_le.idealI) - apply (rule_tac x="Abs_compact_basis (approx 0\w)" in exI) - apply (simp add: approx_less) - apply simp - apply (cut_tac a=x in compact_Rep_compact_basis) - apply (cut_tac a=y in compact_Rep_compact_basis) - apply (drule bifinite_compact_eq_approx) - apply (drule bifinite_compact_eq_approx) - apply (clarify, rename_tac i j) - apply (rule_tac x="Abs_compact_basis (approx (max i j)\w)" in exI) - apply (simp add: approx_less compact_le_def) - apply (erule subst, erule subst) - apply (simp add: monofun_cfun chain_mono [OF chain_approx]) - apply (simp add: compact_le_def) - apply (erule (1) trans_less) -done - -lemma compacts_Rep_compact_basis: - "compacts (Rep_compact_basis b) = {a. a \ b}" -unfolding compacts_def compact_le_def .. - -lemma cont_compacts: "cont compacts" -unfolding compacts_def -apply (rule contI2) -apply (rule monofunI) -apply (simp add: set_cpo_simps) -apply (fast intro: trans_less) -apply (simp add: set_cpo_simps) -apply clarify -apply simp -apply (erule (1) compactD2 [OF compact_Rep_compact_basis]) -done - -lemma compacts_lessD: "compacts x \ compacts y \ x \ y" -apply (subgoal_tac "(\i. approx i\x) \ y", simp) -apply (rule admD, simp, simp) -apply (drule_tac c="Abs_compact_basis (approx i\x)" in subsetD) -apply (simp add: compacts_def Abs_compact_basis_inverse approx_less) -apply (simp add: compacts_def Abs_compact_basis_inverse) -done - -lemma compacts_mono: "x \ y \ compacts x \ compacts y" -unfolding compacts_def by (fast intro: trans_less) - -lemma less_compact_basis_iff: "(x \ y) = (compacts x \ compacts y)" -by (rule iffI [OF compacts_mono compacts_lessD]) - -lemma compact_basis_induct: - "\adm P; \a. P (Rep_compact_basis a)\ \ P x" -apply (erule approx_induct) -apply (drule_tac x="Abs_compact_basis (approx n\x)" in meta_spec) -apply (simp add: Abs_compact_basis_inverse) -done - -text {* approximation on compact bases *} +text {* Take function for compact basis *} definition compact_approx :: "nat \ 'a compact_basis \ 'a compact_basis" where @@ -525,84 +453,128 @@ lemmas approx_Rep_compact_basis = Rep_compact_approx [symmetric] -lemma compact_approx_le: "compact_approx n a \ a" -unfolding compact_le_def -by (simp add: Rep_compact_approx approx_less) - -lemma compact_approx_mono1: - "i \ j \ compact_approx i a \ compact_approx j a" -unfolding compact_le_def -apply (simp add: Rep_compact_approx) -apply (rule chain_mono, simp, assumption) -done - -lemma compact_approx_mono: - "a \ b \ compact_approx n a \ compact_approx n b" -unfolding compact_le_def -apply (simp add: Rep_compact_approx) -apply (erule monofun_cfun_arg) -done +interpretation compact_basis: + basis_take [sq_le compact_approx] +proof + fix n :: nat and a :: "'a compact_basis" + show "compact_approx n a \ a" + unfolding compact_le_def + by (simp add: Rep_compact_approx approx_less) +next + fix n :: nat and a :: "'a compact_basis" + show "compact_approx n (compact_approx n a) = compact_approx n a" + by (simp add: Rep_compact_basis_inject [symmetric] Rep_compact_approx) +next + fix n :: nat and a b :: "'a compact_basis" + assume "a \ b" thus "compact_approx n a \ compact_approx n b" + unfolding compact_le_def Rep_compact_approx + by (rule monofun_cfun_arg) +next + fix n :: nat and a :: "'a compact_basis" + show "\n a. compact_approx n a \ compact_approx (Suc n) a" + unfolding compact_le_def Rep_compact_approx + by (rule chainE, simp) +next + fix n :: nat + show "finite (range (compact_approx n))" + apply (rule finite_imageD [where f="Rep_compact_basis"]) + apply (rule finite_subset [where B="range (\x. approx n\x)"]) + apply (clarsimp simp add: Rep_compact_approx) + apply (rule finite_range_approx) + apply (rule inj_onI, simp add: Rep_compact_basis_inject) + done +next + fix a :: "'a compact_basis" + show "\n. compact_approx n a = a" + apply (simp add: Rep_compact_basis_inject [symmetric]) + apply (simp add: Rep_compact_approx) + apply (rule bifinite_compact_eq_approx) + apply (rule compact_Rep_compact_basis) + done +qed -lemma ex_compact_approx_eq: "\n. compact_approx n a = a" -apply (simp add: Rep_compact_basis_inject [symmetric]) -apply (simp add: Rep_compact_approx) -apply (rule bifinite_compact_eq_approx) -apply (rule compact_Rep_compact_basis) -done - -lemma compact_approx_idem: - "compact_approx n (compact_approx n a) = compact_approx n a" -apply (rule Rep_compact_basis_inject [THEN iffD1]) -apply (simp add: Rep_compact_approx) -done +text {* Ideal completion *} -lemma finite_fixes_compact_approx: "finite {a. compact_approx n a = a}" -apply (subgoal_tac "finite (Rep_compact_basis ` {a. compact_approx n a = a})") -apply (erule finite_imageD) -apply (rule inj_onI, simp add: Rep_compact_basis_inject) -apply (rule finite_subset [OF _ finite_fixes_approx [where i=n]]) -apply (rule subsetI, clarify, rename_tac a) -apply (simp add: Rep_compact_basis_inject [symmetric]) -apply (simp add: Rep_compact_approx) -done - -lemma finite_range_compact_approx: "finite (range (compact_approx n))" -apply (cut_tac n=n in finite_fixes_compact_approx) -apply (simp add: idem_fixes_eq_range compact_approx_idem) -apply (simp add: image_def) -done +definition + compacts :: "'a \ 'a compact_basis set" where + "compacts = (\x. {a. Rep_compact_basis a \ x})" interpretation compact_basis: ideal_completion [sq_le compact_approx Rep_compact_basis compacts] -proof (unfold_locales) - fix n :: nat and a b :: "'a compact_basis" and x :: "'a" - show "compact_approx n a \ a" - by (rule compact_approx_le) - show "compact_approx n (compact_approx n a) = compact_approx n a" - by (rule compact_approx_idem) - show "compact_approx n a \ compact_approx (Suc n) a" - by (rule compact_approx_mono1, simp) - show "finite (range (compact_approx n))" - by (rule finite_range_compact_approx) - show "\n\nat. compact_approx n a = a" - by (rule ex_compact_approx_eq) - show "preorder.ideal sq_le (compacts x)" - by (rule ideal_compacts) +proof + fix w :: 'a + show "preorder.ideal sq_le (compacts w)" + proof (rule sq_le.idealI) + show "\x. x \ compacts w" + unfolding compacts_def + apply (rule_tac x="Abs_compact_basis (approx 0\w)" in exI) + apply (simp add: Abs_compact_basis_inverse approx_less) + done + next + fix x y :: "'a compact_basis" + assume "x \ compacts w" "y \ compacts w" + thus "\z \ compacts w. x \ z \ y \ z" + unfolding compacts_def + apply simp + apply (cut_tac a=x in compact_Rep_compact_basis) + apply (cut_tac a=y in compact_Rep_compact_basis) + apply (drule bifinite_compact_eq_approx) + apply (drule bifinite_compact_eq_approx) + apply (clarify, rename_tac i j) + apply (rule_tac x="Abs_compact_basis (approx (max i j)\w)" in exI) + apply (simp add: compact_le_def) + apply (simp add: Abs_compact_basis_inverse approx_less) + apply (erule subst, erule subst) + apply (simp add: monofun_cfun chain_mono [OF chain_approx]) + done + next + fix x y :: "'a compact_basis" + assume "x \ y" "y \ compacts w" thus "x \ compacts w" + unfolding compacts_def + apply simp + apply (simp add: compact_le_def) + apply (erule (1) trans_less) + done + qed +next show "cont compacts" - by (rule cont_compacts) - show "compacts (Rep_compact_basis a) = {b. b \ a}" - by (rule compacts_Rep_compact_basis) + unfolding compacts_def + apply (rule contI2) + apply (rule monofunI) + apply (simp add: set_cpo_simps) + apply (fast intro: trans_less) + apply (simp add: set_cpo_simps) + apply clarify + apply simp + apply (erule (1) compactD2 [OF compact_Rep_compact_basis]) + done next - fix n :: nat and a b :: "'a compact_basis" - assume "a \ b" - thus "compact_approx n a \ compact_approx n b" - by (rule compact_approx_mono) + fix a :: "'a compact_basis" + show "compacts (Rep_compact_basis a) = {b. b \ a}" + unfolding compacts_def compact_le_def .. next fix x y :: "'a" assume "compacts x \ compacts y" thus "x \ y" - by (rule compacts_lessD) + apply (subgoal_tac "(\i. approx i\x) \ y", simp) + apply (rule admD, simp, simp) + apply (drule_tac c="Abs_compact_basis (approx i\x)" in subsetD) + apply (simp add: compacts_def Abs_compact_basis_inverse approx_less) + apply (simp add: compacts_def Abs_compact_basis_inverse) + done qed +text {* minimal compact element *} + +definition + compact_bot :: "'a::bifinite compact_basis" where + "compact_bot = Abs_compact_basis \" + +lemma Rep_compact_bot: "Rep_compact_basis compact_bot = \" +unfolding compact_bot_def by (simp add: Abs_compact_basis_inverse) + +lemma compact_minimal [simp]: "compact_bot \ a" +unfolding compact_le_def Rep_compact_bot by simp + subsection {* A compact basis for powerdomains *} @@ -705,37 +677,27 @@ lemma approx_pd_idem: "approx_pd n (approx_pd n t) = approx_pd n t" apply (induct t rule: pd_basis_induct) -apply (simp add: compact_approx_idem) +apply (simp add: compact_basis.take_take) apply simp done -lemma range_image_f: "range (image f) = Pow (range f)" -apply (safe, fast) -apply (rule_tac x="f -` x" in range_eqI) -apply (simp, fast) -done - lemma finite_range_approx_pd: "finite (range (approx_pd n))" -apply (subgoal_tac "finite (Rep_pd_basis ` range (approx_pd n))") -apply (erule finite_imageD) +apply (rule finite_imageD [where f="Rep_pd_basis"]) +apply (rule finite_subset [where B="Pow (range (compact_approx n))"]) +apply (clarsimp simp add: Rep_approx_pd) +apply (simp add: compact_basis.finite_range_take) apply (rule inj_onI, simp add: Rep_pd_basis_inject) -apply (subst image_image) -apply (subst Rep_approx_pd) -apply (simp only: range_composition) -apply (rule finite_subset [OF image_mono [OF subset_UNIV]]) -apply (simp add: range_image_f) -apply (rule finite_range_compact_approx) done -lemma ex_approx_pd_eq: "\n. approx_pd n t = t" +lemma approx_pd_covers: "\n. approx_pd n t = t" apply (subgoal_tac "\n. \m\n. approx_pd m t = t", fast) apply (induct t rule: pd_basis_induct) -apply (cut_tac a=a in ex_compact_approx_eq) +apply (cut_tac a=a in compact_basis.take_covers) apply (clarify, rule_tac x=n in exI) apply (clarify, simp) apply (rule antisym_less) -apply (rule compact_approx_le) -apply (drule_tac a=a in compact_approx_mono1) +apply (rule compact_basis.take_less) +apply (drule_tac a=a in compact_basis.take_chain_le) apply simp apply (clarify, rename_tac i j) apply (rule_tac x="max i j" in exI, simp) diff -r 274b80691259 -r c49d427867aa src/HOLCF/ConvexPD.thy --- a/src/HOLCF/ConvexPD.thy Thu Jun 19 22:43:59 2008 +0200 +++ b/src/HOLCF/ConvexPD.thy Thu Jun 19 22:50:58 2008 +0200 @@ -117,16 +117,16 @@ apply (simp add: 4) done -lemma approx_pd_convex_mono1: - "i \ j \ approx_pd i t \\ approx_pd j t" +lemma approx_pd_convex_chain: + "approx_pd n t \\ approx_pd (Suc n) t" apply (induct t rule: pd_basis_induct) -apply (simp add: compact_approx_mono1) +apply (simp add: compact_basis.take_chain) apply (simp add: PDPlus_convex_mono) done lemma approx_pd_convex_le: "approx_pd i t \\ t" apply (induct t rule: pd_basis_induct) -apply (simp add: compact_approx_le) +apply (simp add: compact_basis.take_less) apply (simp add: PDPlus_convex_mono) done @@ -134,7 +134,7 @@ "t \\ u \ approx_pd n t \\ approx_pd n u" apply (erule convex_le_induct) apply (erule (1) convex_le_trans) -apply (simp add: compact_approx_mono) +apply (simp add: compact_basis.take_mono) apply (simp add: PDPlus_convex_mono) done @@ -170,29 +170,16 @@ apply (rule approx_pd_convex_le) apply (rule approx_pd_idem) apply (erule approx_pd_convex_mono) -apply (rule approx_pd_convex_mono1, simp) +apply (rule approx_pd_convex_chain) apply (rule finite_range_approx_pd) -apply (rule ex_approx_pd_eq) +apply (rule approx_pd_covers) apply (rule ideal_Rep_convex_pd) apply (rule cont_Rep_convex_pd) apply (rule Rep_convex_principal) apply (simp only: less_convex_pd_def less_set_eq) done -lemma convex_principal_less_iff [simp]: - "convex_principal t \ convex_principal u \ t \\ u" -by (rule convex_pd.principal_less_iff) - -lemma convex_principal_eq_iff [simp]: - "convex_principal t = convex_principal u \ t \\ u \ u \\ t" -by (rule convex_pd.principal_eq_iff) - -lemma convex_principal_mono: - "t \\ u \ convex_principal t \ convex_principal u" -by (rule convex_pd.principal_mono) - -lemma compact_convex_principal: "compact (convex_principal t)" -by (rule convex_pd.compact_principal) +text {* Convex powerdomain is pointed *} lemma convex_pd_minimal: "convex_principal (PDUnit compact_bot) \ ys" by (induct ys rule: convex_pd.principal_induct, simp, simp) @@ -203,8 +190,7 @@ lemma inst_convex_pd_pcpo: "\ = convex_principal (PDUnit compact_bot)" by (rule convex_pd_minimal [THEN UU_I, symmetric]) - -subsection {* Approximation *} +text {* Convex powerdomain is profinite *} instantiation convex_pd :: (profinite) profinite begin @@ -234,24 +220,6 @@ unfolding approx_convex_pd_def by (rule convex_pd.completion_approx_eq_principal) -lemma compact_imp_convex_principal: - "compact xs \ \t. xs = convex_principal t" -by (rule convex_pd.compact_imp_principal) - -lemma convex_principal_induct: - "\adm P; \t. P (convex_principal t)\ \ P xs" -by (rule convex_pd.principal_induct) - -lemma convex_principal_induct2: - "\\ys. adm (\xs. P xs ys); \xs. adm (\ys. P xs ys); - \t u. P (convex_principal t) (convex_principal u)\ \ P xs ys" -apply (rule_tac x=ys in spec) -apply (rule_tac xs=xs in convex_principal_induct, simp) -apply (rule allI, rename_tac ys) -apply (rule_tac xs=ys in convex_principal_induct, simp) -apply simp -done - subsection {* Monadic unit and plus *} @@ -279,8 +247,7 @@ lemma convex_unit_Rep_compact_basis [simp]: "{Rep_compact_basis a}\ = convex_principal (PDUnit a)" unfolding convex_unit_def -by (simp add: compact_basis.basis_fun_principal - convex_principal_mono PDUnit_convex_mono) +by (simp add: compact_basis.basis_fun_principal PDUnit_convex_mono) lemma convex_plus_principal [simp]: "convex_principal t +\ convex_principal u = convex_principal (PDPlus t u)" @@ -290,28 +257,28 @@ lemma approx_convex_unit [simp]: "approx n\{x}\ = {approx n\x}\" -apply (induct x rule: compact_basis_induct, simp) +apply (induct x rule: compact_basis.principal_induct, simp) apply (simp add: approx_Rep_compact_basis) done lemma approx_convex_plus [simp]: "approx n\(xs +\ ys) = approx n\xs +\ approx n\ys" -by (induct xs ys rule: convex_principal_induct2, simp, simp, simp) +by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) lemma convex_plus_assoc: "(xs +\ ys) +\ zs = xs +\ (ys +\ zs)" -apply (induct xs ys arbitrary: zs rule: convex_principal_induct2, simp, simp) -apply (rule_tac xs=zs in convex_principal_induct, simp) +apply (induct xs ys arbitrary: zs rule: convex_pd.principal_induct2, simp, simp) +apply (rule_tac x=zs in convex_pd.principal_induct, simp) apply (simp add: PDPlus_assoc) done lemma convex_plus_commute: "xs +\ ys = ys +\ xs" -apply (induct xs ys rule: convex_principal_induct2, simp, simp) +apply (induct xs ys rule: convex_pd.principal_induct2, simp, simp) apply (simp add: PDPlus_commute) done lemma convex_plus_absorb: "xs +\ xs = xs" -apply (induct xs rule: convex_principal_induct, simp) +apply (induct xs rule: convex_pd.principal_induct, simp) apply (simp add: PDPlus_absorb) done @@ -334,9 +301,9 @@ "adm (\f. f\{x}\ \ f\ys \ f\{x}\ \ f\zs)") apply (drule admD, rule chain_approx) apply (drule_tac f="approx i" in monofun_cfun_arg) - apply (cut_tac x="approx i\x" in compact_imp_Rep_compact_basis, simp) - apply (cut_tac xs="approx i\ys" in compact_imp_convex_principal, simp) - apply (cut_tac xs="approx i\zs" in compact_imp_convex_principal, simp) + apply (cut_tac x="approx i\x" in compact_basis.compact_imp_principal, simp) + apply (cut_tac x="approx i\ys" in convex_pd.compact_imp_principal, simp) + apply (cut_tac x="approx i\zs" in convex_pd.compact_imp_principal, simp) apply (clarify, simp) apply simp apply simp @@ -352,9 +319,9 @@ "adm (\f. f\xs \ f\{z}\ \ f\ys \ f\{z}\)") apply (drule admD, rule chain_approx) apply (drule_tac f="approx i" in monofun_cfun_arg) - apply (cut_tac xs="approx i\xs" in compact_imp_convex_principal, simp) - apply (cut_tac xs="approx i\ys" in compact_imp_convex_principal, simp) - apply (cut_tac x="approx i\z" in compact_imp_Rep_compact_basis, simp) + apply (cut_tac x="approx i\xs" in convex_pd.compact_imp_principal, simp) + apply (cut_tac x="approx i\ys" in convex_pd.compact_imp_principal, simp) + apply (cut_tac x="approx i\z" in compact_basis.compact_imp_principal, simp) apply (clarify, simp) apply simp apply simp @@ -367,9 +334,9 @@ apply (rule iffI) apply (rule bifinite_less_ext) apply (drule_tac f="approx i" in monofun_cfun_arg, simp) - apply (cut_tac x="approx i\x" in compact_imp_Rep_compact_basis, simp) - apply (cut_tac x="approx i\y" in compact_imp_Rep_compact_basis, simp) - apply (clarify, simp add: compact_le_def) + apply (cut_tac x="approx i\x" in compact_basis.compact_imp_principal, simp) + apply (cut_tac x="approx i\y" in compact_basis.compact_imp_principal, simp) + apply clarsimp apply (erule monofun_cfun_arg) done @@ -388,9 +355,7 @@ lemma compact_convex_plus [simp]: "\compact xs; compact ys\ \ compact (xs +\ ys)" -apply (drule compact_imp_convex_principal)+ -apply (auto simp add: compact_convex_principal) -done +by (auto dest!: convex_pd.compact_imp_principal) subsection {* Induction rules *} @@ -400,8 +365,8 @@ assumes unit: "\x. P {x}\" assumes insert: "\x ys. \P {x}\; P ys\ \ P ({x}\ +\ ys)" shows "P (xs::'a convex_pd)" -apply (induct xs rule: convex_principal_induct, rule P) -apply (induct_tac t rule: pd_basis_induct1) +apply (induct xs rule: convex_pd.principal_induct, rule P) +apply (induct_tac a rule: pd_basis_induct1) apply (simp only: convex_unit_Rep_compact_basis [symmetric]) apply (rule unit) apply (simp only: convex_unit_Rep_compact_basis [symmetric] @@ -414,8 +379,8 @@ assumes unit: "\x. P {x}\" assumes plus: "\xs ys. \P xs; P ys\ \ P (xs +\ ys)" shows "P (xs::'a convex_pd)" -apply (induct xs rule: convex_principal_induct, rule P) -apply (induct_tac t rule: pd_basis_induct) +apply (induct xs rule: convex_pd.principal_induct, rule P) +apply (induct_tac a rule: pd_basis_induct) apply (simp only: convex_unit_Rep_compact_basis [symmetric] unit) apply (simp only: convex_plus_principal [symmetric] plus) done @@ -457,8 +422,8 @@ "t \\ u \ convex_bind_basis t \ convex_bind_basis u" apply (erule convex_le_induct) apply (erule (1) trans_less) -apply (simp add: monofun_LAM compact_le_def monofun_cfun) -apply (simp add: monofun_LAM compact_le_def monofun_cfun) +apply (simp add: monofun_LAM monofun_cfun) +apply (simp add: monofun_LAM monofun_cfun) done definition @@ -474,11 +439,11 @@ lemma convex_bind_unit [simp]: "convex_bind\{x}\\f = f\x" -by (induct x rule: compact_basis_induct, simp, simp) +by (induct x rule: compact_basis.principal_induct, simp, simp) lemma convex_bind_plus [simp]: "convex_bind\(xs +\ ys)\f = convex_bind\xs\f +\ convex_bind\ys\f" -by (induct xs ys rule: convex_principal_induct2, simp, simp, simp) +by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) lemma convex_bind_strict [simp]: "convex_bind\\\f = f\\" unfolding convex_unit_strict [symmetric] by (rule convex_bind_unit) @@ -549,22 +514,36 @@ "convex_to_upper\(convex_principal t) = upper_principal t" unfolding convex_to_upper_def apply (rule convex_pd.basis_fun_principal) -apply (rule upper_principal_mono) +apply (rule upper_pd.principal_mono) apply (erule convex_le_imp_upper_le) done lemma convex_to_upper_unit [simp]: "convex_to_upper\{x}\ = {x}\" -by (induct x rule: compact_basis_induct, simp, simp) +by (induct x rule: compact_basis.principal_induct, simp, simp) lemma convex_to_upper_plus [simp]: "convex_to_upper\(xs +\ ys) = convex_to_upper\xs +\ convex_to_upper\ys" -by (induct xs ys rule: convex_principal_induct2, simp, simp, simp) +by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) lemma approx_convex_to_upper: "approx i\(convex_to_upper\xs) = convex_to_upper\(approx i\xs)" by (induct xs rule: convex_pd_induct, simp, simp, simp) +lemma convex_to_upper_bind [simp]: + "convex_to_upper\(convex_bind\xs\f) = + upper_bind\(convex_to_upper\xs)\(convex_to_upper oo f)" +by (induct xs rule: convex_pd_induct, simp, simp, simp) + +lemma convex_to_upper_map [simp]: + "convex_to_upper\(convex_map\f\xs) = upper_map\f\(convex_to_upper\xs)" +by (simp add: convex_map_def upper_map_def cfcomp_LAM) + +lemma convex_to_upper_join [simp]: + "convex_to_upper\(convex_join\xss) = + upper_bind\(convex_to_upper\xss)\convex_to_upper" +by (simp add: convex_join_def upper_join_def cfcomp_LAM eta_cfun) + text {* Convex to lower *} lemma convex_le_imp_lower_le: "t \\ u \ t \\ u" @@ -578,22 +557,36 @@ "convex_to_lower\(convex_principal t) = lower_principal t" unfolding convex_to_lower_def apply (rule convex_pd.basis_fun_principal) -apply (rule lower_principal_mono) +apply (rule lower_pd.principal_mono) apply (erule convex_le_imp_lower_le) done lemma convex_to_lower_unit [simp]: "convex_to_lower\{x}\ = {x}\" -by (induct x rule: compact_basis_induct, simp, simp) +by (induct x rule: compact_basis.principal_induct, simp, simp) lemma convex_to_lower_plus [simp]: "convex_to_lower\(xs +\ ys) = convex_to_lower\xs +\ convex_to_lower\ys" -by (induct xs ys rule: convex_principal_induct2, simp, simp, simp) +by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) lemma approx_convex_to_lower: "approx i\(convex_to_lower\xs) = convex_to_lower\(approx i\xs)" by (induct xs rule: convex_pd_induct, simp, simp, simp) +lemma convex_to_lower_bind [simp]: + "convex_to_lower\(convex_bind\xs\f) = + lower_bind\(convex_to_lower\xs)\(convex_to_lower oo f)" +by (induct xs rule: convex_pd_induct, simp, simp, simp) + +lemma convex_to_lower_map [simp]: + "convex_to_lower\(convex_map\f\xs) = lower_map\f\(convex_to_lower\xs)" +by (simp add: convex_map_def lower_map_def cfcomp_LAM) + +lemma convex_to_lower_join [simp]: + "convex_to_lower\(convex_join\xss) = + lower_bind\(convex_to_lower\xss)\convex_to_lower" +by (simp add: convex_join_def lower_join_def cfcomp_LAM eta_cfun) + text {* Ordering property *} lemma convex_pd_less_iff: @@ -604,8 +597,8 @@ apply (rule bifinite_less_ext) apply (drule_tac f="approx i" in monofun_cfun_arg) apply (drule_tac f="approx i" in monofun_cfun_arg) - apply (cut_tac xs="approx i\xs" in compact_imp_convex_principal, simp) - apply (cut_tac xs="approx i\ys" in compact_imp_convex_principal, simp) + apply (cut_tac x="approx i\xs" in convex_pd.compact_imp_principal, simp) + apply (cut_tac x="approx i\ys" in convex_pd.compact_imp_principal, simp) apply clarify apply (simp add: approx_convex_to_upper approx_convex_to_lower convex_le_def) done diff -r 274b80691259 -r c49d427867aa src/HOLCF/LowerPD.thy --- a/src/HOLCF/LowerPD.thy Thu Jun 19 22:43:59 2008 +0200 +++ b/src/HOLCF/LowerPD.thy Thu Jun 19 22:50:58 2008 +0200 @@ -72,23 +72,23 @@ apply (simp add: lower_le_PDPlus_iff 3) done -lemma approx_pd_lower_mono1: - "i \ j \ approx_pd i t \\ approx_pd j t" +lemma approx_pd_lower_chain: + "approx_pd n t \\ approx_pd (Suc n) t" apply (induct t rule: pd_basis_induct) -apply (simp add: compact_approx_mono1) +apply (simp add: compact_basis.take_chain) apply (simp add: PDPlus_lower_mono) done lemma approx_pd_lower_le: "approx_pd i t \\ t" apply (induct t rule: pd_basis_induct) -apply (simp add: compact_approx_le) +apply (simp add: compact_basis.take_less) apply (simp add: PDPlus_lower_mono) done lemma approx_pd_lower_mono: "t \\ u \ approx_pd n t \\ approx_pd n u" apply (erule lower_le_induct) -apply (simp add: compact_approx_mono) +apply (simp add: compact_basis.take_mono) apply (simp add: lower_le_PDUnit_PDPlus_iff) apply (simp add: lower_le_PDPlus_iff) done @@ -122,29 +122,16 @@ apply (rule approx_pd_lower_le) apply (rule approx_pd_idem) apply (erule approx_pd_lower_mono) -apply (rule approx_pd_lower_mono1, simp) +apply (rule approx_pd_lower_chain) apply (rule finite_range_approx_pd) -apply (rule ex_approx_pd_eq) +apply (rule approx_pd_covers) apply (rule ideal_Rep_lower_pd) apply (rule cont_Rep_lower_pd) apply (rule Rep_lower_principal) apply (simp only: less_lower_pd_def less_set_eq) done -lemma lower_principal_less_iff [simp]: - "lower_principal t \ lower_principal u \ t \\ u" -by (rule lower_pd.principal_less_iff) - -lemma lower_principal_eq_iff: - "lower_principal t = lower_principal u \ t \\ u \ u \\ t" -by (rule lower_pd.principal_eq_iff) - -lemma lower_principal_mono: - "t \\ u \ lower_principal t \ lower_principal u" -by (rule lower_pd.principal_mono) - -lemma compact_lower_principal: "compact (lower_principal t)" -by (rule lower_pd.compact_principal) +text {* Lower powerdomain is pointed *} lemma lower_pd_minimal: "lower_principal (PDUnit compact_bot) \ ys" by (induct ys rule: lower_pd.principal_induct, simp, simp) @@ -155,8 +142,7 @@ lemma inst_lower_pd_pcpo: "\ = lower_principal (PDUnit compact_bot)" by (rule lower_pd_minimal [THEN UU_I, symmetric]) - -subsection {* Approximation *} +text {* Lower powerdomain is profinite *} instantiation lower_pd :: (profinite) profinite begin @@ -186,24 +172,6 @@ unfolding approx_lower_pd_def by (rule lower_pd.completion_approx_eq_principal) -lemma compact_imp_lower_principal: - "compact xs \ \t. xs = lower_principal t" -by (rule lower_pd.compact_imp_principal) - -lemma lower_principal_induct: - "\adm P; \t. P (lower_principal t)\ \ P xs" -by (rule lower_pd.principal_induct) - -lemma lower_principal_induct2: - "\\ys. adm (\xs. P xs ys); \xs. adm (\ys. P xs ys); - \t u. P (lower_principal t) (lower_principal u)\ \ P xs ys" -apply (rule_tac x=ys in spec) -apply (rule_tac xs=xs in lower_principal_induct, simp) -apply (rule allI, rename_tac ys) -apply (rule_tac xs=ys in lower_principal_induct, simp) -apply simp -done - subsection {* Monadic unit and plus *} @@ -231,8 +199,7 @@ lemma lower_unit_Rep_compact_basis [simp]: "{Rep_compact_basis a}\ = lower_principal (PDUnit a)" unfolding lower_unit_def -by (simp add: compact_basis.basis_fun_principal - lower_principal_mono PDUnit_lower_mono) +by (simp add: compact_basis.basis_fun_principal PDUnit_lower_mono) lemma lower_plus_principal [simp]: "lower_principal t +\ lower_principal u = lower_principal (PDPlus t u)" @@ -242,27 +209,27 @@ lemma approx_lower_unit [simp]: "approx n\{x}\ = {approx n\x}\" -apply (induct x rule: compact_basis_induct, simp) +apply (induct x rule: compact_basis.principal_induct, simp) apply (simp add: approx_Rep_compact_basis) done lemma approx_lower_plus [simp]: "approx n\(xs +\ ys) = (approx n\xs) +\ (approx n\ys)" -by (induct xs ys rule: lower_principal_induct2, simp, simp, simp) +by (induct xs ys rule: lower_pd.principal_induct2, simp, simp, simp) lemma lower_plus_assoc: "(xs +\ ys) +\ zs = xs +\ (ys +\ zs)" -apply (induct xs ys arbitrary: zs rule: lower_principal_induct2, simp, simp) -apply (rule_tac xs=zs in lower_principal_induct, simp) +apply (induct xs ys arbitrary: zs rule: lower_pd.principal_induct2, simp, simp) +apply (rule_tac x=zs in lower_pd.principal_induct, simp) apply (simp add: PDPlus_assoc) done lemma lower_plus_commute: "xs +\ ys = ys +\ xs" -apply (induct xs ys rule: lower_principal_induct2, simp, simp) +apply (induct xs ys rule: lower_pd.principal_induct2, simp, simp) apply (simp add: PDPlus_commute) done lemma lower_plus_absorb: "xs +\ xs = xs" -apply (induct xs rule: lower_principal_induct, simp) +apply (induct xs rule: lower_pd.principal_induct, simp) apply (simp add: PDPlus_absorb) done @@ -279,7 +246,7 @@ lemmas lower_plus_aci = aci_lower_plus.mult_ac_idem lemma lower_plus_less1: "xs \ xs +\ ys" -apply (induct xs ys rule: lower_principal_induct2, simp, simp) +apply (induct xs ys rule: lower_pd.principal_induct2, simp, simp) apply (simp add: PDPlus_lower_less) done @@ -306,9 +273,9 @@ "adm (\f. f\{x}\ \ f\ys \ f\{x}\ \ f\zs)") apply (drule admD, rule chain_approx) apply (drule_tac f="approx i" in monofun_cfun_arg) - apply (cut_tac x="approx i\x" in compact_imp_Rep_compact_basis, simp) - apply (cut_tac xs="approx i\ys" in compact_imp_lower_principal, simp) - apply (cut_tac xs="approx i\zs" in compact_imp_lower_principal, simp) + apply (cut_tac x="approx i\x" in compact_basis.compact_imp_principal, simp) + apply (cut_tac x="approx i\ys" in lower_pd.compact_imp_principal, simp) + apply (cut_tac x="approx i\zs" in lower_pd.compact_imp_principal, simp) apply (clarify, simp add: lower_le_PDUnit_PDPlus_iff) apply simp apply simp @@ -321,9 +288,9 @@ apply (rule iffI) apply (rule bifinite_less_ext) apply (drule_tac f="approx i" in monofun_cfun_arg, simp) - apply (cut_tac x="approx i\x" in compact_imp_Rep_compact_basis, simp) - apply (cut_tac x="approx i\y" in compact_imp_Rep_compact_basis, simp) - apply (clarify, simp add: compact_le_def) + apply (cut_tac x="approx i\x" in compact_basis.compact_imp_principal, simp) + apply (cut_tac x="approx i\y" in compact_basis.compact_imp_principal, simp) + apply clarsimp apply (erule monofun_cfun_arg) done @@ -332,8 +299,14 @@ lower_plus_less_iff lower_unit_less_plus_iff +lemma fooble: + fixes f :: "'a::po \ 'b::po" + assumes f: "\x y. f x \ f y \ x \ y" + shows "f x = f y \ x = y" +unfolding po_eq_conv by (simp add: f) + lemma lower_unit_eq_iff [simp]: "{x}\ = {y}\ \ x = y" -unfolding po_eq_conv by simp +by (rule lower_unit_less_iff [THEN fooble]) lemma lower_unit_strict [simp]: "{\}\ = \" unfolding inst_lower_pd_pcpo Rep_compact_bot [symmetric] by simp @@ -364,9 +337,7 @@ lemma compact_lower_plus [simp]: "\compact xs; compact ys\ \ compact (xs +\ ys)" -apply (drule compact_imp_lower_principal)+ -apply (auto simp add: compact_lower_principal) -done +by (auto dest!: lower_pd.compact_imp_principal) subsection {* Induction rules *} @@ -377,8 +348,8 @@ assumes insert: "\x ys. \P {x}\; P ys\ \ P ({x}\ +\ ys)" shows "P (xs::'a lower_pd)" -apply (induct xs rule: lower_principal_induct, rule P) -apply (induct_tac t rule: pd_basis_induct1) +apply (induct xs rule: lower_pd.principal_induct, rule P) +apply (induct_tac a rule: pd_basis_induct1) apply (simp only: lower_unit_Rep_compact_basis [symmetric]) apply (rule unit) apply (simp only: lower_unit_Rep_compact_basis [symmetric] @@ -391,8 +362,8 @@ assumes unit: "\x. P {x}\" assumes plus: "\xs ys. \P xs; P ys\ \ P (xs +\ ys)" shows "P (xs::'a lower_pd)" -apply (induct xs rule: lower_principal_induct, rule P) -apply (induct_tac t rule: pd_basis_induct) +apply (induct xs rule: lower_pd.principal_induct, rule P) +apply (induct_tac a rule: pd_basis_induct) apply (simp only: lower_unit_Rep_compact_basis [symmetric] unit) apply (simp only: lower_plus_principal [symmetric] plus) done @@ -430,7 +401,7 @@ "t \\ u \ lower_bind_basis t \ lower_bind_basis u" unfolding expand_cfun_less apply (erule lower_le_induct, safe) -apply (simp add: compact_le_def monofun_cfun) +apply (simp add: monofun_cfun) apply (simp add: rev_trans_less [OF lower_plus_less1]) apply (simp add: lower_plus_less_iff) done @@ -448,11 +419,11 @@ lemma lower_bind_unit [simp]: "lower_bind\{x}\\f = f\x" -by (induct x rule: compact_basis_induct, simp, simp) +by (induct x rule: compact_basis.principal_induct, simp, simp) lemma lower_bind_plus [simp]: "lower_bind\(xs +\ ys)\f = lower_bind\xs\f +\ lower_bind\ys\f" -by (induct xs ys rule: lower_principal_induct2, simp, simp, simp) +by (induct xs ys rule: lower_pd.principal_induct2, simp, simp, simp) lemma lower_bind_strict [simp]: "lower_bind\\\f = f\\" unfolding lower_unit_strict [symmetric] by (rule lower_bind_unit) diff -r 274b80691259 -r c49d427867aa src/HOLCF/UpperPD.thy --- a/src/HOLCF/UpperPD.thy Thu Jun 19 22:43:59 2008 +0200 +++ b/src/HOLCF/UpperPD.thy Thu Jun 19 22:50:58 2008 +0200 @@ -70,23 +70,23 @@ apply (simp add: upper_le_PDPlus_iff 3) done -lemma approx_pd_upper_mono1: - "i \ j \ approx_pd i t \\ approx_pd j t" +lemma approx_pd_upper_chain: + "approx_pd n t \\ approx_pd (Suc n) t" apply (induct t rule: pd_basis_induct) -apply (simp add: compact_approx_mono1) +apply (simp add: compact_basis.take_chain) apply (simp add: PDPlus_upper_mono) done lemma approx_pd_upper_le: "approx_pd i t \\ t" apply (induct t rule: pd_basis_induct) -apply (simp add: compact_approx_le) +apply (simp add: compact_basis.take_less) apply (simp add: PDPlus_upper_mono) done lemma approx_pd_upper_mono: "t \\ u \ approx_pd n t \\ approx_pd n u" apply (erule upper_le_induct) -apply (simp add: compact_approx_mono) +apply (simp add: compact_basis.take_mono) apply (simp add: upper_le_PDPlus_PDUnit_iff) apply (simp add: upper_le_PDPlus_iff) done @@ -120,29 +120,16 @@ apply (rule approx_pd_upper_le) apply (rule approx_pd_idem) apply (erule approx_pd_upper_mono) -apply (rule approx_pd_upper_mono1, simp) +apply (rule approx_pd_upper_chain) apply (rule finite_range_approx_pd) -apply (rule ex_approx_pd_eq) +apply (rule approx_pd_covers) apply (rule ideal_Rep_upper_pd) apply (rule cont_Rep_upper_pd) apply (rule Rep_upper_principal) apply (simp only: less_upper_pd_def less_set_eq) done -lemma upper_principal_less_iff [simp]: - "upper_principal t \ upper_principal u \ t \\ u" -by (rule upper_pd.principal_less_iff) - -lemma upper_principal_eq_iff: - "upper_principal t = upper_principal u \ t \\ u \ u \\ t" -by (rule upper_pd.principal_eq_iff) - -lemma upper_principal_mono: - "t \\ u \ upper_principal t \ upper_principal u" -by (rule upper_pd.principal_mono) - -lemma compact_upper_principal: "compact (upper_principal t)" -by (rule upper_pd.compact_principal) +text {* Upper powerdomain is pointed *} lemma upper_pd_minimal: "upper_principal (PDUnit compact_bot) \ ys" by (induct ys rule: upper_pd.principal_induct, simp, simp) @@ -153,8 +140,7 @@ lemma inst_upper_pd_pcpo: "\ = upper_principal (PDUnit compact_bot)" by (rule upper_pd_minimal [THEN UU_I, symmetric]) - -subsection {* Approximation *} +text {* Upper powerdomain is profinite *} instantiation upper_pd :: (profinite) profinite begin @@ -184,24 +170,6 @@ unfolding approx_upper_pd_def by (rule upper_pd.completion_approx_eq_principal) -lemma compact_imp_upper_principal: - "compact xs \ \t. xs = upper_principal t" -by (rule upper_pd.compact_imp_principal) - -lemma upper_principal_induct: - "\adm P; \t. P (upper_principal t)\ \ P xs" -by (rule upper_pd.principal_induct) - -lemma upper_principal_induct2: - "\\ys. adm (\xs. P xs ys); \xs. adm (\ys. P xs ys); - \t u. P (upper_principal t) (upper_principal u)\ \ P xs ys" -apply (rule_tac x=ys in spec) -apply (rule_tac xs=xs in upper_principal_induct, simp) -apply (rule allI, rename_tac ys) -apply (rule_tac xs=ys in upper_principal_induct, simp) -apply simp -done - subsection {* Monadic unit and plus *} @@ -229,8 +197,7 @@ lemma upper_unit_Rep_compact_basis [simp]: "{Rep_compact_basis a}\ = upper_principal (PDUnit a)" unfolding upper_unit_def -by (simp add: compact_basis.basis_fun_principal - upper_principal_mono PDUnit_upper_mono) +by (simp add: compact_basis.basis_fun_principal PDUnit_upper_mono) lemma upper_plus_principal [simp]: "upper_principal t +\ upper_principal u = upper_principal (PDPlus t u)" @@ -240,27 +207,27 @@ lemma approx_upper_unit [simp]: "approx n\{x}\ = {approx n\x}\" -apply (induct x rule: compact_basis_induct, simp) +apply (induct x rule: compact_basis.principal_induct, simp) apply (simp add: approx_Rep_compact_basis) done lemma approx_upper_plus [simp]: "approx n\(xs +\ ys) = (approx n\xs) +\ (approx n\ys)" -by (induct xs ys rule: upper_principal_induct2, simp, simp, simp) +by (induct xs ys rule: upper_pd.principal_induct2, simp, simp, simp) lemma upper_plus_assoc: "(xs +\ ys) +\ zs = xs +\ (ys +\ zs)" -apply (induct xs ys arbitrary: zs rule: upper_principal_induct2, simp, simp) -apply (rule_tac xs=zs in upper_principal_induct, simp) +apply (induct xs ys arbitrary: zs rule: upper_pd.principal_induct2, simp, simp) +apply (rule_tac x=zs in upper_pd.principal_induct, simp) apply (simp add: PDPlus_assoc) done lemma upper_plus_commute: "xs +\ ys = ys +\ xs" -apply (induct xs ys rule: upper_principal_induct2, simp, simp) +apply (induct xs ys rule: upper_pd.principal_induct2, simp, simp) apply (simp add: PDPlus_commute) done lemma upper_plus_absorb: "xs +\ xs = xs" -apply (induct xs rule: upper_principal_induct, simp) +apply (induct xs rule: upper_pd.principal_induct, simp) apply (simp add: PDPlus_absorb) done @@ -277,7 +244,7 @@ lemmas upper_plus_aci = aci_upper_plus.mult_ac_idem lemma upper_plus_less1: "xs +\ ys \ xs" -apply (induct xs ys rule: upper_principal_induct2, simp, simp) +apply (induct xs ys rule: upper_pd.principal_induct2, simp, simp) apply (simp add: PDPlus_upper_less) done @@ -304,9 +271,9 @@ "adm (\f. f\xs \ f\{z}\ \ f\ys \ f\{z}\)") apply (drule admD, rule chain_approx) apply (drule_tac f="approx i" in monofun_cfun_arg) - apply (cut_tac xs="approx i\xs" in compact_imp_upper_principal, simp) - apply (cut_tac xs="approx i\ys" in compact_imp_upper_principal, simp) - apply (cut_tac x="approx i\z" in compact_imp_Rep_compact_basis, simp) + apply (cut_tac x="approx i\xs" in upper_pd.compact_imp_principal, simp) + apply (cut_tac x="approx i\ys" in upper_pd.compact_imp_principal, simp) + apply (cut_tac x="approx i\z" in compact_basis.compact_imp_principal, simp) apply (clarify, simp add: upper_le_PDPlus_PDUnit_iff) apply simp apply simp @@ -319,9 +286,9 @@ apply (rule iffI) apply (rule bifinite_less_ext) apply (drule_tac f="approx i" in monofun_cfun_arg, simp) - apply (cut_tac x="approx i\x" in compact_imp_Rep_compact_basis, simp) - apply (cut_tac x="approx i\y" in compact_imp_Rep_compact_basis, simp) - apply (clarify, simp add: compact_le_def) + apply (cut_tac x="approx i\x" in compact_basis.compact_imp_principal, simp) + apply (cut_tac x="approx i\y" in compact_basis.compact_imp_principal, simp) + apply clarsimp apply (erule monofun_cfun_arg) done @@ -349,8 +316,8 @@ "xs +\ ys = \ \ xs = \ \ ys = \" apply (rule iffI) apply (erule rev_mp) -apply (rule upper_principal_induct2 [where xs=xs and ys=ys], simp, simp) -apply (simp add: inst_upper_pd_pcpo upper_principal_eq_iff +apply (rule upper_pd.principal_induct2 [where x=xs and y=ys], simp, simp) +apply (simp add: inst_upper_pd_pcpo upper_pd.principal_eq_iff upper_le_PDPlus_PDUnit_iff) apply auto done @@ -360,9 +327,7 @@ lemma compact_upper_plus [simp]: "\compact xs; compact ys\ \ compact (xs +\ ys)" -apply (drule compact_imp_upper_principal)+ -apply (auto simp add: compact_upper_principal) -done +by (auto dest!: upper_pd.compact_imp_principal) subsection {* Induction rules *} @@ -372,8 +337,8 @@ assumes unit: "\x. P {x}\" assumes insert: "\x ys. \P {x}\; P ys\ \ P ({x}\ +\ ys)" shows "P (xs::'a upper_pd)" -apply (induct xs rule: upper_principal_induct, rule P) -apply (induct_tac t rule: pd_basis_induct1) +apply (induct xs rule: upper_pd.principal_induct, rule P) +apply (induct_tac a rule: pd_basis_induct1) apply (simp only: upper_unit_Rep_compact_basis [symmetric]) apply (rule unit) apply (simp only: upper_unit_Rep_compact_basis [symmetric] @@ -386,8 +351,8 @@ assumes unit: "\x. P {x}\" assumes plus: "\xs ys. \P xs; P ys\ \ P (xs +\ ys)" shows "P (xs::'a upper_pd)" -apply (induct xs rule: upper_principal_induct, rule P) -apply (induct_tac t rule: pd_basis_induct) +apply (induct xs rule: upper_pd.principal_induct, rule P) +apply (induct_tac a rule: pd_basis_induct) apply (simp only: upper_unit_Rep_compact_basis [symmetric] unit) apply (simp only: upper_plus_principal [symmetric] plus) done @@ -425,7 +390,7 @@ "t \\ u \ upper_bind_basis t \ upper_bind_basis u" unfolding expand_cfun_less apply (erule upper_le_induct, safe) -apply (simp add: compact_le_def monofun_cfun) +apply (simp add: monofun_cfun) apply (simp add: trans_less [OF upper_plus_less1]) apply (simp add: upper_less_plus_iff) done @@ -443,11 +408,11 @@ lemma upper_bind_unit [simp]: "upper_bind\{x}\\f = f\x" -by (induct x rule: compact_basis_induct, simp, simp) +by (induct x rule: compact_basis.principal_induct, simp, simp) lemma upper_bind_plus [simp]: "upper_bind\(xs +\ ys)\f = upper_bind\xs\f +\ upper_bind\ys\f" -by (induct xs ys rule: upper_principal_induct2, simp, simp, simp) +by (induct xs ys rule: upper_pd.principal_induct2, simp, simp, simp) lemma upper_bind_strict [simp]: "upper_bind\\\f = f\\" unfolding upper_unit_strict [symmetric] by (rule upper_bind_unit)