src/HOLCF/Fixrec.thy
author huffman
Sat, 04 Jun 2005 02:11:47 +0200
changeset 16229 77cae9c8e73e
parent 16221 879400e029bf
child 16401 57c35ede00b9
permissions -rw-r--r--
use fixrec_package.ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/Fixrec.thy
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     2
    ID:         $Id$
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     3
    Author:     Amber Telfer and Brian Huffman
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     4
*)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     5
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     6
header "Package for defining recursive functions in HOLCF"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     7
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     8
theory Fixrec
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     9
imports Ssum One Up Fix
16229
77cae9c8e73e use fixrec_package.ML
huffman
parents: 16221
diff changeset
    10
files ("fixrec_package.ML")
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    11
begin
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    12
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    13
subsection {* Maybe monad type *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    14
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    15
types 'a maybe = "one ++ 'a u"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    16
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    17
constdefs
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    18
  fail :: "'a maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    19
  "fail \<equiv> sinl\<cdot>ONE"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    20
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    21
  return :: "'a \<rightarrow> 'a maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    22
  "return \<equiv> sinr oo up"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    23
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    24
lemma maybeE:
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    25
  "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = return\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    26
apply (unfold fail_def return_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    27
apply (rule_tac p=p in ssumE, simp)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    28
apply (rule_tac p=x in oneE, simp, simp)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    29
apply (rule_tac p=y in upE1, simp, simp)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    30
done
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    31
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    32
subsection {* Monadic bind operator *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    33
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    34
constdefs
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    35
  bind :: "'a maybe \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> 'b maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    36
  "bind \<equiv> \<Lambda> m f. sscase\<cdot>sinl\<cdot>(fup\<cdot>f)\<cdot>m"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    37
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    38
syntax
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    39
  "_bind" :: "'a maybe \<Rightarrow> ('a \<rightarrow> 'b maybe) \<Rightarrow> 'b maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    40
    ("(_ >>= _)" [50, 51] 50)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    41
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    42
translations "m >>= k" == "bind\<cdot>m\<cdot>k"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    43
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    44
nonterminals
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    45
  maybebind maybebinds
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    46
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    47
syntax 
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    48
  "_MBIND"  :: "pttrn \<Rightarrow> 'a maybe \<Rightarrow> maybebind"         ("(2_ <-/ _)" 10)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    49
  ""        :: "maybebind \<Rightarrow> maybebinds"                ("_")
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    50
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    51
  "_MBINDS" :: "[maybebind, maybebinds] \<Rightarrow> maybebinds"  ("_;/ _")
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    52
  "_MDO"    :: "[maybebinds, 'a maybe] \<Rightarrow> 'a maybe"     ("(do _;/ (_))" 10)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    53
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    54
translations
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    55
  "_MDO (_MBINDS b bs) e" == "_MDO b (_MDO bs e)"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    56
  "do (x,y) <- m; e" == "m >>= (LAM <x,y>. e)" 
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    57
  "do x <- m; e"            == "m >>= (LAM x. e)"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    58
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    59
text {* monad laws *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    60
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    61
lemma bind_strict [simp]: "UU >>= f = UU"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    62
by (simp add: bind_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    63
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    64
lemma bind_fail [simp]: "fail >>= f = fail"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    65
by (simp add: bind_def fail_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    66
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    67
lemma left_unit [simp]: "(return\<cdot>a) >>= k = k\<cdot>a"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    68
by (simp add: bind_def return_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    69
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    70
lemma right_unit [simp]: "m >>= return = m"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    71
by (rule_tac p=m in maybeE, simp_all)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    72
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    73
lemma bind_assoc [simp]:
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    74
 "(do a <- m; b <- k\<cdot>a; h\<cdot>b) = (do b <- (do a <- m; k\<cdot>a); h\<cdot>b)"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    75
by (rule_tac p=m in maybeE, simp_all)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    76
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    77
subsection {* Run operator *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    78
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    79
constdefs
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    80
  run:: "'a maybe \<rightarrow> 'a"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    81
  "run \<equiv> sscase\<cdot>\<bottom>\<cdot>(fup\<cdot>ID)"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    82
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    83
text {* rewrite rules for run *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    84
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    85
lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    86
by (simp add: run_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    87
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    88
lemma run_fail [simp]: "run\<cdot>fail = \<bottom>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    89
by (simp add: run_def fail_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    90
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    91
lemma run_return [simp]: "run\<cdot>(return\<cdot>x) = x"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    92
by (simp add: run_def return_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    93
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    94
subsection {* Monad plus operator *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    95
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    96
constdefs
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    97
  mplus :: "'a maybe \<rightarrow> 'a maybe \<rightarrow> 'a maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    98
  "mplus \<equiv> \<Lambda> m1 m2. sscase\<cdot>(\<Lambda> x. m2)\<cdot>(fup\<cdot>return)\<cdot>m1"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    99
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   100
syntax "+++" :: "'a maybe \<Rightarrow> 'a maybe \<Rightarrow> 'a maybe" (infixr 65)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   101
translations "x +++ y" == "mplus\<cdot>x\<cdot>y"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   102
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   103
text {* rewrite rules for mplus *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   104
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   105
lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   106
by (simp add: mplus_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   107
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   108
lemma mplus_fail [simp]: "fail +++ m = m"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   109
by (simp add: mplus_def fail_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   110
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   111
lemma mplus_return [simp]: "return\<cdot>x +++ m = return\<cdot>x"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   112
by (simp add: mplus_def return_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   113
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   114
lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   115
by (rule_tac p=x in maybeE, simp_all)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   116
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   117
subsection {* Match functions for built-in types *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   118
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   119
text {* Currently the package only supports lazy constructors *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   120
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   121
constdefs
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   122
  match_cpair :: "'a \<times> 'b \<rightarrow> ('a \<times> 'b) maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   123
  "match_cpair \<equiv> csplit\<cdot>(\<Lambda> x y. return\<cdot><x,y>)"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   124
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   125
  match_up :: "'a u \<rightarrow> 'a maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   126
  "match_up \<equiv> fup\<cdot>return"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   127
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   128
lemma match_cpair_simps [simp]:
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   129
  "match_cpair\<cdot><x,y> = return\<cdot><x,y>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   130
by (simp add: match_cpair_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   131
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   132
lemma match_up_simps [simp]:
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   133
  "match_up\<cdot>(up\<cdot>x) = return\<cdot>x"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   134
  "match_up\<cdot>\<bottom> = \<bottom>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   135
by (simp_all add: match_up_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   136
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   137
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   138
subsection {* Intitializing the fixrec package *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   139
16229
77cae9c8e73e use fixrec_package.ML
huffman
parents: 16221
diff changeset
   140
use "fixrec_package.ML"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   141
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   142
end