# HG changeset patch # User huffman # Date 1211233806 -7200 # Node ID 290d23780204227a914eb4dcb2499f115bbbfbb2 # Parent c8b20f615d6cebc20fd79c66f7d821f74fa90a2f instantiation lift :: (countable) bifinite diff -r c8b20f615d6c -r 290d23780204 src/HOLCF/Lift.thy --- a/src/HOLCF/Lift.thy Mon May 19 23:49:20 2008 +0200 +++ b/src/HOLCF/Lift.thy Mon May 19 23:50:06 2008 +0200 @@ -6,7 +6,7 @@ header {* Lifting types of class type to flat pcpo's *} theory Lift -imports Discrete Up Cprod +imports Discrete Up Cprod Countable begin defaultsort type @@ -204,4 +204,51 @@ end; *} +subsection {* Lifted countable types are bifinite *} + +instantiation lift :: (countable) bifinite +begin + +definition + approx_lift_def: + "approx = (\n. FLIFT x. if to_nat x < n then Def x else \)" + +instance proof + fix x :: "'a lift" + show "chain (\i. approx i\x)" + unfolding approx_lift_def + by (rule chainI, cases x, simp_all) +next + fix x :: "'a lift" + show "(\i. approx i\x) = x" + unfolding approx_lift_def + apply (cases x, simp) + apply (rule thelubI) + apply (rule is_lubI) + apply (rule ub_rangeI, simp) + apply (drule ub_rangeD) + apply (erule rev_trans_less) + apply simp + apply (rule lessI) + done +next + fix i :: nat and x :: "'a lift" + show "approx i\(approx i\x) = approx i\x" + unfolding approx_lift_def + by (cases x, simp, simp) +next + fix i :: nat + show "finite {x::'a lift. approx i\x = x}" + proof (rule finite_subset) + let ?S = "insert (\::'a lift) (Def ` to_nat -` {..x = x} \ ?S" + unfolding approx_lift_def + by (rule subsetI, case_tac x, simp, simp split: split_if_asm) + show "finite ?S" + by (simp add: finite_vimageI) + qed +qed + end + +end