src/HOLCF/Fixrec.thy
author huffman
Sat, 05 Nov 2005 21:56:45 +0100
changeset 18094 404f298220af
parent 16779 ac1dc3d4746a
child 18096 574aa0487069
permissions -rw-r--r--
simplify definitions
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
16551
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
     9
imports Sprod Ssum Up One Tr Fix
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 16401
diff changeset
    10
uses ("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
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
    15
defaultsort cpo
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
    16
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    17
types 'a maybe = "one ++ 'a u"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    18
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    19
constdefs
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    20
  fail :: "'a maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    21
  "fail \<equiv> sinl\<cdot>ONE"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    22
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    23
  return :: "'a \<rightarrow> 'a maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    24
  "return \<equiv> sinr oo up"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    25
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    26
lemma maybeE:
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    27
  "\<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
    28
apply (unfold fail_def return_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    29
apply (rule_tac p=p in ssumE, simp)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    30
apply (rule_tac p=x in oneE, simp, simp)
16754
1b979f8b7e8e renamed upE1 to upE
huffman
parents: 16551
diff changeset
    31
apply (rule_tac p=y in upE, simp, simp)
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    32
done
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
subsection {* Monadic bind operator *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    35
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    36
constdefs
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    37
  bind :: "'a maybe \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> 'b maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    38
  "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
    39
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    40
syntax
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    41
  "_bind" :: "'a maybe \<Rightarrow> ('a \<rightarrow> 'b maybe) \<Rightarrow> 'b maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    42
    ("(_ >>= _)" [50, 51] 50)
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
translations "m >>= k" == "bind\<cdot>m\<cdot>k"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    45
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    46
nonterminals
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    47
  maybebind maybebinds
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    48
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    49
syntax 
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    50
  "_MBIND"  :: "pttrn \<Rightarrow> 'a maybe \<Rightarrow> maybebind"         ("(2_ <-/ _)" 10)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    51
  ""        :: "maybebind \<Rightarrow> maybebinds"                ("_")
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    52
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    53
  "_MBINDS" :: "[maybebind, maybebinds] \<Rightarrow> maybebinds"  ("_;/ _")
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    54
  "_MDO"    :: "[maybebinds, 'a maybe] \<Rightarrow> 'a maybe"     ("(do _;/ (_))" 10)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    55
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    56
translations
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    57
  "_MDO (_MBINDS b bs) e" == "_MDO b (_MDO bs e)"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    58
  "do (x,y) <- m; e" == "m >>= (LAM <x,y>. e)" 
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    59
  "do x <- m; e"            == "m >>= (LAM x. e)"
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
text {* monad laws *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    62
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    63
lemma bind_strict [simp]: "UU >>= f = UU"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    64
by (simp add: bind_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    65
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    66
lemma bind_fail [simp]: "fail >>= f = fail"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    67
by (simp add: bind_def fail_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    68
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    69
lemma left_unit [simp]: "(return\<cdot>a) >>= k = k\<cdot>a"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    70
by (simp add: bind_def return_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    71
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    72
lemma right_unit [simp]: "m >>= return = m"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    73
by (rule_tac p=m in maybeE, simp_all)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    74
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    75
lemma bind_assoc [simp]:
16779
ac1dc3d4746a changed orientation of bind_assoc rule
huffman
parents: 16776
diff changeset
    76
 "(do b <- (do a <- m; k\<cdot>a); h\<cdot>b) = (do a <- m; b <- k\<cdot>a; h\<cdot>b)"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    77
by (rule_tac p=m in maybeE, simp_all)
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
subsection {* Run operator *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    80
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    81
constdefs
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
    82
  run:: "'a::pcpo maybe \<rightarrow> 'a"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    83
  "run \<equiv> sscase\<cdot>\<bottom>\<cdot>(fup\<cdot>ID)"
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
text {* rewrite rules for run *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    86
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    87
lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    88
by (simp add: run_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    89
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    90
lemma run_fail [simp]: "run\<cdot>fail = \<bottom>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    91
by (simp add: run_def fail_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    92
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    93
lemma run_return [simp]: "run\<cdot>(return\<cdot>x) = x"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    94
by (simp add: run_def return_def)
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
subsection {* Monad plus operator *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    97
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    98
constdefs
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    99
  mplus :: "'a maybe \<rightarrow> 'a maybe \<rightarrow> 'a maybe"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   100
  "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
   101
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   102
syntax "+++" :: "'a maybe \<Rightarrow> 'a maybe \<Rightarrow> 'a maybe" (infixr 65)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   103
translations "x +++ y" == "mplus\<cdot>x\<cdot>y"
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
text {* rewrite rules for mplus *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   106
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   107
lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   108
by (simp add: mplus_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   109
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   110
lemma mplus_fail [simp]: "fail +++ m = m"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   111
by (simp add: mplus_def fail_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   112
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   113
lemma mplus_return [simp]: "return\<cdot>x +++ m = return\<cdot>x"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   114
by (simp add: mplus_def return_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   115
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   116
lemma mplus_fail2 [simp]: "m +++ fail = m"
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   117
by (rule_tac p=m in maybeE, simp_all)
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   118
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   119
lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   120
by (rule_tac p=x in maybeE, simp_all)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   121
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   122
subsection {* Match functions for built-in types *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   123
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   124
defaultsort pcpo
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   125
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   126
constdefs
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   127
  match_UU :: "'a \<rightarrow> unit maybe"
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   128
  "match_UU \<equiv> \<Lambda> x. fail"
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   129
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   130
  match_cpair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<times> 'b) maybe"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   131
  "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
   132
16551
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   133
  match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<times> 'b) maybe"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   134
  "match_spair \<equiv> ssplit\<cdot>(\<Lambda> x y. return\<cdot><x,y>)"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   135
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   136
  match_sinl :: "'a \<oplus> 'b \<rightarrow> 'a maybe"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   137
  "match_sinl \<equiv> sscase\<cdot>return\<cdot>(\<Lambda> y. fail)"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   138
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   139
  match_sinr :: "'a \<oplus> 'b \<rightarrow> 'b maybe"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   140
  "match_sinr \<equiv> sscase\<cdot>(\<Lambda> x. fail)\<cdot>return"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   141
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   142
  match_up :: "'a::cpo u \<rightarrow> 'a maybe"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   143
  "match_up \<equiv> fup\<cdot>return"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   144
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   145
  match_ONE :: "one \<rightarrow> unit maybe"
18094
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   146
  "match_ONE \<equiv> \<Lambda> ONE. return\<cdot>()"
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   147
 
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   148
  match_TT :: "tr \<rightarrow> unit maybe"
18094
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   149
  "match_TT \<equiv> \<Lambda> b. If b then return\<cdot>() else fail fi"
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   150
 
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   151
  match_FF :: "tr \<rightarrow> unit maybe"
18094
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   152
  "match_FF \<equiv> \<Lambda> b. If b then fail else return\<cdot>() fi"
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   153
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   154
lemma match_UU_simps [simp]:
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   155
  "match_UU\<cdot>x = fail"
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   156
by (simp add: match_UU_def)
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   157
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   158
lemma match_cpair_simps [simp]:
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   159
  "match_cpair\<cdot><x,y> = return\<cdot><x,y>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   160
by (simp add: match_cpair_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   161
16551
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   162
lemma match_spair_simps [simp]:
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   163
  "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> match_spair\<cdot>(:x,y:) = return\<cdot><x,y>"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   164
  "match_spair\<cdot>\<bottom> = \<bottom>"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   165
by (simp_all add: match_spair_def)
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   166
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   167
lemma match_sinl_simps [simp]:
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   168
  "x \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinl\<cdot>x) = return\<cdot>x"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   169
  "x \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinr\<cdot>x) = fail"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   170
  "match_sinl\<cdot>\<bottom> = \<bottom>"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   171
by (simp_all add: match_sinl_def)
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   172
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   173
lemma match_sinr_simps [simp]:
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   174
  "x \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinr\<cdot>x) = return\<cdot>x"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   175
  "x \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinl\<cdot>x) = fail"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   176
  "match_sinr\<cdot>\<bottom> = \<bottom>"
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   177
by (simp_all add: match_sinr_def)
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   178
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   179
lemma match_up_simps [simp]:
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   180
  "match_up\<cdot>(up\<cdot>x) = return\<cdot>x"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   181
  "match_up\<cdot>\<bottom> = \<bottom>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   182
by (simp_all add: match_up_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   183
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   184
lemma match_ONE_simps [simp]:
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   185
  "match_ONE\<cdot>ONE = return\<cdot>()"
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   186
  "match_ONE\<cdot>\<bottom> = \<bottom>"
18094
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   187
by (simp_all add: match_ONE_def)
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   188
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   189
lemma match_TT_simps [simp]:
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   190
  "match_TT\<cdot>TT = return\<cdot>()"
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   191
  "match_TT\<cdot>FF = fail"
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   192
  "match_TT\<cdot>\<bottom> = \<bottom>"
18094
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   193
by (simp_all add: match_TT_def)
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   194
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   195
lemma match_FF_simps [simp]:
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   196
  "match_FF\<cdot>FF = return\<cdot>()"
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   197
  "match_FF\<cdot>TT = fail"
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   198
  "match_FF\<cdot>\<bottom> = \<bottom>"
18094
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   199
by (simp_all add: match_FF_def)
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   200
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   201
subsection {* Mutual recursion *}
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   202
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   203
text {*
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   204
  The following rules are used to prove unfolding theorems from
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   205
  fixed-point definitions of mutually recursive functions.
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   206
*}
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   207
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   208
lemma cpair_equalI: "\<lbrakk>x \<equiv> cfst\<cdot>p; y \<equiv> csnd\<cdot>p\<rbrakk> \<Longrightarrow> <x,y> \<equiv> p"
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   209
by (simp add: surjective_pairing_Cprod2)
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   210
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   211
lemma cpair_eqD1: "<x,y> = <x',y'> \<Longrightarrow> x = x'"
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   212
by simp
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   213
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   214
lemma cpair_eqD2: "<x,y> = <x',y'> \<Longrightarrow> y = y'"
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   215
by simp
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   216
16463
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16460
diff changeset
   217
text {* lemma for proving rewrite rules *}
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16460
diff changeset
   218
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16460
diff changeset
   219
lemma ssubst_lhs: "\<lbrakk>t = s; P s = Q\<rbrakk> \<Longrightarrow> P t = Q"
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16460
diff changeset
   220
by simp
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16460
diff changeset
   221
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   222
ML {*
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   223
val cpair_equalI = thm "cpair_equalI";
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   224
val cpair_eqD1 = thm "cpair_eqD1";
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   225
val cpair_eqD2 = thm "cpair_eqD2";
16463
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16460
diff changeset
   226
val ssubst_lhs = thm "ssubst_lhs";
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   227
*}
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   228
16758
c32334d98fcd fix typo
huffman
parents: 16754
diff changeset
   229
subsection {* Initializing the fixrec package *}
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   230
16229
77cae9c8e73e use fixrec_package.ML
huffman
parents: 16221
diff changeset
   231
use "fixrec_package.ML"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   232
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   233
end