| author | haftmann | 
| Wed, 10 Aug 2022 18:26:22 +0000 | |
| changeset 75800 | a21debbc7074 | 
| parent 62175 | 8ffc4d0e652d | 
| permissions | -rw-r--r-- | 
| 42151 | 1  | 
(* Title: HOL/HOLCF/ex/Powerdomain_ex.thy  | 
| 29992 | 2  | 
Author: Brian Huffman  | 
3  | 
*)  | 
|
4  | 
||
| 62175 | 5  | 
section \<open>Powerdomain examples\<close>  | 
| 29992 | 6  | 
|
7  | 
theory Powerdomain_ex  | 
|
8  | 
imports HOLCF  | 
|
9  | 
begin  | 
|
10  | 
||
| 62175 | 11  | 
subsection \<open>Monadic sorting example\<close>  | 
| 29992 | 12  | 
|
13  | 
domain ordering = LT | EQ | GT  | 
|
14  | 
||
15  | 
definition  | 
|
16  | 
compare :: "int lift \<rightarrow> int lift \<rightarrow> ordering" where  | 
|
17  | 
"compare = (FLIFT x y. if x < y then LT else if x = y then EQ else GT)"  | 
|
18  | 
||
19  | 
definition  | 
|
20  | 
is_le :: "int lift \<rightarrow> int lift \<rightarrow> tr" where  | 
|
21  | 
"is_le = (\<Lambda> x y. case compare\<cdot>x\<cdot>y of LT \<Rightarrow> TT | EQ \<Rightarrow> TT | GT \<Rightarrow> FF)"  | 
|
22  | 
||
23  | 
definition  | 
|
24  | 
is_less :: "int lift \<rightarrow> int lift \<rightarrow> tr" where  | 
|
25  | 
"is_less = (\<Lambda> x y. case compare\<cdot>x\<cdot>y of LT \<Rightarrow> TT | EQ \<Rightarrow> FF | GT \<Rightarrow> FF)"  | 
|
26  | 
||
27  | 
definition  | 
|
28  | 
r1 :: "(int lift \<times> 'a) \<rightarrow> (int lift \<times> 'a) \<rightarrow> tr convex_pd" where  | 
|
| 35918 | 29  | 
"r1 = (\<Lambda> (x,_) (y,_). case compare\<cdot>x\<cdot>y of  | 
| 29992 | 30  | 
          LT \<Rightarrow> {TT}\<natural> |
 | 
31  | 
          EQ \<Rightarrow> {TT, FF}\<natural> |
 | 
|
32  | 
          GT \<Rightarrow> {FF}\<natural>)"
 | 
|
33  | 
||
34  | 
definition  | 
|
35  | 
r2 :: "(int lift \<times> 'a) \<rightarrow> (int lift \<times> 'a) \<rightarrow> tr convex_pd" where  | 
|
| 35918 | 36  | 
  "r2 = (\<Lambda> (x,_) (y,_). {is_le\<cdot>x\<cdot>y, is_less\<cdot>x\<cdot>y}\<natural>)"
 | 
| 29992 | 37  | 
|
| 35918 | 38  | 
lemma r1_r2: "r1\<cdot>(x,a)\<cdot>(y,b) = (r2\<cdot>(x,a)\<cdot>(y,b) :: tr convex_pd)"  | 
| 29992 | 39  | 
apply (simp add: r1_def r2_def)  | 
40  | 
apply (simp add: is_le_def is_less_def)  | 
|
| 
35781
 
b7738ab762b1
renamed some lemmas generated by the domain package
 
huffman 
parents: 
35769 
diff
changeset
 | 
41  | 
apply (cases "compare\<cdot>x\<cdot>y")  | 
| 29992 | 42  | 
apply simp_all  | 
43  | 
done  | 
|
44  | 
||
45  | 
||
| 62175 | 46  | 
subsection \<open>Picking a leaf from a tree\<close>  | 
| 29992 | 47  | 
|
48  | 
domain 'a tree =  | 
|
49  | 
Node (lazy "'a tree") (lazy "'a tree") |  | 
|
50  | 
Leaf (lazy "'a")  | 
|
51  | 
||
| 
30158
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
52  | 
fixrec  | 
| 29992 | 53  | 
mirror :: "'a tree \<rightarrow> 'a tree"  | 
| 
30158
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
54  | 
where  | 
| 29992 | 55  | 
mirror_Leaf: "mirror\<cdot>(Leaf\<cdot>a) = Leaf\<cdot>a"  | 
| 
30158
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
56  | 
| mirror_Node: "mirror\<cdot>(Node\<cdot>l\<cdot>r) = Node\<cdot>(mirror\<cdot>r)\<cdot>(mirror\<cdot>l)"  | 
| 29992 | 57  | 
|
| 35917 | 58  | 
lemma mirror_strict [simp]: "mirror\<cdot>\<bottom> = \<bottom>"  | 
59  | 
by fixrec_simp  | 
|
| 29992 | 60  | 
|
61  | 
fixrec  | 
|
| 
30158
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
62  | 
pick :: "'a tree \<rightarrow> 'a convex_pd"  | 
| 
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
63  | 
where  | 
| 29992 | 64  | 
  pick_Leaf: "pick\<cdot>(Leaf\<cdot>a) = {a}\<natural>"
 | 
| 
41399
 
ad093e4638e2
changed syntax of powerdomain binary union operators
 
huffman 
parents: 
40774 
diff
changeset
 | 
65  | 
| pick_Node: "pick\<cdot>(Node\<cdot>l\<cdot>r) = pick\<cdot>l \<union>\<natural> pick\<cdot>r"  | 
| 29992 | 66  | 
|
| 35917 | 67  | 
lemma pick_strict [simp]: "pick\<cdot>\<bottom> = \<bottom>"  | 
68  | 
by fixrec_simp  | 
|
| 29992 | 69  | 
|
70  | 
lemma pick_mirror: "pick\<cdot>(mirror\<cdot>t) = pick\<cdot>t"  | 
|
| 
35781
 
b7738ab762b1
renamed some lemmas generated by the domain package
 
huffman 
parents: 
35769 
diff
changeset
 | 
71  | 
by (induct t) (simp_all add: convex_plus_ac)  | 
| 29992 | 72  | 
|
| 
30158
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
73  | 
fixrec tree1 :: "int lift tree"  | 
| 
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
74  | 
where "tree1 = Node\<cdot>(Node\<cdot>(Leaf\<cdot>(Def 1))\<cdot>(Leaf\<cdot>(Def 2)))  | 
| 
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
75  | 
\<cdot>(Node\<cdot>(Leaf\<cdot>(Def 3))\<cdot>(Leaf\<cdot>(Def 4)))"  | 
| 29992 | 76  | 
|
| 
30158
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
77  | 
fixrec tree2 :: "int lift tree"  | 
| 
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
78  | 
where "tree2 = Node\<cdot>(Node\<cdot>(Leaf\<cdot>(Def 1))\<cdot>(Leaf\<cdot>(Def 2)))  | 
| 
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
79  | 
\<cdot>(Node\<cdot>\<bottom>\<cdot>(Leaf\<cdot>(Def 4)))"  | 
| 29992 | 80  | 
|
| 
30158
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
81  | 
fixrec tree3 :: "int lift tree"  | 
| 
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
82  | 
where "tree3 = Node\<cdot>(Node\<cdot>(Leaf\<cdot>(Def 1))\<cdot>tree3)  | 
| 
 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 
huffman 
parents: 
29992 
diff
changeset
 | 
83  | 
\<cdot>(Node\<cdot>(Leaf\<cdot>(Def 3))\<cdot>(Leaf\<cdot>(Def 4)))"  | 
| 29992 | 84  | 
|
| 35769 | 85  | 
declare tree1.simps tree2.simps tree3.simps [simp del]  | 
| 29992 | 86  | 
|
87  | 
lemma pick_tree1:  | 
|
88  | 
  "pick\<cdot>tree1 = {Def 1, Def 2, Def 3, Def 4}\<natural>"
 | 
|
| 35769 | 89  | 
apply (subst tree1.simps)  | 
| 29992 | 90  | 
apply simp  | 
91  | 
apply (simp add: convex_plus_ac)  | 
|
92  | 
done  | 
|
93  | 
||
94  | 
lemma pick_tree2:  | 
|
95  | 
  "pick\<cdot>tree2 = {Def 1, Def 2, \<bottom>, Def 4}\<natural>"
 | 
|
| 35769 | 96  | 
apply (subst tree2.simps)  | 
| 29992 | 97  | 
apply simp  | 
98  | 
apply (simp add: convex_plus_ac)  | 
|
99  | 
done  | 
|
100  | 
||
101  | 
lemma pick_tree3:  | 
|
102  | 
  "pick\<cdot>tree3 = {Def 1, \<bottom>, Def 3, Def 4}\<natural>"
 | 
|
| 35769 | 103  | 
apply (subst tree3.simps)  | 
| 29992 | 104  | 
apply simp  | 
| 35769 | 105  | 
apply (induct rule: tree3.induct)  | 
| 29992 | 106  | 
apply simp  | 
107  | 
apply simp  | 
|
108  | 
apply (simp add: convex_plus_ac)  | 
|
109  | 
apply simp  | 
|
110  | 
apply (simp add: convex_plus_ac)  | 
|
111  | 
done  | 
|
112  | 
||
113  | 
end  |