| author | huffman | 
| Thu, 18 Aug 2011 22:31:52 -0700 | |
| changeset 44291 | dbd9965745fd | 
| parent 42151 | 4da4fc77664b | 
| child 58880 | 0baae4311a9f | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/Bifinite.thy | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 2 | Author: Brian Huffman | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 3 | *) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 4 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 5 | header {* Profinite and bifinite cpos *}
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 6 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 7 | theory Bifinite | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
41370diff
changeset | 8 | imports Map_Functions "~~/src/HOL/Library/Countable" | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 9 | begin | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 10 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 11 | default_sort cpo | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 12 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 13 | subsection {* Chains of finite deflations *}
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 14 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 15 | locale approx_chain = | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 16 | fixes approx :: "nat \<Rightarrow> 'a \<rightarrow> 'a" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 17 | assumes chain_approx [simp]: "chain (\<lambda>i. approx i)" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 18 | assumes lub_approx [simp]: "(\<Squnion>i. approx i) = ID" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 19 | assumes finite_deflation_approx [simp]: "\<And>i. finite_deflation (approx i)" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 20 | begin | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 21 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 22 | lemma deflation_approx: "deflation (approx i)" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 23 | using finite_deflation_approx by (rule finite_deflation_imp_deflation) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 24 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 25 | lemma approx_idem: "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 26 | using deflation_approx by (rule deflation.idem) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 27 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 28 | lemma approx_below: "approx i\<cdot>x \<sqsubseteq> x" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 29 | using deflation_approx by (rule deflation.below) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 30 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 31 | lemma finite_range_approx: "finite (range (\<lambda>x. approx i\<cdot>x))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 32 | apply (rule finite_deflation.finite_range) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 33 | apply (rule finite_deflation_approx) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 34 | done | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 35 | |
| 41370 
17b09240893c
declare more simp rules, rewrite proofs in Isar-style
 huffman parents: 
41299diff
changeset | 36 | lemma compact_approx [simp]: "compact (approx n\<cdot>x)" | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 37 | apply (rule finite_deflation.compact) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 38 | apply (rule finite_deflation_approx) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 39 | done | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 40 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 41 | lemma compact_eq_approx: "compact x \<Longrightarrow> \<exists>i. approx i\<cdot>x = x" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 42 | by (rule admD2, simp_all) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 43 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 44 | end | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 45 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 46 | subsection {* Omega-profinite and bifinite domains *}
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 47 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 48 | class bifinite = pcpo + | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 49 | assumes bifinite: "\<exists>(a::nat \<Rightarrow> 'a \<rightarrow> 'a). approx_chain a" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 50 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 51 | class profinite = cpo + | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 52 | assumes profinite: "\<exists>(a::nat \<Rightarrow> 'a\<^sub>\<bottom> \<rightarrow> 'a\<^sub>\<bottom>). approx_chain a" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 53 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 54 | subsection {* Building approx chains *}
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 55 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 56 | lemma approx_chain_iso: | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 57 | assumes a: "approx_chain a" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 58 | assumes [simp]: "\<And>x. f\<cdot>(g\<cdot>x) = x" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 59 | assumes [simp]: "\<And>y. g\<cdot>(f\<cdot>y) = y" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 60 | shows "approx_chain (\<lambda>i. f oo a i oo g)" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 61 | proof - | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 62 | have 1: "f oo g = ID" by (simp add: cfun_eqI) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 63 | have 2: "ep_pair f g" by (simp add: ep_pair_def) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 64 | from 1 2 show ?thesis | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 65 | using a unfolding approx_chain_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 66 | by (simp add: lub_APP ep_pair.finite_deflation_e_d_p) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 67 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 68 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 69 | lemma approx_chain_u_map: | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 70 | assumes "approx_chain a" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 71 | shows "approx_chain (\<lambda>i. u_map\<cdot>(a i))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 72 | using assms unfolding approx_chain_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 73 | by (simp add: lub_APP u_map_ID finite_deflation_u_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 74 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 75 | lemma approx_chain_sfun_map: | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 76 | assumes "approx_chain a" and "approx_chain b" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 77 | shows "approx_chain (\<lambda>i. sfun_map\<cdot>(a i)\<cdot>(b i))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 78 | using assms unfolding approx_chain_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 79 | by (simp add: lub_APP sfun_map_ID finite_deflation_sfun_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 80 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 81 | lemma approx_chain_sprod_map: | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 82 | assumes "approx_chain a" and "approx_chain b" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 83 | shows "approx_chain (\<lambda>i. sprod_map\<cdot>(a i)\<cdot>(b i))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 84 | using assms unfolding approx_chain_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 85 | by (simp add: lub_APP sprod_map_ID finite_deflation_sprod_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 86 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 87 | lemma approx_chain_ssum_map: | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 88 | assumes "approx_chain a" and "approx_chain b" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 89 | shows "approx_chain (\<lambda>i. ssum_map\<cdot>(a i)\<cdot>(b i))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 90 | using assms unfolding approx_chain_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 91 | by (simp add: lub_APP ssum_map_ID finite_deflation_ssum_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 92 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 93 | lemma approx_chain_cfun_map: | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 94 | assumes "approx_chain a" and "approx_chain b" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 95 | shows "approx_chain (\<lambda>i. cfun_map\<cdot>(a i)\<cdot>(b i))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 96 | using assms unfolding approx_chain_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 97 | by (simp add: lub_APP cfun_map_ID finite_deflation_cfun_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 98 | |
| 41297 | 99 | lemma approx_chain_prod_map: | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 100 | assumes "approx_chain a" and "approx_chain b" | 
| 41297 | 101 | shows "approx_chain (\<lambda>i. prod_map\<cdot>(a i)\<cdot>(b i))" | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 102 | using assms unfolding approx_chain_def | 
| 41297 | 103 | by (simp add: lub_APP prod_map_ID finite_deflation_prod_map) | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 104 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 105 | text {* Approx chains for countable discrete types. *}
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 106 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 107 | definition discr_approx :: "nat \<Rightarrow> 'a::countable discr u \<rightarrow> 'a discr u" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 108 | where "discr_approx = (\<lambda>i. \<Lambda>(up\<cdot>x). if to_nat (undiscr x) < i then up\<cdot>x else \<bottom>)" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 109 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 110 | lemma chain_discr_approx [simp]: "chain discr_approx" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 111 | unfolding discr_approx_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 112 | by (rule chainI, simp add: monofun_cfun monofun_LAM) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 113 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 114 | lemma lub_discr_approx [simp]: "(\<Squnion>i. discr_approx i) = ID" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 115 | apply (rule cfun_eqI) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 116 | apply (simp add: contlub_cfun_fun) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 117 | apply (simp add: discr_approx_def) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 118 | apply (case_tac x, simp) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 119 | apply (rule lub_eqI) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 120 | apply (rule is_lubI) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 121 | apply (rule ub_rangeI, simp) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 122 | apply (drule ub_rangeD) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 123 | apply (erule rev_below_trans) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 124 | apply simp | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 125 | apply (rule lessI) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 126 | done | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 127 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 128 | lemma inj_on_undiscr [simp]: "inj_on undiscr A" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 129 | using Discr_undiscr by (rule inj_on_inverseI) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 130 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 131 | lemma finite_deflation_discr_approx: "finite_deflation (discr_approx i)" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 132 | proof | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 133 | fix x :: "'a discr u" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 134 | show "discr_approx i\<cdot>x \<sqsubseteq> x" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 135 | unfolding discr_approx_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 136 | by (cases x, simp, simp) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 137 | show "discr_approx i\<cdot>(discr_approx i\<cdot>x) = discr_approx i\<cdot>x" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 138 | unfolding discr_approx_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 139 | by (cases x, simp, simp) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 140 |   show "finite {x::'a discr u. discr_approx i\<cdot>x = x}"
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 141 | proof (rule finite_subset) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 142 |     let ?S = "insert (\<bottom>::'a discr u) ((\<lambda>x. up\<cdot>x) ` undiscr -` to_nat -` {..<i})"
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 143 |     show "{x::'a discr u. discr_approx i\<cdot>x = x} \<subseteq> ?S"
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 144 | unfolding discr_approx_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 145 | by (rule subsetI, case_tac x, simp, simp split: split_if_asm) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 146 | show "finite ?S" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 147 | by (simp add: finite_vimageI) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 148 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 149 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 150 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 151 | lemma discr_approx: "approx_chain discr_approx" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 152 | using chain_discr_approx lub_discr_approx finite_deflation_discr_approx | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 153 | by (rule approx_chain.intro) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 154 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 155 | subsection {* Class instance proofs *}
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 156 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 157 | instance bifinite \<subseteq> profinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 158 | proof | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 159 | show "\<exists>(a::nat \<Rightarrow> 'a\<^sub>\<bottom> \<rightarrow> 'a\<^sub>\<bottom>). approx_chain a" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 160 | using bifinite [where 'a='a] | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 161 | by (fast intro!: approx_chain_u_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 162 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 163 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 164 | instance u :: (profinite) bifinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 165 | by default (rule profinite) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 166 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 167 | text {*
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 168 |   Types @{typ "'a \<rightarrow> 'b"} and @{typ "'a u \<rightarrow>! 'b"} are isomorphic.
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 169 | *} | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 170 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 171 | definition "encode_cfun = (\<Lambda> f. sfun_abs\<cdot>(fup\<cdot>f))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 172 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 173 | definition "decode_cfun = (\<Lambda> g x. sfun_rep\<cdot>g\<cdot>(up\<cdot>x))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 174 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 175 | lemma decode_encode_cfun [simp]: "decode_cfun\<cdot>(encode_cfun\<cdot>x) = x" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 176 | unfolding encode_cfun_def decode_cfun_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 177 | by (simp add: eta_cfun) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 178 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 179 | lemma encode_decode_cfun [simp]: "encode_cfun\<cdot>(decode_cfun\<cdot>y) = y" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 180 | unfolding encode_cfun_def decode_cfun_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 181 | apply (simp add: sfun_eq_iff strictify_cancel) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 182 | apply (rule cfun_eqI, case_tac x, simp_all) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 183 | done | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 184 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 185 | instance cfun :: (profinite, bifinite) bifinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 186 | proof | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 187 | obtain a :: "nat \<Rightarrow> 'a\<^sub>\<bottom> \<rightarrow> 'a\<^sub>\<bottom>" where a: "approx_chain a" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 188 | using profinite .. | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 189 | obtain b :: "nat \<Rightarrow> 'b \<rightarrow> 'b" where b: "approx_chain b" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 190 | using bifinite .. | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 191 | have "approx_chain (\<lambda>i. decode_cfun oo sfun_map\<cdot>(a i)\<cdot>(b i) oo encode_cfun)" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 192 | using a b by (simp add: approx_chain_iso approx_chain_sfun_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 193 |   thus "\<exists>(a::nat \<Rightarrow> ('a \<rightarrow> 'b) \<rightarrow> ('a \<rightarrow> 'b)). approx_chain a"
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 194 | by - (rule exI) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 195 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 196 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 197 | text {*
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 198 |   Types @{typ "('a * 'b) u"} and @{typ "'a u \<otimes> 'b u"} are isomorphic.
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 199 | *} | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 200 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 201 | definition "encode_prod_u = (\<Lambda>(up\<cdot>(x, y)). (:up\<cdot>x, up\<cdot>y:))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 202 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 203 | definition "decode_prod_u = (\<Lambda>(:up\<cdot>x, up\<cdot>y:). up\<cdot>(x, y))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 204 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 205 | lemma decode_encode_prod_u [simp]: "decode_prod_u\<cdot>(encode_prod_u\<cdot>x) = x" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 206 | unfolding encode_prod_u_def decode_prod_u_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 207 | by (case_tac x, simp, rename_tac y, case_tac y, simp) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 208 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 209 | lemma encode_decode_prod_u [simp]: "encode_prod_u\<cdot>(decode_prod_u\<cdot>y) = y" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 210 | unfolding encode_prod_u_def decode_prod_u_def | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 211 | apply (case_tac y, simp, rename_tac a b) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 212 | apply (case_tac a, simp, case_tac b, simp, simp) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 213 | done | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 214 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 215 | instance prod :: (profinite, profinite) profinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 216 | proof | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 217 | obtain a :: "nat \<Rightarrow> 'a\<^sub>\<bottom> \<rightarrow> 'a\<^sub>\<bottom>" where a: "approx_chain a" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 218 | using profinite .. | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 219 | obtain b :: "nat \<Rightarrow> 'b\<^sub>\<bottom> \<rightarrow> 'b\<^sub>\<bottom>" where b: "approx_chain b" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 220 | using profinite .. | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 221 | have "approx_chain (\<lambda>i. decode_prod_u oo sprod_map\<cdot>(a i)\<cdot>(b i) oo encode_prod_u)" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 222 | using a b by (simp add: approx_chain_iso approx_chain_sprod_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 223 |   thus "\<exists>(a::nat \<Rightarrow> ('a \<times> 'b)\<^sub>\<bottom> \<rightarrow> ('a \<times> 'b)\<^sub>\<bottom>). approx_chain a"
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 224 | by - (rule exI) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 225 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 226 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 227 | instance prod :: (bifinite, bifinite) bifinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 228 | proof | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 229 |   show "\<exists>(a::nat \<Rightarrow> ('a \<times> 'b) \<rightarrow> ('a \<times> 'b)). approx_chain a"
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 230 | using bifinite [where 'a='a] and bifinite [where 'a='b] | 
| 41297 | 231 | by (fast intro!: approx_chain_prod_map) | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 232 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 233 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 234 | instance sfun :: (bifinite, bifinite) bifinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 235 | proof | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 236 |   show "\<exists>(a::nat \<Rightarrow> ('a \<rightarrow>! 'b) \<rightarrow> ('a \<rightarrow>! 'b)). approx_chain a"
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 237 | using bifinite [where 'a='a] and bifinite [where 'a='b] | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 238 | by (fast intro!: approx_chain_sfun_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 239 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 240 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 241 | instance sprod :: (bifinite, bifinite) bifinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 242 | proof | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 243 |   show "\<exists>(a::nat \<Rightarrow> ('a \<otimes> 'b) \<rightarrow> ('a \<otimes> 'b)). approx_chain a"
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 244 | using bifinite [where 'a='a] and bifinite [where 'a='b] | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 245 | by (fast intro!: approx_chain_sprod_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 246 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 247 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 248 | instance ssum :: (bifinite, bifinite) bifinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 249 | proof | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 250 |   show "\<exists>(a::nat \<Rightarrow> ('a \<oplus> 'b) \<rightarrow> ('a \<oplus> 'b)). approx_chain a"
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 251 | using bifinite [where 'a='a] and bifinite [where 'a='b] | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 252 | by (fast intro!: approx_chain_ssum_map) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 253 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 254 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 255 | lemma approx_chain_unit: "approx_chain (\<bottom> :: nat \<Rightarrow> unit \<rightarrow> unit)" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41413diff
changeset | 256 | by (simp add: approx_chain_def cfun_eq_iff finite_deflation_bottom) | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 257 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 258 | instance unit :: bifinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 259 | by default (fast intro!: approx_chain_unit) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 260 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 261 | instance discr :: (countable) profinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 262 | by default (fast intro!: discr_approx) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 263 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 264 | instance lift :: (countable) bifinite | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 265 | proof | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 266 | note [simp] = cont_Abs_lift cont_Rep_lift Rep_lift_inverse Abs_lift_inverse | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 267 |   obtain a :: "nat \<Rightarrow> ('a discr)\<^sub>\<bottom> \<rightarrow> ('a discr)\<^sub>\<bottom>" where a: "approx_chain a"
 | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 268 | using profinite .. | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 269 | hence "approx_chain (\<lambda>i. (\<Lambda> y. Abs_lift y) oo a i oo (\<Lambda> x. Rep_lift x))" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 270 | by (rule approx_chain_iso) simp_all | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 271 | thus "\<exists>(a::nat \<Rightarrow> 'a lift \<rightarrow> 'a lift). approx_chain a" | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 272 | by - (rule exI) | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 273 | qed | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 274 | |
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: diff
changeset | 275 | end |