# HG changeset patch # User huffman # Date 1220541858 -7200 # Node ID 218252dfd81e4d7dfb85bcbd0f3d69cefb90d9d7 # Parent 236e07d8821ef8dcc26096732b7744f21aeedb9e reorganize subsections diff -r 236e07d8821e -r 218252dfd81e src/HOLCF/Completion.thy --- a/src/HOLCF/Completion.thy Thu Sep 04 17:21:49 2008 +0200 +++ b/src/HOLCF/Completion.thy Thu Sep 04 17:24:18 2008 +0200 @@ -163,7 +163,7 @@ apply (erule (1) trans_less) done -subsection {* Defining functions in terms of basis elements *} +subsection {* Lemmas about least upper bounds *} lemma finite_directed_contains_lub: "\finite S; directed S\ \ \u\S. S <<| u" @@ -191,6 +191,8 @@ lemma is_lub_thelub0: "\\u. S <<| u; S <| x\ \ lub S \ x" by (erule exE, drule lubI, erule is_lub_lub) +subsection {* Locale for ideal completion *} + locale basis_take = preorder + fixes take :: "nat \ 'a::type \ 'a" assumes take_less: "take n a \ a" @@ -221,6 +223,61 @@ lemma finite_take_rep: "finite (take n ` rep x)" by (rule finite_subset [OF image_mono [OF subset_UNIV] finite_range_take]) +lemma rep_mono: "x \ y \ rep x \ rep y" +apply (frule bin_chain) +apply (drule rep_contlub) +apply (simp only: thelubI [OF lub_bin_chain]) +apply (rule subsetI, rule UN_I [where a=0], simp_all) +done + +lemma less_def: "x \ y \ rep x \ rep y" +by (rule iffI [OF rep_mono subset_repD]) + +lemma rep_eq: "rep x = {a. principal a \ x}" +unfolding less_def rep_principal +apply safe +apply (erule (1) idealD3 [OF ideal_rep]) +apply (erule subsetD, simp add: r_refl) +done + +lemma mem_rep_iff_principal_less: "a \ rep x \ principal a \ x" +by (simp add: rep_eq) + +lemma principal_less_iff_mem_rep: "principal a \ x \ a \ rep x" +by (simp add: rep_eq) + +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" +unfolding po_eq_conv [where 'a='b] principal_less_iff .. + +lemma repD: "a \ rep x \ principal a \ x" +by (simp add: rep_eq) + +lemma principal_mono: "a \ b \ principal a \ principal b" +by (simp only: principal_less_iff) + +lemma lessI: "(\a. principal a \ x \ principal a \ u) \ x \ u" +unfolding principal_less_iff_mem_rep +by (simp add: less_def subset_eq) + +lemma lub_principal_rep: "principal ` rep x <<| x" +apply (rule is_lubI) +apply (rule ub_imageI) +apply (erule repD) +apply (subst less_def) +apply (rule subsetI) +apply (drule (1) ub_imageD) +apply (simp add: rep_eq) +done + +subsection {* Defining functions in terms of basis elements *} + +definition + basis_fun :: "('a::type \ 'c::cpo) \ 'b \ 'c" where + "basis_fun = (\f. (\ x. lub (f ` rep x)))" + lemma basis_fun_lemma0: fixes f :: "'a::type \ 'c::cpo" assumes f_mono: "\a b. a \ b \ f a \ f b" @@ -278,59 +335,6 @@ shows "\u. f ` rep x <<| u" by (rule exI, rule basis_fun_lemma2, erule f_mono) -lemma rep_mono: "x \ y \ rep x \ rep y" -apply (frule bin_chain) -apply (drule rep_contlub) -apply (simp only: thelubI [OF lub_bin_chain]) -apply (rule subsetI, rule UN_I [where a=0], simp_all) -done - -lemma less_def: "x \ y \ rep x \ rep y" -by (rule iffI [OF rep_mono subset_repD]) - -lemma rep_eq: "rep x = {a. principal a \ x}" -unfolding less_def rep_principal -apply safe -apply (erule (1) idealD3 [OF ideal_rep]) -apply (erule subsetD, simp add: r_refl) -done - -lemma mem_rep_iff_principal_less: "a \ rep x \ principal a \ x" -by (simp add: rep_eq) - -lemma principal_less_iff_mem_rep: "principal a \ x \ a \ rep x" -by (simp add: rep_eq) - -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" -unfolding po_eq_conv [where 'a='b] principal_less_iff .. - -lemma repD: "a \ rep x \ principal a \ x" -by (simp add: rep_eq) - -lemma principal_mono: "a \ b \ principal a \ principal b" -by (simp only: principal_less_iff) - -lemma lessI: "(\a. principal a \ x \ principal a \ u) \ x \ u" -unfolding principal_less_iff_mem_rep -by (simp add: less_def subset_eq) - -lemma lub_principal_rep: "principal ` rep x <<| x" -apply (rule is_lubI) -apply (rule ub_imageI) -apply (erule repD) -apply (subst less_def) -apply (rule subsetI) -apply (drule (1) ub_imageD) -apply (simp add: rep_eq) -done - -definition - basis_fun :: "('a::type \ 'c::cpo) \ 'b \ 'c" where - "basis_fun = (\f. (\ x. lub (f ` rep x)))" - lemma basis_fun_beta: fixes f :: "'a::type \ 'c::cpo" assumes f_mono: "\a b. a \ b \ f a \ f b" @@ -380,6 +384,8 @@ lemma compact_principal [simp]: "compact (principal a)" by (rule compactI2, simp add: principal_less_iff_mem_rep rep_contlub) +subsection {* Bifiniteness of ideal completions *} + definition completion_approx :: "nat \ 'b \ 'b" where "completion_approx = (\i. basis_fun (\a. principal (take i a)))"