src/HOLCF/ex/Powerdomain_ex.thy
author blanchet
Mon, 19 Apr 2010 18:14:45 +0200
changeset 36230 43d10a494c91
parent 35918 68397d86d454
child 36452 d37c6eed8117
permissions -rw-r--r--
added warning about inconsistent context to Metis; it makes more sense here than in Sledgehammer, because Sledgehammer is unsound and there's no point in having people panicking about the consistency of their context when their context is in fact consistent
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
definition
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    18
  compare :: "int lift \<rightarrow> int lift \<rightarrow> ordering" where
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    19
  "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
    20
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    21
definition
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    22
  is_le :: "int lift \<rightarrow> int lift \<rightarrow> tr" where
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    23
  "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
    24
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    25
definition
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    26
  is_less :: "int lift \<rightarrow> int lift \<rightarrow> tr" where
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    27
  "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
    28
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    29
definition
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    30
  r1 :: "(int lift \<times> 'a) \<rightarrow> (int lift \<times> 'a) \<rightarrow> tr convex_pd" where
35918
68397d86d454 use Pair instead of cpair
huffman
parents: 35917
diff changeset
    31
  "r1 = (\<Lambda> (x,_) (y,_). case compare\<cdot>x\<cdot>y of
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    32
          LT \<Rightarrow> {TT}\<natural> |
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    33
          EQ \<Rightarrow> {TT, FF}\<natural> |
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    34
          GT \<Rightarrow> {FF}\<natural>)"
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    35
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    36
definition
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    37
  r2 :: "(int lift \<times> 'a) \<rightarrow> (int lift \<times> 'a) \<rightarrow> tr convex_pd" where
35918
68397d86d454 use Pair instead of cpair
huffman
parents: 35917
diff changeset
    38
  "r2 = (\<Lambda> (x,_) (y,_). {is_le\<cdot>x\<cdot>y, is_less\<cdot>x\<cdot>y}\<natural>)"
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    39
35918
68397d86d454 use Pair instead of cpair
huffman
parents: 35917
diff changeset
    40
lemma r1_r2: "r1\<cdot>(x,a)\<cdot>(y,b) = (r2\<cdot>(x,a)\<cdot>(y,b) :: tr convex_pd)"
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    41
apply (simp add: r1_def r2_def)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    42
apply (simp add: is_le_def is_less_def)
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35769
diff changeset
    43
apply (cases "compare\<cdot>x\<cdot>y")
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    44
apply simp_all
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    45
done
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    46
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    47
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    48
subsection {* Picking a leaf from a tree *}
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    49
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    50
domain 'a tree =
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    51
  Node (lazy "'a tree") (lazy "'a tree") |
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    52
  Leaf (lazy "'a")
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    53
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    54
fixrec
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    55
  mirror :: "'a tree \<rightarrow> 'a tree"
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    56
where
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    57
  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
    58
| 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
    59
35917
85b0efdcdae9 use fixrec_simp instead of fixpat
huffman
parents: 35781
diff changeset
    60
lemma mirror_strict [simp]: "mirror\<cdot>\<bottom> = \<bottom>"
85b0efdcdae9 use fixrec_simp instead of fixpat
huffman
parents: 35781
diff changeset
    61
by fixrec_simp
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    62
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    63
fixrec
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    64
  pick :: "'a tree \<rightarrow> 'a convex_pd"
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    65
where
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    66
  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
    67
| 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
    68
35917
85b0efdcdae9 use fixrec_simp instead of fixpat
huffman
parents: 35781
diff changeset
    69
lemma pick_strict [simp]: "pick\<cdot>\<bottom> = \<bottom>"
85b0efdcdae9 use fixrec_simp instead of fixpat
huffman
parents: 35781
diff changeset
    70
by fixrec_simp
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    71
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    72
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
    73
by (induct t) (simp_all add: convex_plus_ac)
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    74
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    75
fixrec tree1 :: "int lift tree"
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    76
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
    77
                   \<cdot>(Node\<cdot>(Leaf\<cdot>(Def 3))\<cdot>(Leaf\<cdot>(Def 4)))"
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    78
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    79
fixrec tree2 :: "int lift tree"
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    80
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
    81
                   \<cdot>(Node\<cdot>\<bottom>\<cdot>(Leaf\<cdot>(Def 4)))"
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    82
30158
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    83
fixrec tree3 :: "int lift tree"
83c50c62cf23 fixrec package uses new-style syntax and local-theory interface
huffman
parents: 29992
diff changeset
    84
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
    85
                   \<cdot>(Node\<cdot>(Leaf\<cdot>(Def 3))\<cdot>(Leaf\<cdot>(Def 4)))"
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    86
35769
500c32e5fadc fixrec now generates qualified theorem names
huffman
parents: 35169
diff changeset
    87
declare tree1.simps tree2.simps tree3.simps [simp del]
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    88
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    89
lemma pick_tree1:
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    90
  "pick\<cdot>tree1 = {Def 1, Def 2, Def 3, Def 4}\<natural>"
35769
500c32e5fadc fixrec now generates qualified theorem names
huffman
parents: 35169
diff changeset
    91
apply (subst tree1.simps)
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    92
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    93
apply (simp add: convex_plus_ac)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    94
done
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    95
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    96
lemma pick_tree2:
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    97
  "pick\<cdot>tree2 = {Def 1, Def 2, \<bottom>, Def 4}\<natural>"
35769
500c32e5fadc fixrec now generates qualified theorem names
huffman
parents: 35169
diff changeset
    98
apply (subst tree2.simps)
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
    99
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   100
apply (simp add: convex_plus_ac)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   101
done
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   102
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   103
lemma pick_tree3:
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   104
  "pick\<cdot>tree3 = {Def 1, \<bottom>, Def 3, Def 4}\<natural>"
35769
500c32e5fadc fixrec now generates qualified theorem names
huffman
parents: 35169
diff changeset
   105
apply (subst tree3.simps)
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   106
apply simp
35769
500c32e5fadc fixrec now generates qualified theorem names
huffman
parents: 35169
diff changeset
   107
apply (induct rule: tree3.induct)
29992
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   108
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   109
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   110
apply (simp add: convex_plus_ac)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   111
apply simp
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   112
apply (simp add: convex_plus_ac)
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   113
done
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   114
5deee36e33c4 add Powerdomain_ex.thy
huffman
parents:
diff changeset
   115
end