src/HOL/HOLCF/Compact_Basis.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 81620 2cb49d09f059
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41289
diff changeset
     1
(*  Title:      HOL/HOLCF/Compact_Basis.thy
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
     3
*)
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
     4
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 59797
diff changeset
     5
section \<open>A compact basis for powerdomains\<close>
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
     6
41284
6d66975b711f renamed CompactBasis.thy to Compact_Basis.thy
huffman
parents: 40774
diff changeset
     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
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 59797
diff changeset
    11
subsection \<open>A compact basis for powerdomains\<close>
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    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
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    16
proof
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    17
  show "{a} \<in> ?pd_basis" for a
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    18
    by (simp add: pd_basis_def)
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    19
qed
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    20
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    21
lemma finite_Rep_pd_basis [simp]: "finite (Rep_pd_basis u)"
81619
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    22
using Rep_pd_basis [of u, unfolded pd_basis_def] by simp
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    23
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    24
lemma Rep_pd_basis_nonempty [simp]: "Rep_pd_basis u \<noteq> {}"
81619
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    25
using Rep_pd_basis [of u, unfolded pd_basis_def] by simp
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    26
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 59797
diff changeset
    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
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    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
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    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
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    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
a712bf5ccab0 tuned whitespace;
wenzelm
parents: 62175
diff changeset
    40
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 59797
diff changeset
    41
subsection \<open>Unit and plus constructors\<close>
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    42
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    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
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    45
  "PDUnit = (\<lambda>x. Abs_pd_basis {x})"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    46
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    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
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    49
  "PDPlus t u = Abs_pd_basis (Rep_pd_basis t \<union> Rep_pd_basis u)"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    50
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    51
lemma Rep_PDUnit:
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    52
  "Rep_pd_basis (PDUnit x) = {x}"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    53
unfolding PDUnit_def by (rule Abs_pd_basis_inverse) (simp add: pd_basis_def)
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    54
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    55
lemma Rep_PDPlus:
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    56
  "Rep_pd_basis (PDPlus u v) = Rep_pd_basis u \<union> Rep_pd_basis v"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    57
unfolding PDPlus_def by (rule Abs_pd_basis_inverse) (simp add: pd_basis_def)
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    58
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    59
lemma PDUnit_inject [simp]: "(PDUnit a = PDUnit b) = (a = b)"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    60
unfolding Rep_pd_basis_inject [symmetric] Rep_PDUnit by simp
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    61
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    62
lemma PDPlus_assoc: "PDPlus (PDPlus t u) v = PDPlus t (PDPlus u v)"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    63
unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_assoc)
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    64
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    65
lemma PDPlus_commute: "PDPlus t u = PDPlus u t"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    66
unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_commute)
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    67
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    68
lemma PDPlus_absorb: "PDPlus t t = t"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    69
unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_absorb)
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    70
81620
2cb49d09f059 clarified induct rules: proper case_names;
wenzelm
parents: 81619
diff changeset
    71
lemma pd_basis_induct1 [case_names PDUnit PDPlus]:
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    72
  assumes PDUnit: "\<And>a. P (PDUnit a)"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    73
  assumes PDPlus: "\<And>a t. P t \<Longrightarrow> P (PDPlus (PDUnit a) t)"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    74
  shows "P x"
81619
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    75
proof (induct x)
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    76
  case (Abs_pd_basis y)
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    77
  then have "finite y" and "y \<noteq> {}" by (simp_all add: pd_basis_def)
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    78
  then show ?case
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    79
  proof (induct rule: finite_ne_induct)
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    80
    case (singleton x)
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    81
    show ?case by (rule PDUnit [unfolded PDUnit_def])
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    82
  next
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    83
    case (insert x F)
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    84
    from insert(4) have "P (PDPlus (PDUnit x) (Abs_pd_basis F))" by (rule PDPlus)
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    85
    with insert(1,2) show ?case
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    86
      by (simp add: PDUnit_def PDPlus_def Abs_pd_basis_inverse [unfolded pd_basis_def])
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    87
  qed
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    88
qed
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    89
81620
2cb49d09f059 clarified induct rules: proper case_names;
wenzelm
parents: 81619
diff changeset
    90
lemma pd_basis_induct [case_names PDUnit PDPlus]:
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    91
  assumes PDUnit: "\<And>a. P (PDUnit a)"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    92
  assumes PDPlus: "\<And>t u. \<lbrakk>P t; P u\<rbrakk> \<Longrightarrow> P (PDPlus t u)"
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    93
  shows "P x"
81619
0c0b2031e42e tuned proofs;
wenzelm
parents: 81583
diff changeset
    94
  by (induct x rule: pd_basis_induct1) (fact PDUnit, fact PDPlus [OF PDUnit])
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    95
81577
a712bf5ccab0 tuned whitespace;
wenzelm
parents: 62175
diff changeset
    96
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 59797
diff changeset
    97
subsection \<open>Fold operator\<close>
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    98
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
    99
definition
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
   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
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49834
diff changeset
   102
  where "fold_pd g f t = semilattice_set.F f (g ` Rep_pd_basis t)"
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
   103
26927
8684b5240f11 rename locales;
huffman
parents: 26806
diff changeset
   104
lemma fold_pd_PDUnit:
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49834
diff changeset
   105
  assumes "semilattice f"
26927
8684b5240f11 rename locales;
huffman
parents: 26806
diff changeset
   106
  shows "fold_pd g f (PDUnit x) = g x"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49834
diff changeset
   107
proof -
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49834
diff changeset
   108
  from assms interpret semilattice_set f by (rule semilattice_set.intro)
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49834
diff changeset
   109
  show ?thesis by (simp add: fold_pd_def Rep_PDUnit)
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49834
diff changeset
   110
qed
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
   111
26927
8684b5240f11 rename locales;
huffman
parents: 26806
diff changeset
   112
lemma fold_pd_PDPlus:
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49834
diff changeset
   113
  assumes "semilattice f"
26927
8684b5240f11 rename locales;
huffman
parents: 26806
diff changeset
   114
  shows "fold_pd g f (PDPlus t u) = f (fold_pd g f t) (fold_pd g f u)"
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27406
diff changeset
   115
proof -
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49834
diff changeset
   116
  from assms interpret semilattice_set f by (rule semilattice_set.intro)
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49834
diff changeset
   117
  show ?thesis by (simp add: image_Un fold_pd_def Rep_PDPlus union)
28611
983c1855a7af More occurrences of 'includes' gone.
ballarin
parents: 27406
diff changeset
   118
qed
25904
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
   119
8161f137b0e9 new theory of powerdomains
huffman
parents:
diff changeset
   120
end
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49834
diff changeset
   121