src/HOLCF/Fixrec.thy
author huffman
Wed, 19 May 2010 16:08:41 -0700
changeset 36989 9316a18ec931
parent 36452 d37c6eed8117
child 37063 0cd15d8c90a0
permissions -rw-r--r--
remove unnecessary constant Fixrec.bind
huffman@16221
     1
(*  Title:      HOLCF/Fixrec.thy
huffman@16221
     2
    Author:     Amber Telfer and Brian Huffman
huffman@16221
     3
*)
huffman@16221
     4
huffman@16221
     5
header "Package for defining recursive functions in HOLCF"
huffman@16221
     6
huffman@16221
     7
theory Fixrec
huffman@35939
     8
imports Cprod Sprod Ssum Up One Tr Fix
huffman@35513
     9
uses
huffman@35513
    10
  ("Tools/holcf_library.ML")
huffman@35513
    11
  ("Tools/fixrec.ML")
huffman@16221
    12
begin
huffman@16221
    13
huffman@16221
    14
subsection {* Maybe monad type *}
huffman@16221
    15
wenzelm@36452
    16
default_sort cpo
huffman@16776
    17
huffman@19092
    18
pcpodef (open) 'a maybe = "UNIV::(one ++ 'a u) set"
wenzelm@29063
    19
by simp_all
huffman@16221
    20
huffman@29141
    21
definition
huffman@29141
    22
  fail :: "'a maybe" where
huffman@29141
    23
  "fail = Abs_maybe (sinl\<cdot>ONE)"
huffman@16221
    24
huffman@29141
    25
definition
wenzelm@25131
    26
  return :: "'a \<rightarrow> 'a maybe" where
huffman@29141
    27
  "return = (\<Lambda> x. Abs_maybe (sinr\<cdot>(up\<cdot>x)))"
huffman@19092
    28
wenzelm@25131
    29
definition
wenzelm@25131
    30
  maybe_when :: "'b \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a maybe \<rightarrow> 'b::pcpo" where
wenzelm@25131
    31
  "maybe_when = (\<Lambda> f r m. sscase\<cdot>(\<Lambda> x. f)\<cdot>(fup\<cdot>r)\<cdot>(Rep_maybe m))"
huffman@16221
    32
huffman@16221
    33
lemma maybeE:
huffman@16221
    34
  "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = return\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
huffman@16221
    35
apply (unfold fail_def return_def)
huffman@19092
    36
apply (cases p, rename_tac r)
huffman@19092
    37
apply (rule_tac p=r in ssumE, simp add: Abs_maybe_strict)
huffman@16221
    38
apply (rule_tac p=x in oneE, simp, simp)
huffman@19092
    39
apply (rule_tac p=y in upE, simp, simp add: cont_Abs_maybe)
huffman@16221
    40
done
huffman@16221
    41
huffman@18293
    42
lemma return_defined [simp]: "return\<cdot>x \<noteq> \<bottom>"
huffman@19092
    43
by (simp add: return_def cont_Abs_maybe Abs_maybe_defined)
huffman@18293
    44
huffman@18293
    45
lemma fail_defined [simp]: "fail \<noteq> \<bottom>"
huffman@19092
    46
by (simp add: fail_def Abs_maybe_defined)
huffman@18293
    47
huffman@18293
    48
lemma return_eq [simp]: "(return\<cdot>x = return\<cdot>y) = (x = y)"
huffman@19092
    49
by (simp add: return_def cont_Abs_maybe Abs_maybe_inject)
huffman@18293
    50
huffman@18293
    51
lemma return_neq_fail [simp]:
huffman@18293
    52
  "return\<cdot>x \<noteq> fail" "fail \<noteq> return\<cdot>x"
huffman@19092
    53
by (simp_all add: return_def fail_def cont_Abs_maybe Abs_maybe_inject)
huffman@19092
    54
huffman@19092
    55
lemma maybe_when_rews [simp]:
huffman@19092
    56
  "maybe_when\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>"
huffman@19092
    57
  "maybe_when\<cdot>f\<cdot>r\<cdot>fail = f"
huffman@19092
    58
  "maybe_when\<cdot>f\<cdot>r\<cdot>(return\<cdot>x) = r\<cdot>x"
huffman@19092
    59
by (simp_all add: return_def fail_def maybe_when_def cont_Rep_maybe
huffman@29530
    60
                  cont2cont_LAM
huffman@19092
    61
                  cont_Abs_maybe Abs_maybe_inverse Rep_maybe_strict)
huffman@19092
    62
huffman@19092
    63
translations
huffman@29141
    64
  "case m of XCONST fail \<Rightarrow> t1 | XCONST return\<cdot>x \<Rightarrow> t2"
huffman@29141
    65
    == "CONST maybe_when\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m"
huffman@18293
    66
huffman@18293
    67
huffman@18097
    68
subsubsection {* Run operator *}
huffman@16221
    69
wenzelm@25131
    70
definition
huffman@28891
    71
  run :: "'a maybe \<rightarrow> 'a::pcpo" where
wenzelm@25131
    72
  "run = maybe_when\<cdot>\<bottom>\<cdot>ID"
huffman@16221
    73
huffman@16221
    74
text {* rewrite rules for run *}
huffman@16221
    75
huffman@16221
    76
lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>"
huffman@16221
    77
by (simp add: run_def)
huffman@16221
    78
huffman@16221
    79
lemma run_fail [simp]: "run\<cdot>fail = \<bottom>"
huffman@19092
    80
by (simp add: run_def)
huffman@16221
    81
huffman@16221
    82
lemma run_return [simp]: "run\<cdot>(return\<cdot>x) = x"
huffman@19092
    83
by (simp add: run_def)
huffman@16221
    84
huffman@18097
    85
subsubsection {* Monad plus operator *}
huffman@16221
    86
wenzelm@25131
    87
definition
wenzelm@25131
    88
  mplus :: "'a maybe \<rightarrow> 'a maybe \<rightarrow> 'a maybe" where
wenzelm@25131
    89
  "mplus = (\<Lambda> m1 m2. case m1 of fail \<Rightarrow> m2 | return\<cdot>x \<Rightarrow> m1)"
huffman@18097
    90
wenzelm@25131
    91
abbreviation
wenzelm@25131
    92
  mplus_syn :: "['a maybe, 'a maybe] \<Rightarrow> 'a maybe"  (infixr "+++" 65)  where
wenzelm@25131
    93
  "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2"
huffman@16221
    94
huffman@16221
    95
text {* rewrite rules for mplus *}
huffman@16221
    96
huffman@16221
    97
lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>"
huffman@16221
    98
by (simp add: mplus_def)
huffman@16221
    99
huffman@16221
   100
lemma mplus_fail [simp]: "fail +++ m = m"
huffman@19092
   101
by (simp add: mplus_def)
huffman@16221
   102
huffman@16221
   103
lemma mplus_return [simp]: "return\<cdot>x +++ m = return\<cdot>x"
huffman@19092
   104
by (simp add: mplus_def)
huffman@16221
   105
huffman@16460
   106
lemma mplus_fail2 [simp]: "m +++ fail = m"
huffman@16460
   107
by (rule_tac p=m in maybeE, simp_all)
huffman@16460
   108
huffman@16221
   109
lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
huffman@16221
   110
by (rule_tac p=x in maybeE, simp_all)
huffman@16221
   111
huffman@18097
   112
subsubsection {* Fatbar combinator *}
huffman@18097
   113
wenzelm@25131
   114
definition
wenzelm@25131
   115
  fatbar :: "('a \<rightarrow> 'b maybe) \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> ('a \<rightarrow> 'b maybe)" where
wenzelm@25131
   116
  "fatbar = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)"
huffman@18097
   117
wenzelm@25131
   118
abbreviation
wenzelm@25131
   119
  fatbar_syn :: "['a \<rightarrow> 'b maybe, 'a \<rightarrow> 'b maybe] \<Rightarrow> 'a \<rightarrow> 'b maybe" (infixr "\<parallel>" 60)  where
wenzelm@25131
   120
  "m1 \<parallel> m2 == fatbar\<cdot>m1\<cdot>m2"
huffman@18097
   121
huffman@18097
   122
lemma fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> (m \<parallel> ms)\<cdot>x = \<bottom>"
huffman@18097
   123
by (simp add: fatbar_def)
huffman@18097
   124
huffman@18097
   125
lemma fatbar2: "m\<cdot>x = fail \<Longrightarrow> (m \<parallel> ms)\<cdot>x = ms\<cdot>x"
huffman@18097
   126
by (simp add: fatbar_def)
huffman@18097
   127
huffman@18097
   128
lemma fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> (m \<parallel> ms)\<cdot>x = return\<cdot>y"
huffman@18097
   129
by (simp add: fatbar_def)
huffman@18097
   130
huffman@18097
   131
lemmas fatbar_simps = fatbar1 fatbar2 fatbar3
huffman@18097
   132
huffman@18112
   133
lemma run_fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = \<bottom>"
huffman@18112
   134
by (simp add: fatbar_def)
huffman@18112
   135
huffman@18112
   136
lemma run_fatbar2: "m\<cdot>x = fail \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = run\<cdot>(ms\<cdot>x)"
huffman@18112
   137
by (simp add: fatbar_def)
huffman@18112
   138
huffman@18112
   139
lemma run_fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = y"
huffman@18112
   140
by (simp add: fatbar_def)
huffman@18112
   141
huffman@18112
   142
lemmas run_fatbar_simps [simp] = run_fatbar1 run_fatbar2 run_fatbar3
huffman@18112
   143
huffman@18293
   144
subsection {* Case branch combinator *}
huffman@18097
   145
huffman@29141
   146
definition
huffman@29141
   147
  branch :: "('a \<rightarrow> 'b maybe) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c maybe)" where
huffman@36989
   148
  "branch p \<equiv> \<Lambda> r x. maybe_when\<cdot>fail\<cdot>(\<Lambda> y. return\<cdot>(r\<cdot>y))\<cdot>(p\<cdot>x)"
huffman@18097
   149
huffman@18293
   150
lemma branch_rews:
huffman@18293
   151
  "p\<cdot>x = \<bottom> \<Longrightarrow> branch p\<cdot>r\<cdot>x = \<bottom>"
huffman@18293
   152
  "p\<cdot>x = fail \<Longrightarrow> branch p\<cdot>r\<cdot>x = fail"
huffman@18293
   153
  "p\<cdot>x = return\<cdot>y \<Longrightarrow> branch p\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>y)"
huffman@18293
   154
by (simp_all add: branch_def)
huffman@18097
   155
huffman@18293
   156
lemma branch_return [simp]: "branch return\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>x)"
huffman@18293
   157
by (simp add: branch_def)
huffman@18097
   158
huffman@28891
   159
subsubsection {* Cases operator *}
huffman@28891
   160
huffman@28891
   161
definition
huffman@28891
   162
  cases :: "'a maybe \<rightarrow> 'a::pcpo" where
huffman@28891
   163
  "cases = maybe_when\<cdot>\<bottom>\<cdot>ID"
huffman@28891
   164
huffman@28891
   165
text {* rewrite rules for cases *}
huffman@28891
   166
huffman@28891
   167
lemma cases_strict [simp]: "cases\<cdot>\<bottom> = \<bottom>"
huffman@28891
   168
by (simp add: cases_def)
huffman@28891
   169
huffman@28891
   170
lemma cases_fail [simp]: "cases\<cdot>fail = \<bottom>"
huffman@28891
   171
by (simp add: cases_def)
huffman@28891
   172
huffman@28891
   173
lemma cases_return [simp]: "cases\<cdot>(return\<cdot>x) = x"
huffman@28891
   174
by (simp add: cases_def)
huffman@18097
   175
huffman@18097
   176
subsection {* Case syntax *}
huffman@18097
   177
huffman@18097
   178
nonterminals
huffman@18097
   179
  Case_syn  Cases_syn
huffman@18097
   180
huffman@18097
   181
syntax
huffman@18097
   182
  "_Case_syntax":: "['a, Cases_syn] => 'b"               ("(Case _ of/ _)" 10)
huffman@18097
   183
  "_Case1"      :: "['a, 'b] => Case_syn"                ("(2_ =>/ _)" 10)
huffman@18097
   184
  ""            :: "Case_syn => Cases_syn"               ("_")
huffman@18097
   185
  "_Case2"      :: "[Case_syn, Cases_syn] => Cases_syn"  ("_/ | _")
huffman@18097
   186
huffman@18097
   187
syntax (xsymbols)
huffman@18097
   188
  "_Case1"      :: "['a, 'b] => Case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
huffman@18097
   189
huffman@18293
   190
translations
huffman@28891
   191
  "_Case_syntax x ms" == "CONST Fixrec.cases\<cdot>(ms\<cdot>x)"
huffman@18293
   192
  "_Case2 m ms" == "m \<parallel> ms"
huffman@18293
   193
huffman@18293
   194
text {* Parsing Case expressions *}
huffman@18293
   195
huffman@18097
   196
syntax
huffman@18112
   197
  "_pat" :: "'a"
wenzelm@29323
   198
  "_variable" :: "'a"
huffman@26046
   199
  "_noargs" :: "'a"
huffman@18097
   200
huffman@18097
   201
translations
wenzelm@29323
   202
  "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_variable p r)"
wenzelm@29323
   203
  "_variable (_args x y) r" => "CONST csplit\<cdot>(_variable x (_variable y r))"
wenzelm@29323
   204
  "_variable _noargs r" => "CONST unit_when\<cdot>r"
huffman@18097
   205
huffman@18112
   206
parse_translation {*
wenzelm@35118
   207
(* rewrite (_pat x) => (return) *)
wenzelm@35118
   208
(* rewrite (_variable x t) => (Abs_CFun (%x. t)) *)
wenzelm@35118
   209
 [(@{syntax_const "_pat"}, fn _ => Syntax.const @{const_syntax Fixrec.return}),
wenzelm@35118
   210
  mk_binder_tr (@{syntax_const "_variable"}, @{const_syntax Abs_CFun})];
huffman@18112
   211
*}
huffman@18112
   212
huffman@18112
   213
text {* Printing Case expressions *}
huffman@18112
   214
huffman@18293
   215
syntax
huffman@18293
   216
  "_match" :: "'a"
huffman@18112
   217
huffman@18112
   218
print_translation {*
huffman@18112
   219
  let
wenzelm@25131
   220
    fun dest_LAM (Const (@{const_syntax Rep_CFun},_) $ Const (@{const_syntax unit_when},_) $ t) =
wenzelm@35118
   221
          (Syntax.const @{syntax_const "_noargs"}, t)
wenzelm@25131
   222
    |   dest_LAM (Const (@{const_syntax Rep_CFun},_) $ Const (@{const_syntax csplit},_) $ t) =
huffman@18293
   223
          let
huffman@18293
   224
            val (v1, t1) = dest_LAM t;
huffman@18293
   225
            val (v2, t2) = dest_LAM t1;
wenzelm@35118
   226
          in (Syntax.const @{syntax_const "_args"} $ v1 $ v2, t2) end
wenzelm@25131
   227
    |   dest_LAM (Const (@{const_syntax Abs_CFun},_) $ t) =
huffman@18293
   228
          let
wenzelm@35118
   229
            val abs =
wenzelm@35118
   230
              case t of Abs abs => abs
huffman@18293
   231
                | _ => ("x", dummyT, incr_boundvars 1 t $ Bound 0);
huffman@18293
   232
            val (x, t') = atomic_abs_tr' abs;
wenzelm@35118
   233
          in (Syntax.const @{syntax_const "_variable"} $ x, t') end
huffman@18293
   234
    |   dest_LAM _ = raise Match; (* too few vars: abort translation *)
huffman@18112
   235
wenzelm@25131
   236
    fun Case1_tr' [Const(@{const_syntax branch},_) $ p, r] =
wenzelm@35118
   237
          let val (v, t) = dest_LAM r in
wenzelm@35118
   238
            Syntax.const @{syntax_const "_Case1"} $
wenzelm@35118
   239
              (Syntax.const @{syntax_const "_match"} $ p $ v) $ t
wenzelm@35118
   240
          end;
huffman@18112
   241
wenzelm@25131
   242
  in [(@{const_syntax Rep_CFun}, Case1_tr')] end;
huffman@18112
   243
*}
huffman@18097
   244
huffman@18293
   245
translations
huffman@35455
   246
  "x" <= "_match (CONST Fixrec.return) (_variable x)"
huffman@18293
   247
huffman@18293
   248
huffman@18293
   249
subsection {* Pattern combinators for data constructors *}
huffman@18293
   250
huffman@18293
   251
types ('a, 'b) pat = "'a \<rightarrow> 'b maybe"
huffman@18097
   252
wenzelm@25131
   253
definition
wenzelm@25131
   254
  cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
huffman@35926
   255
  "cpair_pat p1 p2 = (\<Lambda>(x, y).
huffman@36989
   256
    maybe_when\<cdot>fail\<cdot>(\<Lambda> a. maybe_when\<cdot>fail\<cdot>(\<Lambda> b. return\<cdot>(a, b))\<cdot>(p2\<cdot>y))\<cdot>(p1\<cdot>x))"
huffman@18097
   257
wenzelm@25131
   258
definition
huffman@18293
   259
  spair_pat ::
wenzelm@25131
   260
  "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a::pcpo \<otimes> 'b::pcpo, 'c \<times> 'd) pat" where
huffman@35926
   261
  "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>(x, y))"
huffman@18097
   262
wenzelm@25131
   263
definition
wenzelm@25131
   264
  sinl_pat :: "('a, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
wenzelm@25131
   265
  "sinl_pat p = sscase\<cdot>p\<cdot>(\<Lambda> x. fail)"
huffman@18097
   266
wenzelm@25131
   267
definition
wenzelm@25131
   268
  sinr_pat :: "('b, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
wenzelm@25131
   269
  "sinr_pat p = sscase\<cdot>(\<Lambda> x. fail)\<cdot>p"
huffman@18097
   270
wenzelm@25131
   271
definition
wenzelm@25131
   272
  up_pat :: "('a, 'b) pat \<Rightarrow> ('a u, 'b) pat" where
wenzelm@25131
   273
  "up_pat p = fup\<cdot>p"
huffman@18097
   274
wenzelm@25131
   275
definition
wenzelm@25131
   276
  TT_pat :: "(tr, unit) pat" where
wenzelm@25131
   277
  "TT_pat = (\<Lambda> b. If b then return\<cdot>() else fail fi)"
huffman@18097
   278
wenzelm@25131
   279
definition
wenzelm@25131
   280
  FF_pat :: "(tr, unit) pat" where
wenzelm@25131
   281
  "FF_pat = (\<Lambda> b. If b then fail else return\<cdot>() fi)"
huffman@18097
   282
wenzelm@25131
   283
definition
wenzelm@25131
   284
  ONE_pat :: "(one, unit) pat" where
wenzelm@25131
   285
  "ONE_pat = (\<Lambda> ONE. return\<cdot>())"
huffman@18097
   286
huffman@18293
   287
text {* Parse translations (patterns) *}
huffman@18293
   288
translations
huffman@35926
   289
  "_pat (XCONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"
huffman@26046
   290
  "_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
huffman@26046
   291
  "_pat (XCONST sinl\<cdot>x)" => "CONST sinl_pat (_pat x)"
huffman@26046
   292
  "_pat (XCONST sinr\<cdot>x)" => "CONST sinr_pat (_pat x)"
huffman@26046
   293
  "_pat (XCONST up\<cdot>x)" => "CONST up_pat (_pat x)"
huffman@26046
   294
  "_pat (XCONST TT)" => "CONST TT_pat"
huffman@26046
   295
  "_pat (XCONST FF)" => "CONST FF_pat"
huffman@26046
   296
  "_pat (XCONST ONE)" => "CONST ONE_pat"
huffman@26046
   297
huffman@26046
   298
text {* CONST version is also needed for constructors with special syntax *}
huffman@26046
   299
translations
huffman@35926
   300
  "_pat (CONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"
huffman@26046
   301
  "_pat (CONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
huffman@18097
   302
huffman@18293
   303
text {* Parse translations (variables) *}
huffman@18097
   304
translations
huffman@35926
   305
  "_variable (XCONST Pair x y) r" => "_variable (_args x y) r"
wenzelm@29323
   306
  "_variable (XCONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
wenzelm@29323
   307
  "_variable (XCONST sinl\<cdot>x) r" => "_variable x r"
wenzelm@29323
   308
  "_variable (XCONST sinr\<cdot>x) r" => "_variable x r"
wenzelm@29323
   309
  "_variable (XCONST up\<cdot>x) r" => "_variable x r"
wenzelm@29323
   310
  "_variable (XCONST TT) r" => "_variable _noargs r"
wenzelm@29323
   311
  "_variable (XCONST FF) r" => "_variable _noargs r"
wenzelm@29323
   312
  "_variable (XCONST ONE) r" => "_variable _noargs r"
huffman@26046
   313
huffman@26046
   314
translations
huffman@35926
   315
  "_variable (CONST Pair x y) r" => "_variable (_args x y) r"
wenzelm@29323
   316
  "_variable (CONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
huffman@18097
   317
huffman@18112
   318
text {* Print translations *}
huffman@18097
   319
translations
huffman@35926
   320
  "CONST Pair (_match p1 v1) (_match p2 v2)"
wenzelm@25131
   321
      <= "_match (CONST cpair_pat p1 p2) (_args v1 v2)"
wenzelm@25131
   322
  "CONST spair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)"
wenzelm@25131
   323
      <= "_match (CONST spair_pat p1 p2) (_args v1 v2)"
wenzelm@25131
   324
  "CONST sinl\<cdot>(_match p1 v1)" <= "_match (CONST sinl_pat p1) v1"
wenzelm@25131
   325
  "CONST sinr\<cdot>(_match p1 v1)" <= "_match (CONST sinr_pat p1) v1"
wenzelm@25131
   326
  "CONST up\<cdot>(_match p1 v1)" <= "_match (CONST up_pat p1) v1"
huffman@26046
   327
  "CONST TT" <= "_match (CONST TT_pat) _noargs"
huffman@26046
   328
  "CONST FF" <= "_match (CONST FF_pat) _noargs"
huffman@26046
   329
  "CONST ONE" <= "_match (CONST ONE_pat) _noargs"
huffman@18097
   330
huffman@18293
   331
lemma cpair_pat1:
huffman@35926
   332
  "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = \<bottom>"
huffman@18293
   333
apply (simp add: branch_def cpair_pat_def)
huffman@18293
   334
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
huffman@18293
   335
done
huffman@18097
   336
huffman@18293
   337
lemma cpair_pat2:
huffman@35926
   338
  "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = fail"
huffman@18293
   339
apply (simp add: branch_def cpair_pat_def)
huffman@18293
   340
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
huffman@18293
   341
done
huffman@18097
   342
huffman@18293
   343
lemma cpair_pat3:
huffman@18293
   344
  "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow>
huffman@35926
   345
   branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = branch q\<cdot>s\<cdot>y"
huffman@18293
   346
apply (simp add: branch_def cpair_pat_def)
huffman@18293
   347
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
huffman@18293
   348
apply (rule_tac p="q\<cdot>y" in maybeE, simp_all)
huffman@18293
   349
done
huffman@18097
   350
huffman@18293
   351
lemmas cpair_pat [simp] =
huffman@18293
   352
  cpair_pat1 cpair_pat2 cpair_pat3
huffman@18097
   353
huffman@18293
   354
lemma spair_pat [simp]:
huffman@18293
   355
  "branch (spair_pat p1 p2)\<cdot>r\<cdot>\<bottom> = \<bottom>"
huffman@18293
   356
  "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk>
huffman@18293
   357
     \<Longrightarrow> branch (spair_pat p1 p2)\<cdot>r\<cdot>(:x, y:) =
huffman@35926
   358
         branch (cpair_pat p1 p2)\<cdot>r\<cdot>(x, y)"
huffman@18293
   359
by (simp_all add: branch_def spair_pat_def)
huffman@18097
   360
huffman@18293
   361
lemma sinl_pat [simp]:
huffman@18293
   362
  "branch (sinl_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"
huffman@18293
   363
  "x \<noteq> \<bottom> \<Longrightarrow> branch (sinl_pat p)\<cdot>r\<cdot>(sinl\<cdot>x) = branch p\<cdot>r\<cdot>x"
huffman@18293
   364
  "y \<noteq> \<bottom> \<Longrightarrow> branch (sinl_pat p)\<cdot>r\<cdot>(sinr\<cdot>y) = fail"
huffman@18293
   365
by (simp_all add: branch_def sinl_pat_def)
huffman@18097
   366
huffman@18293
   367
lemma sinr_pat [simp]:
huffman@18293
   368
  "branch (sinr_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"
huffman@18293
   369
  "x \<noteq> \<bottom> \<Longrightarrow> branch (sinr_pat p)\<cdot>r\<cdot>(sinl\<cdot>x) = fail"
huffman@18293
   370
  "y \<noteq> \<bottom> \<Longrightarrow> branch (sinr_pat p)\<cdot>r\<cdot>(sinr\<cdot>y) = branch p\<cdot>r\<cdot>y"
huffman@18293
   371
by (simp_all add: branch_def sinr_pat_def)
huffman@18097
   372
huffman@18293
   373
lemma up_pat [simp]:
huffman@18293
   374
  "branch (up_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"
huffman@18293
   375
  "branch (up_pat p)\<cdot>r\<cdot>(up\<cdot>x) = branch p\<cdot>r\<cdot>x"
huffman@18293
   376
by (simp_all add: branch_def up_pat_def)
huffman@18097
   377
huffman@18293
   378
lemma TT_pat [simp]:
huffman@18293
   379
  "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
huffman@18293
   380
  "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = return\<cdot>r"
huffman@18293
   381
  "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = fail"
huffman@18293
   382
by (simp_all add: branch_def TT_pat_def)
huffman@18097
   383
huffman@18293
   384
lemma FF_pat [simp]:
huffman@18293
   385
  "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
huffman@18293
   386
  "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = fail"
huffman@18293
   387
  "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = return\<cdot>r"
huffman@18293
   388
by (simp_all add: branch_def FF_pat_def)
huffman@18293
   389
huffman@18293
   390
lemma ONE_pat [simp]:
huffman@18293
   391
  "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
huffman@18293
   392
  "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>ONE = return\<cdot>r"
huffman@18293
   393
by (simp_all add: branch_def ONE_pat_def)
huffman@18293
   394
huffman@18293
   395
huffman@18293
   396
subsection {* Wildcards, as-patterns, and lazy patterns *}
huffman@18112
   397
wenzelm@25131
   398
definition
wenzelm@25131
   399
  wild_pat :: "'a \<rightarrow> unit maybe" where
wenzelm@25131
   400
  "wild_pat = (\<Lambda> x. return\<cdot>())"
huffman@18112
   401
wenzelm@25131
   402
definition
wenzelm@25131
   403
  as_pat :: "('a \<rightarrow> 'b maybe) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) maybe" where
huffman@36989
   404
  "as_pat p = (\<Lambda> x. maybe_when\<cdot>fail\<cdot>(\<Lambda> a. return\<cdot>(x, a))\<cdot>(p\<cdot>x))"
huffman@18293
   405
wenzelm@25131
   406
definition
wenzelm@25131
   407
  lazy_pat :: "('a \<rightarrow> 'b::pcpo maybe) \<Rightarrow> ('a \<rightarrow> 'b maybe)" where
huffman@28891
   408
  "lazy_pat p = (\<Lambda> x. return\<cdot>(cases\<cdot>(p\<cdot>x)))"
huffman@18293
   409
huffman@18293
   410
text {* Parse translations (patterns) *}
huffman@18112
   411
translations
huffman@26046
   412
  "_pat _" => "CONST wild_pat"
huffman@18112
   413
huffman@18293
   414
text {* Parse translations (variables) *}
huffman@18293
   415
translations
wenzelm@29323
   416
  "_variable _ r" => "_variable _noargs r"
huffman@18293
   417
huffman@18293
   418
text {* Print translations *}
huffman@18293
   419
translations
huffman@26046
   420
  "_" <= "_match (CONST wild_pat) _noargs"
huffman@19327
   421
huffman@18293
   422
lemma wild_pat [simp]: "branch wild_pat\<cdot>(unit_when\<cdot>r)\<cdot>x = return\<cdot>r"
huffman@18293
   423
by (simp add: branch_def wild_pat_def)
huffman@18293
   424
huffman@18293
   425
lemma as_pat [simp]:
huffman@18293
   426
  "branch (as_pat p)\<cdot>(csplit\<cdot>r)\<cdot>x = branch p\<cdot>(r\<cdot>x)\<cdot>x"
huffman@18293
   427
apply (simp add: branch_def as_pat_def)
huffman@18293
   428
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
huffman@18293
   429
done
huffman@18293
   430
huffman@18293
   431
lemma lazy_pat [simp]:
huffman@18293
   432
  "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>\<bottom>)"
huffman@18293
   433
  "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>\<bottom>)"
huffman@18293
   434
  "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>s"
huffman@18293
   435
apply (simp_all add: branch_def lazy_pat_def)
huffman@18293
   436
apply (rule_tac [!] p="p\<cdot>x" in maybeE, simp_all)
huffman@18293
   437
done
huffman@18293
   438
huffman@18112
   439
huffman@16221
   440
subsection {* Match functions for built-in types *}
huffman@16221
   441
wenzelm@36452
   442
default_sort pcpo
huffman@16776
   443
wenzelm@25131
   444
definition
huffman@30912
   445
  match_UU :: "'a \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
huffman@30912
   446
where
huffman@30912
   447
  "match_UU = strictify\<cdot>(\<Lambda> x k. fail)"
huffman@16776
   448
wenzelm@25131
   449
definition
huffman@30912
   450
  match_cpair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
huffman@30912
   451
where
huffman@30912
   452
  "match_cpair = (\<Lambda> x k. csplit\<cdot>k\<cdot>x)"
huffman@16221
   453
wenzelm@25131
   454
definition
huffman@30912
   455
  match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
huffman@30912
   456
where
huffman@30912
   457
  "match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)"
huffman@16551
   458
wenzelm@25131
   459
definition
huffman@30912
   460
  match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
huffman@30912
   461
where
huffman@30912
   462
  "match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)"
huffman@16551
   463
wenzelm@25131
   464
definition
huffman@30912
   465
  match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
huffman@30912
   466
where
huffman@30912
   467
  "match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)"
huffman@16551
   468
wenzelm@25131
   469
definition
huffman@30912
   470
  match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
huffman@30912
   471
where
huffman@30912
   472
  "match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)"
huffman@16221
   473
wenzelm@25131
   474
definition
huffman@30912
   475
  match_ONE :: "one \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
huffman@30912
   476
where
huffman@30912
   477
  "match_ONE = (\<Lambda> ONE k. k)"
huffman@30912
   478
huffman@30912
   479
definition
huffman@30912
   480
  match_TT :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
huffman@30912
   481
where
huffman@30912
   482
  "match_TT = (\<Lambda> x k. If x then k else fail fi)"
huffman@18094
   483
 
wenzelm@25131
   484
definition
huffman@30912
   485
  match_FF :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
huffman@30912
   486
where
huffman@30912
   487
  "match_FF = (\<Lambda> x k. If x then fail else k fi)"
huffman@16460
   488
huffman@16776
   489
lemma match_UU_simps [simp]:
huffman@31008
   490
  "match_UU\<cdot>\<bottom>\<cdot>k = \<bottom>"
huffman@31008
   491
  "x \<noteq> \<bottom> \<Longrightarrow> match_UU\<cdot>x\<cdot>k = fail"
huffman@31008
   492
by (simp_all add: match_UU_def)
huffman@16776
   493
huffman@16221
   494
lemma match_cpair_simps [simp]:
huffman@33401
   495
  "match_cpair\<cdot>(x, y)\<cdot>k = k\<cdot>x\<cdot>y"
huffman@33401
   496
by (simp_all add: match_cpair_def)
huffman@16221
   497
huffman@16551
   498
lemma match_spair_simps [simp]:
huffman@30912
   499
  "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> match_spair\<cdot>(:x, y:)\<cdot>k = k\<cdot>x\<cdot>y"
huffman@30912
   500
  "match_spair\<cdot>\<bottom>\<cdot>k = \<bottom>"
huffman@16551
   501
by (simp_all add: match_spair_def)
huffman@16551
   502
huffman@16551
   503
lemma match_sinl_simps [simp]:
huffman@30912
   504
  "x \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinl\<cdot>x)\<cdot>k = k\<cdot>x"
huffman@30914
   505
  "y \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinr\<cdot>y)\<cdot>k = fail"
huffman@30912
   506
  "match_sinl\<cdot>\<bottom>\<cdot>k = \<bottom>"
huffman@16551
   507
by (simp_all add: match_sinl_def)
huffman@16551
   508
huffman@16551
   509
lemma match_sinr_simps [simp]:
huffman@30912
   510
  "x \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinl\<cdot>x)\<cdot>k = fail"
huffman@30914
   511
  "y \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinr\<cdot>y)\<cdot>k = k\<cdot>y"
huffman@30912
   512
  "match_sinr\<cdot>\<bottom>\<cdot>k = \<bottom>"
huffman@16551
   513
by (simp_all add: match_sinr_def)
huffman@16551
   514
huffman@16221
   515
lemma match_up_simps [simp]:
huffman@30912
   516
  "match_up\<cdot>(up\<cdot>x)\<cdot>k = k\<cdot>x"
huffman@30912
   517
  "match_up\<cdot>\<bottom>\<cdot>k = \<bottom>"
huffman@16221
   518
by (simp_all add: match_up_def)
huffman@16221
   519
huffman@16460
   520
lemma match_ONE_simps [simp]:
huffman@30912
   521
  "match_ONE\<cdot>ONE\<cdot>k = k"
huffman@30912
   522
  "match_ONE\<cdot>\<bottom>\<cdot>k = \<bottom>"
huffman@18094
   523
by (simp_all add: match_ONE_def)
huffman@16460
   524
huffman@16460
   525
lemma match_TT_simps [simp]:
huffman@30912
   526
  "match_TT\<cdot>TT\<cdot>k = k"
huffman@30912
   527
  "match_TT\<cdot>FF\<cdot>k = fail"
huffman@30912
   528
  "match_TT\<cdot>\<bottom>\<cdot>k = \<bottom>"
huffman@18094
   529
by (simp_all add: match_TT_def)
huffman@16460
   530
huffman@16460
   531
lemma match_FF_simps [simp]:
huffman@30912
   532
  "match_FF\<cdot>FF\<cdot>k = k"
huffman@30912
   533
  "match_FF\<cdot>TT\<cdot>k = fail"
huffman@30912
   534
  "match_FF\<cdot>\<bottom>\<cdot>k = \<bottom>"
huffman@18094
   535
by (simp_all add: match_FF_def)
huffman@16460
   536
huffman@16401
   537
subsection {* Mutual recursion *}
huffman@16401
   538
huffman@16401
   539
text {*
huffman@16401
   540
  The following rules are used to prove unfolding theorems from
huffman@16401
   541
  fixed-point definitions of mutually recursive functions.
huffman@16401
   542
*}
huffman@16401
   543
huffman@31095
   544
lemma Pair_equalI: "\<lbrakk>x \<equiv> fst p; y \<equiv> snd p\<rbrakk> \<Longrightarrow> (x, y) \<equiv> p"
huffman@16401
   545
by simp
huffman@16401
   546
huffman@31095
   547
lemma Pair_eqD1: "(x, y) = (x', y') \<Longrightarrow> x = x'"
huffman@16401
   548
by simp
huffman@16401
   549
huffman@31095
   550
lemma Pair_eqD2: "(x, y) = (x', y') \<Longrightarrow> y = y'"
huffman@31095
   551
by simp
huffman@31095
   552
huffman@31095
   553
lemma def_cont_fix_eq:
huffman@31095
   554
  "\<lbrakk>f \<equiv> fix\<cdot>(Abs_CFun F); cont F\<rbrakk> \<Longrightarrow> f = F f"
huffman@31095
   555
by (simp, subst fix_eq, simp)
huffman@31095
   556
huffman@31095
   557
lemma def_cont_fix_ind:
huffman@31095
   558
  "\<lbrakk>f \<equiv> fix\<cdot>(Abs_CFun F); cont F; adm P; P \<bottom>; \<And>x. P x \<Longrightarrow> P (F x)\<rbrakk> \<Longrightarrow> P f"
huffman@31095
   559
by (simp add: fix_ind)
huffman@31095
   560
huffman@16463
   561
text {* lemma for proving rewrite rules *}
huffman@16463
   562
huffman@16463
   563
lemma ssubst_lhs: "\<lbrakk>t = s; P s = Q\<rbrakk> \<Longrightarrow> P t = Q"
huffman@16463
   564
by simp
huffman@16463
   565
huffman@16221
   566
huffman@16758
   567
subsection {* Initializing the fixrec package *}
huffman@16221
   568
huffman@35513
   569
use "Tools/holcf_library.ML"
haftmann@31738
   570
use "Tools/fixrec.ML"
huffman@16221
   571
haftmann@31738
   572
setup {* Fixrec.setup *}
huffman@30131
   573
huffman@30131
   574
setup {*
haftmann@31738
   575
  Fixrec.add_matchers
huffman@30131
   576
    [ (@{const_name up}, @{const_name match_up}),
huffman@30131
   577
      (@{const_name sinl}, @{const_name match_sinl}),
huffman@30131
   578
      (@{const_name sinr}, @{const_name match_sinr}),
huffman@30131
   579
      (@{const_name spair}, @{const_name match_spair}),
huffman@33401
   580
      (@{const_name Pair}, @{const_name match_cpair}),
huffman@30131
   581
      (@{const_name ONE}, @{const_name match_ONE}),
huffman@30131
   582
      (@{const_name TT}, @{const_name match_TT}),
huffman@31008
   583
      (@{const_name FF}, @{const_name match_FF}),
huffman@31008
   584
      (@{const_name UU}, @{const_name match_UU}) ]
huffman@30131
   585
*}
huffman@30131
   586
huffman@36989
   587
hide_const (open) return fail run cases
huffman@19104
   588
huffman@33414
   589
lemmas [fixrec_simp] =
huffman@33414
   590
  run_strict run_fail run_return
huffman@33418
   591
  mplus_strict mplus_fail mplus_return
huffman@33414
   592
  spair_strict_iff
huffman@33414
   593
  sinl_defined_iff
huffman@33414
   594
  sinr_defined_iff
huffman@33414
   595
  up_defined
huffman@33414
   596
  ONE_defined
huffman@33414
   597
  dist_eq_tr(1,2)
huffman@33414
   598
  match_UU_simps
huffman@33414
   599
  match_cpair_simps
huffman@33414
   600
  match_spair_simps
huffman@33414
   601
  match_sinl_simps
huffman@33414
   602
  match_sinr_simps
huffman@33414
   603
  match_up_simps
huffman@33414
   604
  match_ONE_simps
huffman@33414
   605
  match_TT_simps
huffman@33414
   606
  match_FF_simps
huffman@33414
   607
huffman@16221
   608
end