src/HOLCF/ex/Powerdomain_ex.thy
author wenzelm
Tue, 21 Jul 2009 01:03:18 +0200
changeset 32091 30e2ffbba718
parent 30158 83c50c62cf23
child 35169 31cbcb019003
permissions -rw-r--r--
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/ex/Powerdomain_ex.thy
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
     3
*)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
     4
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
     5
header {* Powerdomain examples *}
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
     6
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
     7
theory Powerdomain_ex
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
     8
imports HOLCF
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
     9
begin
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    10
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    11
defaultsort bifinite
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    12
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    13
subsection {* Monadic sorting example *}
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    14
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    15
domain ordering = LT | EQ | GT
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    16
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    17
declare ordering.rews [simp]
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    18
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    19
definition
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    20
  compare :: "int lift \<rightarrow> int lift \<rightarrow> ordering" where
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    21
  "compare = (FLIFT x y. if x < y then LT else if x = y then EQ else GT)"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    22
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    23
definition
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    24
  is_le :: "int lift \<rightarrow> int lift \<rightarrow> tr" where
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    25
  "is_le = (\<Lambda> x y. case compare\<cdot>x\<cdot>y of LT \<Rightarrow> TT | EQ \<Rightarrow> TT | GT \<Rightarrow> FF)"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    26
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    27
definition
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    28
  is_less :: "int lift \<rightarrow> int lift \<rightarrow> tr" where
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    29
  "is_less = (\<Lambda> x y. case compare\<cdot>x\<cdot>y of LT \<Rightarrow> TT | EQ \<Rightarrow> FF | GT \<Rightarrow> FF)"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    30
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    31
definition
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    32
  r1 :: "(int lift \<times> 'a) \<rightarrow> (int lift \<times> 'a) \<rightarrow> tr convex_pd" where
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    33
  "r1 = (\<Lambda> \<langle>x,_\<rangle> \<langle>y,_\<rangle>. case compare\<cdot>x\<cdot>y of
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    34
          LT \<Rightarrow> {TT}\<natural> |
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    35
          EQ \<Rightarrow> {TT, FF}\<natural> |
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    36
          GT \<Rightarrow> {FF}\<natural>)"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    37
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    38
definition
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    39
  r2 :: "(int lift \<times> 'a) \<rightarrow> (int lift \<times> 'a) \<rightarrow> tr convex_pd" where
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    40
  "r2 = (\<Lambda> \<langle>x,_\<rangle> \<langle>y,_\<rangle>. {is_le\<cdot>x\<cdot>y, is_less\<cdot>x\<cdot>y}\<natural>)"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    41
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    42
lemma r1_r2: "r1\<cdot>\<langle>x,a\<rangle>\<cdot>\<langle>y,b\<rangle> = (r2\<cdot>\<langle>x,a\<rangle>\<cdot>\<langle>y,b\<rangle> :: tr convex_pd)"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    43
apply (simp add: r1_def r2_def)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    44
apply (simp add: is_le_def is_less_def)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    45
apply (cases "compare\<cdot>x\<cdot>y" rule: ordering.casedist)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    46
apply simp_all
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    47
done
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    48
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    49
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    50
subsection {* Picking a leaf from a tree *}
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    51
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    52
domain 'a tree =
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    53
  Node (lazy "'a tree") (lazy "'a tree") |
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    54
  Leaf (lazy "'a")
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    55
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    56
fixrec
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    57
  mirror :: "'a tree \<rightarrow> 'a tree"
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    58
where
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    59
  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
    60
| mirror_Node: "mirror\<cdot>(Node\<cdot>l\<cdot>r) = Node\<cdot>(mirror\<cdot>r)\<cdot>(mirror\<cdot>l)"
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    61
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    62
fixpat
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    63
  mirror_strict [simp]: "mirror\<cdot>\<bottom>"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    64
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    65
fixrec
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    66
  pick :: "'a tree \<rightarrow> 'a convex_pd"
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    67
where
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    68
  pick_Leaf: "pick\<cdot>(Leaf\<cdot>a) = {a}\<natural>"
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    69
| pick_Node: "pick\<cdot>(Node\<cdot>l\<cdot>r) = pick\<cdot>l +\<natural> pick\<cdot>r"
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    70
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    71
fixpat
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    72
  pick_strict [simp]: "pick\<cdot>\<bottom>"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    73
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    74
lemma pick_mirror: "pick\<cdot>(mirror\<cdot>t) = pick\<cdot>t"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    75
by (induct t rule: tree.ind)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    76
   (simp_all add: convex_plus_ac)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    77
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    78
fixrec tree1 :: "int lift tree"
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    79
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
    80
                   \<cdot>(Node\<cdot>(Leaf\<cdot>(Def 3))\<cdot>(Leaf\<cdot>(Def 4)))"
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    81
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    82
fixrec tree2 :: "int lift tree"
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    83
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
    84
                   \<cdot>(Node\<cdot>\<bottom>\<cdot>(Leaf\<cdot>(Def 4)))"
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    85
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    86
fixrec tree3 :: "int lift tree"
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    87
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
    88
                   \<cdot>(Node\<cdot>(Leaf\<cdot>(Def 3))\<cdot>(Leaf\<cdot>(Def 4)))"
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    89
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    90
declare tree1_simps tree2_simps tree3_simps [simp del]
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    91
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    92
lemma pick_tree1:
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    93
  "pick\<cdot>tree1 = {Def 1, Def 2, Def 3, Def 4}\<natural>"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    94
apply (subst tree1_simps)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    95
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    96
apply (simp add: convex_plus_ac)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    97
done
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    98
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    99
lemma pick_tree2:
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   100
  "pick\<cdot>tree2 = {Def 1, Def 2, \<bottom>, Def 4}\<natural>"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   101
apply (subst tree2_simps)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   102
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   103
apply (simp add: convex_plus_ac)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   104
done
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   105
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   106
lemma pick_tree3:
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   107
  "pick\<cdot>tree3 = {Def 1, \<bottom>, Def 3, Def 4}\<natural>"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   108
apply (subst tree3_simps)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   109
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   110
apply (induct rule: tree3_induct)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   111
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   112
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   113
apply (simp add: convex_plus_ac)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   114
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   115
apply (simp add: convex_plus_ac)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   116
done
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   117
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   118
end