| author | wenzelm | 
| Sat, 17 Aug 2019 11:23:20 +0200 | |
| changeset 70557 | 5d6e9c65ea67 | 
| parent 62175 | 8ffc4d0e652d | 
| child 81577 | a712bf5ccab0 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/HOLCF/Algebraic.thy | 
| 27409 | 2 | Author: Brian Huffman | 
| 3 | *) | |
| 4 | ||
| 62175 | 5 | section \<open>Algebraic deflations\<close> | 
| 27409 | 6 | |
| 7 | theory Algebraic | |
| 40502 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: 
40327diff
changeset | 8 | imports Universal Map_Functions | 
| 27409 | 9 | begin | 
| 10 | ||
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 11 | default_sort bifinite | 
| 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 12 | |
| 62175 | 13 | subsection \<open>Type constructor for finite deflations\<close> | 
| 27409 | 14 | |
| 49834 | 15 | typedef 'a fin_defl = "{d::'a \<rightarrow> 'a. finite_deflation d}"
 | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41394diff
changeset | 16 | by (fast intro: finite_deflation_bottom) | 
| 27409 | 17 | |
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 18 | instantiation fin_defl :: (bifinite) below | 
| 27409 | 19 | begin | 
| 20 | ||
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 21 | definition below_fin_defl_def: | 
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 22 | "below \<equiv> \<lambda>x y. Rep_fin_defl x \<sqsubseteq> Rep_fin_defl y" | 
| 27409 | 23 | |
| 24 | instance .. | |
| 25 | end | |
| 26 | ||
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 27 | instance fin_defl :: (bifinite) po | 
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 28 | using type_definition_fin_defl below_fin_defl_def | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 29 | by (rule typedef_po) | 
| 27409 | 30 | |
| 31 | lemma finite_deflation_Rep_fin_defl: "finite_deflation (Rep_fin_defl d)" | |
| 32 | using Rep_fin_defl by simp | |
| 33 | ||
| 31164 
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
 huffman parents: 
31076diff
changeset | 34 | lemma deflation_Rep_fin_defl: "deflation (Rep_fin_defl d)" | 
| 
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
 huffman parents: 
31076diff
changeset | 35 | using finite_deflation_Rep_fin_defl | 
| 
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
 huffman parents: 
31076diff
changeset | 36 | by (rule finite_deflation_imp_deflation) | 
| 
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
 huffman parents: 
31076diff
changeset | 37 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29252diff
changeset | 38 | interpretation Rep_fin_defl: finite_deflation "Rep_fin_defl d" | 
| 27409 | 39 | by (rule finite_deflation_Rep_fin_defl) | 
| 40 | ||
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 41 | lemma fin_defl_belowI: | 
| 27409 | 42 | "(\<And>x. Rep_fin_defl a\<cdot>x = x \<Longrightarrow> Rep_fin_defl b\<cdot>x = x) \<Longrightarrow> a \<sqsubseteq> b" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 43 | unfolding below_fin_defl_def | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 44 | by (rule Rep_fin_defl.belowI) | 
| 27409 | 45 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 46 | lemma fin_defl_belowD: | 
| 27409 | 47 | "\<lbrakk>a \<sqsubseteq> b; Rep_fin_defl a\<cdot>x = x\<rbrakk> \<Longrightarrow> Rep_fin_defl b\<cdot>x = x" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 48 | unfolding below_fin_defl_def | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 49 | by (rule Rep_fin_defl.belowD) | 
| 27409 | 50 | |
| 51 | lemma fin_defl_eqI: | |
| 52 | "(\<And>x. Rep_fin_defl a\<cdot>x = x \<longleftrightarrow> Rep_fin_defl b\<cdot>x = x) \<Longrightarrow> a = b" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 53 | apply (rule below_antisym) | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 54 | apply (rule fin_defl_belowI, simp) | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 55 | apply (rule fin_defl_belowI, simp) | 
| 27409 | 56 | done | 
| 57 | ||
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 58 | lemma Rep_fin_defl_mono: "a \<sqsubseteq> b \<Longrightarrow> Rep_fin_defl a \<sqsubseteq> Rep_fin_defl b" | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 59 | unfolding below_fin_defl_def . | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 60 | |
| 27409 | 61 | lemma Abs_fin_defl_mono: | 
| 62 | "\<lbrakk>finite_deflation a; finite_deflation b; a \<sqsubseteq> b\<rbrakk> | |
| 63 | \<Longrightarrow> Abs_fin_defl a \<sqsubseteq> Abs_fin_defl b" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 64 | unfolding below_fin_defl_def | 
| 27409 | 65 | by (simp add: Abs_fin_defl_inverse) | 
| 66 | ||
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 67 | lemma (in finite_deflation) compact_belowI: | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 68 | assumes "\<And>x. compact x \<Longrightarrow> d\<cdot>x = x \<Longrightarrow> f\<cdot>x = x" shows "d \<sqsubseteq> f" | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 69 | by (rule belowI, rule assms, erule subst, rule compact) | 
| 27409 | 70 | |
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 71 | lemma compact_Rep_fin_defl [simp]: "compact (Rep_fin_defl a)" | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 72 | using finite_deflation_Rep_fin_defl | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 73 | by (rule finite_deflation_imp_compact) | 
| 33586 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 74 | |
| 62175 | 75 | subsection \<open>Defining algebraic deflations by ideal completion\<close> | 
| 27409 | 76 | |
| 49834 | 77 | typedef 'a defl = "{S::'a fin_defl set. below.ideal S}"
 | 
| 40888 
28cd51cff70c
reformulate lemma preorder.ex_ideal, and use it for typedefs
 huffman parents: 
40774diff
changeset | 78 | by (rule below.ex_ideal) | 
| 27409 | 79 | |
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 80 | instantiation defl :: (bifinite) below | 
| 27409 | 81 | begin | 
| 82 | ||
| 83 | definition | |
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 84 | "x \<sqsubseteq> y \<longleftrightarrow> Rep_defl x \<subseteq> Rep_defl y" | 
| 27409 | 85 | |
| 86 | instance .. | |
| 87 | end | |
| 88 | ||
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 89 | instance defl :: (bifinite) po | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 90 | using type_definition_defl below_defl_def | 
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 91 | by (rule below.typedef_ideal_po) | 
| 27409 | 92 | |
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 93 | instance defl :: (bifinite) cpo | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 94 | using type_definition_defl below_defl_def | 
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 95 | by (rule below.typedef_ideal_cpo) | 
| 27409 | 96 | |
| 97 | definition | |
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 98 | defl_principal :: "'a fin_defl \<Rightarrow> 'a defl" where | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 99 |   "defl_principal t = Abs_defl {u. u \<sqsubseteq> t}"
 | 
| 27409 | 100 | |
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 101 | lemma fin_defl_countable: "\<exists>f::'a fin_defl \<Rightarrow> nat. inj f" | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: 
40888diff
changeset | 102 | proof - | 
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 103 | obtain f :: "'a compact_basis \<Rightarrow> nat" where inj_f: "inj f" | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: 
40888diff
changeset | 104 | using compact_basis.countable .. | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: 
40888diff
changeset | 105 |   have *: "\<And>d. finite (f ` Rep_compact_basis -` {x. Rep_fin_defl d\<cdot>x = x})"
 | 
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 106 | apply (rule finite_imageI) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 107 | apply (rule finite_vimageI) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 108 | apply (rule Rep_fin_defl.finite_fixes) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 109 | apply (simp add: inj_on_def Rep_compact_basis_inject) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 110 | done | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 111 |   have range_eq: "range Rep_compact_basis = {x. compact x}"
 | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 112 | using type_definition_compact_basis by (rule type_definition.Rep_range) | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: 
40888diff
changeset | 113 | have "inj (\<lambda>d. set_encode | 
| 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: 
40888diff
changeset | 114 |     (f ` Rep_compact_basis -` {x. Rep_fin_defl d\<cdot>x = x}))"
 | 
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 115 | apply (rule inj_onI) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 116 | apply (simp only: set_encode_eq *) | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: 
40888diff
changeset | 117 | apply (simp only: inj_image_eq_iff inj_f) | 
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 118 | apply (drule_tac f="image Rep_compact_basis" in arg_cong) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 119 | apply (simp del: vimage_Collect_eq add: range_eq set_eq_iff) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 120 | apply (rule Rep_fin_defl_inject [THEN iffD1]) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 121 | apply (rule below_antisym) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 122 | apply (rule Rep_fin_defl.compact_belowI, rename_tac z) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 123 | apply (drule_tac x=z in spec, simp) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 124 | apply (rule Rep_fin_defl.compact_belowI, rename_tac z) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 125 | apply (drule_tac x=z in spec, simp) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 126 | done | 
| 41286 
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
 huffman parents: 
40888diff
changeset | 127 | thus ?thesis by - (rule exI) | 
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 128 | qed | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 129 | |
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 130 | interpretation defl: ideal_completion below defl_principal Rep_defl | 
| 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 131 | using type_definition_defl below_defl_def | 
| 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 132 | using defl_principal_def fin_defl_countable | 
| 39984 | 133 | by (rule below.typedef_ideal_completion) | 
| 27409 | 134 | |
| 62175 | 135 | text \<open>Algebraic deflations are pointed\<close> | 
| 27409 | 136 | |
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 137 | lemma defl_minimal: "defl_principal (Abs_fin_defl \<bottom>) \<sqsubseteq> x" | 
| 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 138 | apply (induct x rule: defl.principal_induct, simp) | 
| 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 139 | apply (rule defl.principal_mono) | 
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 140 | apply (simp add: below_fin_defl_def) | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41394diff
changeset | 141 | apply (simp add: Abs_fin_defl_inverse finite_deflation_bottom) | 
| 27409 | 142 | done | 
| 143 | ||
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 144 | instance defl :: (bifinite) pcpo | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 145 | by intro_classes (fast intro: defl_minimal) | 
| 27409 | 146 | |
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 147 | lemma inst_defl_pcpo: "\<bottom> = defl_principal (Abs_fin_defl \<bottom>)" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41394diff
changeset | 148 | by (rule defl_minimal [THEN bottomI, symmetric]) | 
| 27409 | 149 | |
| 62175 | 150 | subsection \<open>Applying algebraic deflations\<close> | 
| 27409 | 151 | |
| 152 | definition | |
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41286diff
changeset | 153 | cast :: "'a defl \<rightarrow> 'a \<rightarrow> 'a" | 
| 27409 | 154 | where | 
| 41394 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 huffman parents: 
41290diff
changeset | 155 | "cast = defl.extension Rep_fin_defl" | 
| 27409 | 156 | |
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 157 | lemma cast_defl_principal: | 
| 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 158 | "cast\<cdot>(defl_principal a) = Rep_fin_defl a" | 
| 27409 | 159 | unfolding cast_def | 
| 41394 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 huffman parents: 
41290diff
changeset | 160 | apply (rule defl.extension_principal) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30729diff
changeset | 161 | apply (simp only: below_fin_defl_def) | 
| 27409 | 162 | done | 
| 163 | ||
| 164 | lemma deflation_cast: "deflation (cast\<cdot>d)" | |
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 165 | apply (induct d rule: defl.principal_induct) | 
| 27409 | 166 | apply (rule adm_subst [OF _ adm_deflation], simp) | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 167 | apply (simp add: cast_defl_principal) | 
| 27409 | 168 | apply (rule finite_deflation_imp_deflation) | 
| 169 | apply (rule finite_deflation_Rep_fin_defl) | |
| 170 | done | |
| 171 | ||
| 172 | lemma finite_deflation_cast: | |
| 173 | "compact d \<Longrightarrow> finite_deflation (cast\<cdot>d)" | |
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 174 | apply (drule defl.compact_imp_principal, clarify) | 
| 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 175 | apply (simp add: cast_defl_principal) | 
| 27409 | 176 | apply (rule finite_deflation_Rep_fin_defl) | 
| 177 | done | |
| 178 | ||
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29252diff
changeset | 179 | interpretation cast: deflation "cast\<cdot>d" | 
| 27409 | 180 | by (rule deflation_cast) | 
| 181 | ||
| 33586 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 182 | declare cast.idem [simp] | 
| 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 183 | |
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 184 | lemma compact_cast [simp]: "compact d \<Longrightarrow> compact (cast\<cdot>d)" | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 185 | apply (rule finite_deflation_imp_compact) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 186 | apply (erule finite_deflation_cast) | 
| 31164 
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
 huffman parents: 
31076diff
changeset | 187 | done | 
| 
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
 huffman parents: 
31076diff
changeset | 188 | |
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 189 | lemma cast_below_cast: "cast\<cdot>A \<sqsubseteq> cast\<cdot>B \<longleftrightarrow> A \<sqsubseteq> B" | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 190 | apply (induct A rule: defl.principal_induct, simp) | 
| 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 191 | apply (induct B rule: defl.principal_induct, simp) | 
| 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 192 | apply (simp add: cast_defl_principal below_fin_defl_def) | 
| 31164 
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
 huffman parents: 
31076diff
changeset | 193 | done | 
| 
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
 huffman parents: 
31076diff
changeset | 194 | |
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 195 | lemma compact_cast_iff: "compact (cast\<cdot>d) \<longleftrightarrow> compact d" | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 196 | apply (rule iffI) | 
| 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 197 | apply (simp only: compact_def cast_below_cast [symmetric]) | 
| 40327 | 198 | apply (erule adm_subst [OF cont_Rep_cfun2]) | 
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 199 | apply (erule compact_cast) | 
| 39972 
4244ff4f9649
add lemmas finite_deflation_imp_compact, cast_below_cast_iff
 Brian Huffman <brianh@cs.pdx.edu> parents: 
39971diff
changeset | 200 | done | 
| 
4244ff4f9649
add lemmas finite_deflation_imp_compact, cast_below_cast_iff
 Brian Huffman <brianh@cs.pdx.edu> parents: 
39971diff
changeset | 201 | |
| 31164 
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
 huffman parents: 
31076diff
changeset | 202 | lemma cast_below_imp_below: "cast\<cdot>A \<sqsubseteq> cast\<cdot>B \<Longrightarrow> A \<sqsubseteq> B" | 
| 39974 
b525988432e9
major reorganization/simplification of HOLCF type classes:
 huffman parents: 
39972diff
changeset | 203 | by (simp only: cast_below_cast) | 
| 31164 
f550c4cf3f3a
continuity proofs for approx function on deflations; lemma cast_below_imp_below
 huffman parents: 
31076diff
changeset | 204 | |
| 33586 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 205 | lemma cast_eq_imp_eq: "cast\<cdot>A = cast\<cdot>B \<Longrightarrow> A = B" | 
| 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 206 | by (simp add: below_antisym cast_below_imp_below) | 
| 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 207 | |
| 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 208 | lemma cast_strict1 [simp]: "cast\<cdot>\<bottom> = \<bottom>" | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 209 | apply (subst inst_defl_pcpo) | 
| 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39985diff
changeset | 210 | apply (subst cast_defl_principal) | 
| 33586 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 211 | apply (rule Abs_fin_defl_inverse) | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41394diff
changeset | 212 | apply (simp add: finite_deflation_bottom) | 
| 33586 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 213 | done | 
| 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 214 | |
| 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 215 | lemma cast_strict2 [simp]: "cast\<cdot>A\<cdot>\<bottom> = \<bottom>" | 
| 41430 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 huffman parents: 
41394diff
changeset | 216 | by (rule cast.below [THEN bottomI]) | 
| 33586 
0e745228d605
add in_deflation relation, more lemmas about cast
 huffman parents: 
31164diff
changeset | 217 | |
| 62175 | 218 | subsection \<open>Deflation combinators\<close> | 
| 41290 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 219 | |
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 220 | definition | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 221 | "defl_fun1 e p f = | 
| 41394 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 huffman parents: 
41290diff
changeset | 222 | defl.extension (\<lambda>a. | 
| 41290 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 223 | defl_principal (Abs_fin_defl | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 224 | (e oo f\<cdot>(Rep_fin_defl a) oo p)))" | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 225 | |
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 226 | definition | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 227 | "defl_fun2 e p f = | 
| 41394 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 huffman parents: 
41290diff
changeset | 228 | defl.extension (\<lambda>a. | 
| 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 huffman parents: 
41290diff
changeset | 229 | defl.extension (\<lambda>b. | 
| 41290 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 230 | defl_principal (Abs_fin_defl | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 231 | (e oo f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo p))))" | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 232 | |
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 233 | lemma cast_defl_fun1: | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 234 | assumes ep: "ep_pair e p" | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 235 | assumes f: "\<And>a. finite_deflation a \<Longrightarrow> finite_deflation (f\<cdot>a)" | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 236 | shows "cast\<cdot>(defl_fun1 e p f\<cdot>A) = e oo f\<cdot>(cast\<cdot>A) oo p" | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 237 | proof - | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 238 | have 1: "\<And>a. finite_deflation (e oo f\<cdot>(Rep_fin_defl a) oo p)" | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 239 | apply (rule ep_pair.finite_deflation_e_d_p [OF ep]) | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 240 | apply (rule f, rule finite_deflation_Rep_fin_defl) | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 241 | done | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 242 | show ?thesis | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 243 | by (induct A rule: defl.principal_induct, simp) | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 244 | (simp only: defl_fun1_def | 
| 41394 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 huffman parents: 
41290diff
changeset | 245 | defl.extension_principal | 
| 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 huffman parents: 
41290diff
changeset | 246 | defl.extension_mono | 
| 41290 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 247 | defl.principal_mono | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 248 | Abs_fin_defl_mono [OF 1 1] | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 249 | monofun_cfun below_refl | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 250 | Rep_fin_defl_mono | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 251 | cast_defl_principal | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 252 | Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1]) | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 253 | qed | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 254 | |
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 255 | lemma cast_defl_fun2: | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 256 | assumes ep: "ep_pair e p" | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 257 | assumes f: "\<And>a b. finite_deflation a \<Longrightarrow> finite_deflation b \<Longrightarrow> | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 258 | finite_deflation (f\<cdot>a\<cdot>b)" | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 259 | shows "cast\<cdot>(defl_fun2 e p f\<cdot>A\<cdot>B) = e oo f\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo p" | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 260 | proof - | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 261 | have 1: "\<And>a b. finite_deflation | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 262 | (e oo f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo p)" | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 263 | apply (rule ep_pair.finite_deflation_e_d_p [OF ep]) | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 264 | apply (rule f, (rule finite_deflation_Rep_fin_defl)+) | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 265 | done | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 266 | show ?thesis | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 267 | apply (induct A rule: defl.principal_induct, simp) | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 268 | apply (induct B rule: defl.principal_induct, simp) | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 269 | by (simp only: defl_fun2_def | 
| 41394 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 huffman parents: 
41290diff
changeset | 270 | defl.extension_principal | 
| 
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
 huffman parents: 
41290diff
changeset | 271 | defl.extension_mono | 
| 41290 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 272 | defl.principal_mono | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 273 | Abs_fin_defl_mono [OF 1 1] | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 274 | monofun_cfun below_refl | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 275 | Rep_fin_defl_mono | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 276 | cast_defl_principal | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 277 | Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1]) | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 278 | qed | 
| 
e9c9577d88b5
replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
 huffman parents: 
41287diff
changeset | 279 | |
| 27409 | 280 | end |