| author | wenzelm |
| Thu, 24 Apr 2025 23:29:57 +0200 | |
| changeset 82584 | 7ab0fb5d9919 |
| parent 81620 | 2cb49d09f059 |
| permissions | -rw-r--r-- |
| 42151 | 1 |
(* Title: HOL/HOLCF/Compact_Basis.thy |
| 25904 | 2 |
Author: Brian Huffman |
3 |
*) |
|
4 |
||
| 62175 | 5 |
section \<open>A compact basis for powerdomains\<close> |
| 25904 | 6 |
|
| 41284 | 7 |
theory Compact_Basis |
|
41289
f655912ac235
minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents:
41288
diff
changeset
|
8 |
imports Universal |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
9 |
begin |
|
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
10 |
|
| 62175 | 11 |
subsection \<open>A compact basis for powerdomains\<close> |
| 25904 | 12 |
|
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
13 |
definition "pd_basis = {S::'a::bifinite compact_basis set. finite S \<and> S \<noteq> {}}"
|
|
45694
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
42151
diff
changeset
|
14 |
|
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
15 |
typedef 'a::bifinite pd_basis = "pd_basis :: 'a compact_basis set set" |
| 81619 | 16 |
proof |
17 |
show "{a} \<in> ?pd_basis" for a
|
|
18 |
by (simp add: pd_basis_def) |
|
19 |
qed |
|
| 25904 | 20 |
|
21 |
lemma finite_Rep_pd_basis [simp]: "finite (Rep_pd_basis u)" |
|
| 81619 | 22 |
using Rep_pd_basis [of u, unfolded pd_basis_def] by simp |
| 25904 | 23 |
|
24 |
lemma Rep_pd_basis_nonempty [simp]: "Rep_pd_basis u \<noteq> {}"
|
|
| 81619 | 25 |
using Rep_pd_basis [of u, unfolded pd_basis_def] by simp |
| 25904 | 26 |
|
| 62175 | 27 |
text \<open>The powerdomain basis type is countable.\<close> |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39967
diff
changeset
|
28 |
|
| 81619 | 29 |
lemma pd_basis_countable: "\<exists>f::'a::bifinite pd_basis \<Rightarrow> nat. inj f" (is "Ex ?P") |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39967
diff
changeset
|
30 |
proof - |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39967
diff
changeset
|
31 |
obtain g :: "'a compact_basis \<Rightarrow> nat" where "inj g" |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39967
diff
changeset
|
32 |
using compact_basis.countable .. |
| 81619 | 33 |
hence image_g_eq: "g ` A = g ` B \<longleftrightarrow> A = B" for A B |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39967
diff
changeset
|
34 |
by (rule inj_image_eq_iff) |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39967
diff
changeset
|
35 |
have "inj (\<lambda>t. set_encode (g ` Rep_pd_basis t))" |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39967
diff
changeset
|
36 |
by (simp add: inj_on_def set_encode_eq image_g_eq Rep_pd_basis_inject) |
| 81619 | 37 |
thus ?thesis by (rule exI [of ?P]) |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39967
diff
changeset
|
38 |
qed |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39967
diff
changeset
|
39 |
|
| 81577 | 40 |
|
| 62175 | 41 |
subsection \<open>Unit and plus constructors\<close> |
| 25904 | 42 |
|
43 |
definition |
|
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
44 |
PDUnit :: "'a::bifinite compact_basis \<Rightarrow> 'a pd_basis" where |
| 25904 | 45 |
"PDUnit = (\<lambda>x. Abs_pd_basis {x})"
|
46 |
||
47 |
definition |
|
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
48 |
PDPlus :: "'a::bifinite pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> 'a pd_basis" where |
| 25904 | 49 |
"PDPlus t u = Abs_pd_basis (Rep_pd_basis t \<union> Rep_pd_basis u)" |
50 |
||
51 |
lemma Rep_PDUnit: |
|
52 |
"Rep_pd_basis (PDUnit x) = {x}"
|
|
53 |
unfolding PDUnit_def by (rule Abs_pd_basis_inverse) (simp add: pd_basis_def) |
|
54 |
||
55 |
lemma Rep_PDPlus: |
|
56 |
"Rep_pd_basis (PDPlus u v) = Rep_pd_basis u \<union> Rep_pd_basis v" |
|
57 |
unfolding PDPlus_def by (rule Abs_pd_basis_inverse) (simp add: pd_basis_def) |
|
58 |
||
59 |
lemma PDUnit_inject [simp]: "(PDUnit a = PDUnit b) = (a = b)" |
|
60 |
unfolding Rep_pd_basis_inject [symmetric] Rep_PDUnit by simp |
|
61 |
||
62 |
lemma PDPlus_assoc: "PDPlus (PDPlus t u) v = PDPlus t (PDPlus u v)" |
|
63 |
unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_assoc) |
|
64 |
||
65 |
lemma PDPlus_commute: "PDPlus t u = PDPlus u t" |
|
66 |
unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_commute) |
|
67 |
||
68 |
lemma PDPlus_absorb: "PDPlus t t = t" |
|
69 |
unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_absorb) |
|
70 |
||
| 81620 | 71 |
lemma pd_basis_induct1 [case_names PDUnit PDPlus]: |
| 25904 | 72 |
assumes PDUnit: "\<And>a. P (PDUnit a)" |
73 |
assumes PDPlus: "\<And>a t. P t \<Longrightarrow> P (PDPlus (PDUnit a) t)" |
|
74 |
shows "P x" |
|
| 81619 | 75 |
proof (induct x) |
76 |
case (Abs_pd_basis y) |
|
77 |
then have "finite y" and "y \<noteq> {}" by (simp_all add: pd_basis_def)
|
|
78 |
then show ?case |
|
79 |
proof (induct rule: finite_ne_induct) |
|
80 |
case (singleton x) |
|
81 |
show ?case by (rule PDUnit [unfolded PDUnit_def]) |
|
82 |
next |
|
83 |
case (insert x F) |
|
84 |
from insert(4) have "P (PDPlus (PDUnit x) (Abs_pd_basis F))" by (rule PDPlus) |
|
85 |
with insert(1,2) show ?case |
|
86 |
by (simp add: PDUnit_def PDPlus_def Abs_pd_basis_inverse [unfolded pd_basis_def]) |
|
87 |
qed |
|
88 |
qed |
|
| 25904 | 89 |
|
| 81620 | 90 |
lemma pd_basis_induct [case_names PDUnit PDPlus]: |
| 25904 | 91 |
assumes PDUnit: "\<And>a. P (PDUnit a)" |
92 |
assumes PDPlus: "\<And>t u. \<lbrakk>P t; P u\<rbrakk> \<Longrightarrow> P (PDPlus t u)" |
|
93 |
shows "P x" |
|
| 81619 | 94 |
by (induct x rule: pd_basis_induct1) (fact PDUnit, fact PDPlus [OF PDUnit]) |
| 25904 | 95 |
|
| 81577 | 96 |
|
| 62175 | 97 |
subsection \<open>Fold operator\<close> |
| 25904 | 98 |
|
99 |
definition |
|
100 |
fold_pd :: |
|
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
101 |
"('a::bifinite compact_basis \<Rightarrow> 'b::type) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a pd_basis \<Rightarrow> 'b"
|
| 51489 | 102 |
where "fold_pd g f t = semilattice_set.F f (g ` Rep_pd_basis t)" |
| 25904 | 103 |
|
| 26927 | 104 |
lemma fold_pd_PDUnit: |
| 51489 | 105 |
assumes "semilattice f" |
| 26927 | 106 |
shows "fold_pd g f (PDUnit x) = g x" |
| 51489 | 107 |
proof - |
108 |
from assms interpret semilattice_set f by (rule semilattice_set.intro) |
|
109 |
show ?thesis by (simp add: fold_pd_def Rep_PDUnit) |
|
110 |
qed |
|
| 25904 | 111 |
|
| 26927 | 112 |
lemma fold_pd_PDPlus: |
| 51489 | 113 |
assumes "semilattice f" |
| 26927 | 114 |
shows "fold_pd g f (PDPlus t u) = f (fold_pd g f t) (fold_pd g f u)" |
| 28611 | 115 |
proof - |
| 51489 | 116 |
from assms interpret semilattice_set f by (rule semilattice_set.intro) |
117 |
show ?thesis by (simp add: image_Un fold_pd_def Rep_PDPlus union) |
|
| 28611 | 118 |
qed |
| 25904 | 119 |
|
120 |
end |
|
| 51489 | 121 |