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