# HG changeset patch # User huffman # Date 1268069585 28800 # Node ID f87132febfac789dc0d5f6bbb931e1150dc493d4 # Parent 05ca920cd94b23fd7eb50d2452e5c27a4066c02b move lemmas from Domain.thy to Domain_Aux.thy diff -r 05ca920cd94b -r f87132febfac src/HOLCF/Domain.thy --- a/src/HOLCF/Domain.thy Mon Mar 08 08:12:48 2010 -0800 +++ b/src/HOLCF/Domain.thy Mon Mar 08 09:33:05 2010 -0800 @@ -19,107 +19,6 @@ defaultsort pcpo -subsection {* Continuous isomorphisms *} - -text {* A locale for continuous isomorphisms *} - -locale iso = - fixes abs :: "'a \ 'b" - fixes rep :: "'b \ 'a" - assumes abs_iso [simp]: "rep\(abs\x) = x" - assumes rep_iso [simp]: "abs\(rep\y) = y" -begin - -lemma swap: "iso rep abs" - by (rule iso.intro [OF rep_iso abs_iso]) - -lemma abs_below: "(abs\x \ abs\y) = (x \ y)" -proof - assume "abs\x \ abs\y" - then have "rep\(abs\x) \ rep\(abs\y)" by (rule monofun_cfun_arg) - then show "x \ y" by simp -next - assume "x \ y" - then show "abs\x \ abs\y" by (rule monofun_cfun_arg) -qed - -lemma rep_below: "(rep\x \ rep\y) = (x \ y)" - by (rule iso.abs_below [OF swap]) - -lemma abs_eq: "(abs\x = abs\y) = (x = y)" - by (simp add: po_eq_conv abs_below) - -lemma rep_eq: "(rep\x = rep\y) = (x = y)" - by (rule iso.abs_eq [OF swap]) - -lemma abs_strict: "abs\\ = \" -proof - - have "\ \ rep\\" .. - then have "abs\\ \ abs\(rep\\)" by (rule monofun_cfun_arg) - then have "abs\\ \ \" by simp - then show ?thesis by (rule UU_I) -qed - -lemma rep_strict: "rep\\ = \" - by (rule iso.abs_strict [OF swap]) - -lemma abs_defin': "abs\x = \ \ x = \" -proof - - have "x = rep\(abs\x)" by simp - also assume "abs\x = \" - also note rep_strict - finally show "x = \" . -qed - -lemma rep_defin': "rep\z = \ \ z = \" - by (rule iso.abs_defin' [OF swap]) - -lemma abs_defined: "z \ \ \ abs\z \ \" - by (erule contrapos_nn, erule abs_defin') - -lemma rep_defined: "z \ \ \ rep\z \ \" - by (rule iso.abs_defined [OF iso.swap]) (rule iso_axioms) - -lemma abs_defined_iff: "(abs\x = \) = (x = \)" - by (auto elim: abs_defin' intro: abs_strict) - -lemma rep_defined_iff: "(rep\x = \) = (x = \)" - by (rule iso.abs_defined_iff [OF iso.swap]) (rule iso_axioms) - -lemma casedist_rule: "rep\x = \ \ P \ x = \ \ P" - by (simp add: rep_defined_iff) - -lemma compact_abs_rev: "compact (abs\x) \ compact x" -proof (unfold compact_def) - assume "adm (\y. \ abs\x \ y)" - with cont_Rep_CFun2 - have "adm (\y. \ abs\x \ abs\y)" by (rule adm_subst) - then show "adm (\y. \ x \ y)" using abs_below by simp -qed - -lemma compact_rep_rev: "compact (rep\x) \ compact x" - by (rule iso.compact_abs_rev [OF iso.swap]) (rule iso_axioms) - -lemma compact_abs: "compact x \ compact (abs\x)" - by (rule compact_rep_rev) simp - -lemma compact_rep: "compact x \ compact (rep\x)" - by (rule iso.compact_abs [OF iso.swap]) (rule iso_axioms) - -lemma iso_swap: "(x = abs\y) = (rep\x = y)" -proof - assume "x = abs\y" - then have "rep\x = rep\(abs\y)" by simp - then show "rep\x = y" by simp -next - assume "rep\x = y" - then have "abs\(rep\x) = abs\y" by simp - then show "x = abs\y" by simp -qed - -end - - subsection {* Casedist *} lemma ex_one_defined_iff: @@ -214,102 +113,6 @@ ssum_map_sinl ssum_map_sinr sprod_map_spair u_map_up -subsection {* Take functions and finiteness *} - -lemma lub_ID_take_lemma: - assumes "chain t" and "(\n. t n) = ID" - assumes "\n. t n\x = t n\y" shows "x = y" -proof - - have "(\n. t n\x) = (\n. t n\y)" - using assms(3) by simp - then have "(\n. t n)\x = (\n. t n)\y" - using assms(1) by (simp add: lub_distribs) - then show "x = y" - using assms(2) by simp -qed - -lemma lub_ID_reach: - assumes "chain t" and "(\n. t n) = ID" - shows "(\n. t n\x) = x" -using assms by (simp add: lub_distribs) - -text {* - Let a ``decisive'' function be a deflation that maps every input to - either itself or bottom. Then if a domain's take functions are all - decisive, then all values in the domain are finite. -*} - -definition - decisive :: "('a::pcpo \ 'a) \ bool" -where - "decisive d \ (\x. d\x = x \ d\x = \)" - -lemma decisiveI: "(\x. d\x = x \ d\x = \) \ decisive d" - unfolding decisive_def by simp - -lemma decisive_cases: - assumes "decisive d" obtains "d\x = x" | "d\x = \" -using assms unfolding decisive_def by auto - -lemma decisive_bottom: "decisive \" - unfolding decisive_def by simp - -lemma decisive_ID: "decisive ID" - unfolding decisive_def by simp - -lemma decisive_ssum_map: - assumes f: "decisive f" - assumes g: "decisive g" - shows "decisive (ssum_map\f\g)" -apply (rule decisiveI, rename_tac s) -apply (case_tac s, simp_all) -apply (rule_tac x=x in decisive_cases [OF f], simp_all) -apply (rule_tac x=y in decisive_cases [OF g], simp_all) -done - -lemma decisive_sprod_map: - assumes f: "decisive f" - assumes g: "decisive g" - shows "decisive (sprod_map\f\g)" -apply (rule decisiveI, rename_tac s) -apply (case_tac s, simp_all) -apply (rule_tac x=x in decisive_cases [OF f], simp_all) -apply (rule_tac x=y in decisive_cases [OF g], simp_all) -done - -lemma decisive_abs_rep: - fixes abs rep - assumes iso: "iso abs rep" - assumes d: "decisive d" - shows "decisive (abs oo d oo rep)" -apply (rule decisiveI) -apply (rule_tac x="rep\x" in decisive_cases [OF d]) -apply (simp add: iso.rep_iso [OF iso]) -apply (simp add: iso.abs_strict [OF iso]) -done - -lemma lub_ID_finite: - assumes chain: "chain d" - assumes lub: "(\n. d n) = ID" - assumes decisive: "\n. decisive (d n)" - shows "\n. d n\x = x" -proof - - have 1: "chain (\n. d n\x)" using chain by simp - have 2: "(\n. d n\x) = x" using chain lub by (rule lub_ID_reach) - have "\n. d n\x = x \ d n\x = \" - using decisive unfolding decisive_def by simp - hence "range (\n. d n\x) \ {x, \}" - by auto - hence "finite (range (\n. d n\x))" - by (rule finite_subset, simp) - with 1 have "finite_chain (\n. d n\x)" - by (rule finite_range_imp_finch) - then have "\n. (\n. d n\x) = d n\x" - unfolding finite_chain_def by (auto simp add: maxinch_is_thelub) - with 2 show "\n. d n\x = x" by (auto elim: sym) -qed - - subsection {* Installing the domain package *} lemmas con_strict_rules = diff -r 05ca920cd94b -r f87132febfac src/HOLCF/Domain_Aux.thy --- a/src/HOLCF/Domain_Aux.thy Mon Mar 08 08:12:48 2010 -0800 +++ b/src/HOLCF/Domain_Aux.thy Mon Mar 08 09:33:05 2010 -0800 @@ -10,6 +10,107 @@ ("Tools/Domain/domain_take_proofs.ML") begin +subsection {* Continuous isomorphisms *} + +text {* A locale for continuous isomorphisms *} + +locale iso = + fixes abs :: "'a \ 'b" + fixes rep :: "'b \ 'a" + assumes abs_iso [simp]: "rep\(abs\x) = x" + assumes rep_iso [simp]: "abs\(rep\y) = y" +begin + +lemma swap: "iso rep abs" + by (rule iso.intro [OF rep_iso abs_iso]) + +lemma abs_below: "(abs\x \ abs\y) = (x \ y)" +proof + assume "abs\x \ abs\y" + then have "rep\(abs\x) \ rep\(abs\y)" by (rule monofun_cfun_arg) + then show "x \ y" by simp +next + assume "x \ y" + then show "abs\x \ abs\y" by (rule monofun_cfun_arg) +qed + +lemma rep_below: "(rep\x \ rep\y) = (x \ y)" + by (rule iso.abs_below [OF swap]) + +lemma abs_eq: "(abs\x = abs\y) = (x = y)" + by (simp add: po_eq_conv abs_below) + +lemma rep_eq: "(rep\x = rep\y) = (x = y)" + by (rule iso.abs_eq [OF swap]) + +lemma abs_strict: "abs\\ = \" +proof - + have "\ \ rep\\" .. + then have "abs\\ \ abs\(rep\\)" by (rule monofun_cfun_arg) + then have "abs\\ \ \" by simp + then show ?thesis by (rule UU_I) +qed + +lemma rep_strict: "rep\\ = \" + by (rule iso.abs_strict [OF swap]) + +lemma abs_defin': "abs\x = \ \ x = \" +proof - + have "x = rep\(abs\x)" by simp + also assume "abs\x = \" + also note rep_strict + finally show "x = \" . +qed + +lemma rep_defin': "rep\z = \ \ z = \" + by (rule iso.abs_defin' [OF swap]) + +lemma abs_defined: "z \ \ \ abs\z \ \" + by (erule contrapos_nn, erule abs_defin') + +lemma rep_defined: "z \ \ \ rep\z \ \" + by (rule iso.abs_defined [OF iso.swap]) (rule iso_axioms) + +lemma abs_defined_iff: "(abs\x = \) = (x = \)" + by (auto elim: abs_defin' intro: abs_strict) + +lemma rep_defined_iff: "(rep\x = \) = (x = \)" + by (rule iso.abs_defined_iff [OF iso.swap]) (rule iso_axioms) + +lemma casedist_rule: "rep\x = \ \ P \ x = \ \ P" + by (simp add: rep_defined_iff) + +lemma compact_abs_rev: "compact (abs\x) \ compact x" +proof (unfold compact_def) + assume "adm (\y. \ abs\x \ y)" + with cont_Rep_CFun2 + have "adm (\y. \ abs\x \ abs\y)" by (rule adm_subst) + then show "adm (\y. \ x \ y)" using abs_below by simp +qed + +lemma compact_rep_rev: "compact (rep\x) \ compact x" + by (rule iso.compact_abs_rev [OF iso.swap]) (rule iso_axioms) + +lemma compact_abs: "compact x \ compact (abs\x)" + by (rule compact_rep_rev) simp + +lemma compact_rep: "compact x \ compact (rep\x)" + by (rule iso.compact_abs [OF iso.swap]) (rule iso_axioms) + +lemma iso_swap: "(x = abs\y) = (rep\x = y)" +proof + assume "x = abs\y" + then have "rep\x = rep\(abs\y)" by simp + then show "rep\x = y" by simp +next + assume "rep\x = y" + then have "abs\(rep\x) = abs\y" by simp + then show "x = abs\y" by simp +qed + +end + + subsection {* Proofs about take functions *} text {* @@ -45,6 +146,104 @@ ultimately show ?thesis by simp qed +lemma lub_ID_take_lemma: + assumes "chain t" and "(\n. t n) = ID" + assumes "\n. t n\x = t n\y" shows "x = y" +proof - + have "(\n. t n\x) = (\n. t n\y)" + using assms(3) by simp + then have "(\n. t n)\x = (\n. t n)\y" + using assms(1) by (simp add: lub_distribs) + then show "x = y" + using assms(2) by simp +qed + +lemma lub_ID_reach: + assumes "chain t" and "(\n. t n) = ID" + shows "(\n. t n\x) = x" +using assms by (simp add: lub_distribs) + + +subsection {* Finiteness *} + +text {* + Let a ``decisive'' function be a deflation that maps every input to + either itself or bottom. Then if a domain's take functions are all + decisive, then all values in the domain are finite. +*} + +definition + decisive :: "('a::pcpo \ 'a) \ bool" +where + "decisive d \ (\x. d\x = x \ d\x = \)" + +lemma decisiveI: "(\x. d\x = x \ d\x = \) \ decisive d" + unfolding decisive_def by simp + +lemma decisive_cases: + assumes "decisive d" obtains "d\x = x" | "d\x = \" +using assms unfolding decisive_def by auto + +lemma decisive_bottom: "decisive \" + unfolding decisive_def by simp + +lemma decisive_ID: "decisive ID" + unfolding decisive_def by simp + +lemma decisive_ssum_map: + assumes f: "decisive f" + assumes g: "decisive g" + shows "decisive (ssum_map\f\g)" +apply (rule decisiveI, rename_tac s) +apply (case_tac s, simp_all) +apply (rule_tac x=x in decisive_cases [OF f], simp_all) +apply (rule_tac x=y in decisive_cases [OF g], simp_all) +done + +lemma decisive_sprod_map: + assumes f: "decisive f" + assumes g: "decisive g" + shows "decisive (sprod_map\f\g)" +apply (rule decisiveI, rename_tac s) +apply (case_tac s, simp_all) +apply (rule_tac x=x in decisive_cases [OF f], simp_all) +apply (rule_tac x=y in decisive_cases [OF g], simp_all) +done + +lemma decisive_abs_rep: + fixes abs rep + assumes iso: "iso abs rep" + assumes d: "decisive d" + shows "decisive (abs oo d oo rep)" +apply (rule decisiveI) +apply (rule_tac x="rep\x" in decisive_cases [OF d]) +apply (simp add: iso.rep_iso [OF iso]) +apply (simp add: iso.abs_strict [OF iso]) +done + +lemma lub_ID_finite: + assumes chain: "chain d" + assumes lub: "(\n. d n) = ID" + assumes decisive: "\n. decisive (d n)" + shows "\n. d n\x = x" +proof - + have 1: "chain (\n. d n\x)" using chain by simp + have 2: "(\n. d n\x) = x" using chain lub by (rule lub_ID_reach) + have "\n. d n\x = x \ d n\x = \" + using decisive unfolding decisive_def by simp + hence "range (\n. d n\x) \ {x, \}" + by auto + hence "finite (range (\n. d n\x))" + by (rule finite_subset, simp) + with 1 have "finite_chain (\n. d n\x)" + by (rule finite_range_imp_finch) + then have "\n. (\n. d n\x) = d n\x" + unfolding finite_chain_def by (auto simp add: maxinch_is_thelub) + with 2 show "\n. d n\x = x" by (auto elim: sym) +qed + +subsection {* ML setup *} + use "Tools/Domain/domain_take_proofs.ML" end