src/HOLCF/Fixrec.thy
author bulwahn
Fri, 12 Mar 2010 12:14:30 +0100
changeset 35756 cfde251d03a5
parent 35527 f4282471461d
child 35920 9ef9a20cfba1
permissions -rw-r--r--
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
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
    Author:     Amber Telfer and Brian Huffman
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     3
*)
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
header "Package for defining recursive functions in HOLCF"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     6
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
     7
theory Fixrec
16551
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
     8
imports Sprod Ssum Up One Tr Fix
35527
f4282471461d fixrec and repdef modules import holcf_library
huffman
parents: 35469
diff changeset
     9
uses
f4282471461d fixrec and repdef modules import holcf_library
huffman
parents: 35469
diff changeset
    10
  ("Tools/holcf_library.ML")
f4282471461d fixrec and repdef modules import holcf_library
huffman
parents: 35469
diff changeset
    11
  ("Tools/fixrec.ML")
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    12
begin
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    13
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    14
subsection {* Maybe monad type *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    15
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
    16
defaultsort cpo
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
    17
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    18
pcpodef (open) 'a maybe = "UNIV::(one ++ 'a u) set"
29063
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 28891
diff changeset
    19
by simp_all
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    20
29141
d5582ab1311f constdefs -> definition
huffman
parents: 29138
diff changeset
    21
definition
d5582ab1311f constdefs -> definition
huffman
parents: 29138
diff changeset
    22
  fail :: "'a maybe" where
d5582ab1311f constdefs -> definition
huffman
parents: 29138
diff changeset
    23
  "fail = Abs_maybe (sinl\<cdot>ONE)"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    24
29141
d5582ab1311f constdefs -> definition
huffman
parents: 29138
diff changeset
    25
definition
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    26
  return :: "'a \<rightarrow> 'a maybe" where
29141
d5582ab1311f constdefs -> definition
huffman
parents: 29138
diff changeset
    27
  "return = (\<Lambda> x. Abs_maybe (sinr\<cdot>(up\<cdot>x)))"
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    28
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    29
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    30
  maybe_when :: "'b \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a maybe \<rightarrow> 'b::pcpo" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    31
  "maybe_when = (\<Lambda> f r m. sscase\<cdot>(\<Lambda> x. f)\<cdot>(fup\<cdot>r)\<cdot>(Rep_maybe m))"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    32
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    33
lemma maybeE:
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    34
  "\<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
    35
apply (unfold fail_def return_def)
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    36
apply (cases p, rename_tac r)
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    37
apply (rule_tac p=r in ssumE, simp add: Abs_maybe_strict)
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    38
apply (rule_tac p=x in oneE, simp, simp)
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    39
apply (rule_tac p=y in upE, simp, simp add: cont_Abs_maybe)
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    40
done
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    41
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
    42
lemma return_defined [simp]: "return\<cdot>x \<noteq> \<bottom>"
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    43
by (simp add: return_def cont_Abs_maybe Abs_maybe_defined)
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
    44
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
    45
lemma fail_defined [simp]: "fail \<noteq> \<bottom>"
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    46
by (simp add: fail_def Abs_maybe_defined)
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
    47
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
    48
lemma return_eq [simp]: "(return\<cdot>x = return\<cdot>y) = (x = y)"
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    49
by (simp add: return_def cont_Abs_maybe Abs_maybe_inject)
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
    50
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
    51
lemma return_neq_fail [simp]:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
    52
  "return\<cdot>x \<noteq> fail" "fail \<noteq> return\<cdot>x"
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    53
by (simp_all add: return_def fail_def cont_Abs_maybe Abs_maybe_inject)
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    54
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    55
lemma maybe_when_rews [simp]:
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    56
  "maybe_when\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>"
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    57
  "maybe_when\<cdot>f\<cdot>r\<cdot>fail = f"
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    58
  "maybe_when\<cdot>f\<cdot>r\<cdot>(return\<cdot>x) = r\<cdot>x"
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    59
by (simp_all add: return_def fail_def maybe_when_def cont_Rep_maybe
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29322
diff changeset
    60
                  cont2cont_LAM
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    61
                  cont_Abs_maybe Abs_maybe_inverse Rep_maybe_strict)
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    62
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    63
translations
29141
d5582ab1311f constdefs -> definition
huffman
parents: 29138
diff changeset
    64
  "case m of XCONST fail \<Rightarrow> t1 | XCONST return\<cdot>x \<Rightarrow> t2"
d5582ab1311f constdefs -> definition
huffman
parents: 29138
diff changeset
    65
    == "CONST maybe_when\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
    66
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
    67
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
    68
subsubsection {* Monadic bind operator *}
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    69
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    70
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    71
  bind :: "'a maybe \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> 'b maybe" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    72
  "bind = (\<Lambda> m f. case m of fail \<Rightarrow> fail | return\<cdot>x \<Rightarrow> f\<cdot>x)"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    73
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    74
text {* monad laws *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    75
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    76
lemma bind_strict [simp]: "bind\<cdot>\<bottom>\<cdot>f = \<bottom>"
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    77
by (simp add: bind_def)
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    78
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    79
lemma bind_fail [simp]: "bind\<cdot>fail\<cdot>f = fail"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    80
by (simp add: bind_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    81
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    82
lemma left_unit [simp]: "bind\<cdot>(return\<cdot>a)\<cdot>k = k\<cdot>a"
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    83
by (simp add: bind_def)
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    84
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    85
lemma right_unit [simp]: "bind\<cdot>m\<cdot>return = m"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    86
by (rule_tac p=m in maybeE, simp_all)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    87
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    88
lemma bind_assoc:
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
    89
 "bind\<cdot>(bind\<cdot>m\<cdot>k)\<cdot>h = bind\<cdot>m\<cdot>(\<Lambda> a. bind\<cdot>(k\<cdot>a)\<cdot>h)"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    90
by (rule_tac p=m in maybeE, simp_all)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    91
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
    92
subsubsection {* Run operator *}
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
    93
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    94
definition
28891
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
    95
  run :: "'a maybe \<rightarrow> 'a::pcpo" where
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    96
  "run = maybe_when\<cdot>\<bottom>\<cdot>ID"
16221
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
text {* rewrite rules for run *}
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
lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   101
by (simp add: run_def)
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
lemma run_fail [simp]: "run\<cdot>fail = \<bottom>"
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
   104
by (simp add: run_def)
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   105
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   106
lemma run_return [simp]: "run\<cdot>(return\<cdot>x) = x"
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
   107
by (simp add: run_def)
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   108
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   109
subsubsection {* Monad plus operator *}
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   110
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   111
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   112
  mplus :: "'a maybe \<rightarrow> 'a maybe \<rightarrow> 'a maybe" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   113
  "mplus = (\<Lambda> m1 m2. case m1 of fail \<Rightarrow> m2 | return\<cdot>x \<Rightarrow> m1)"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   114
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   115
abbreviation
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   116
  mplus_syn :: "['a maybe, 'a maybe] \<Rightarrow> 'a maybe"  (infixr "+++" 65)  where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   117
  "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2"
16221
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 {* rewrite rules for mplus *}
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
lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   122
by (simp add: mplus_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   123
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   124
lemma mplus_fail [simp]: "fail +++ m = m"
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
   125
by (simp add: mplus_def)
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   126
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   127
lemma mplus_return [simp]: "return\<cdot>x +++ m = return\<cdot>x"
19092
e32cf29f01fc make maybe into a real type constructor; remove monad syntax
huffman
parents: 18293
diff changeset
   128
by (simp add: mplus_def)
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   129
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   130
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
   131
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
   132
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   133
lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   134
by (rule_tac p=x in maybeE, simp_all)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   135
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   136
subsubsection {* Fatbar combinator *}
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   137
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   138
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   139
  fatbar :: "('a \<rightarrow> 'b maybe) \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> ('a \<rightarrow> 'b maybe)" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   140
  "fatbar = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   141
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   142
abbreviation
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   143
  fatbar_syn :: "['a \<rightarrow> 'b maybe, 'a \<rightarrow> 'b maybe] \<Rightarrow> 'a \<rightarrow> 'b maybe" (infixr "\<parallel>" 60)  where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   144
  "m1 \<parallel> m2 == fatbar\<cdot>m1\<cdot>m2"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   145
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   146
lemma fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> (m \<parallel> ms)\<cdot>x = \<bottom>"
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   147
by (simp add: fatbar_def)
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   148
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   149
lemma fatbar2: "m\<cdot>x = fail \<Longrightarrow> (m \<parallel> ms)\<cdot>x = ms\<cdot>x"
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   150
by (simp add: fatbar_def)
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   151
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   152
lemma fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> (m \<parallel> ms)\<cdot>x = return\<cdot>y"
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   153
by (simp add: fatbar_def)
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   154
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   155
lemmas fatbar_simps = fatbar1 fatbar2 fatbar3
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   156
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   157
lemma run_fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = \<bottom>"
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   158
by (simp add: fatbar_def)
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   159
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   160
lemma run_fatbar2: "m\<cdot>x = fail \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = run\<cdot>(ms\<cdot>x)"
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   161
by (simp add: fatbar_def)
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   162
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   163
lemma run_fatbar3: "m\<cdot>x = return\<cdot>y \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = y"
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   164
by (simp add: fatbar_def)
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   165
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   166
lemmas run_fatbar_simps [simp] = run_fatbar1 run_fatbar2 run_fatbar3
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   167
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   168
subsection {* Case branch combinator *}
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   169
29141
d5582ab1311f constdefs -> definition
huffman
parents: 29138
diff changeset
   170
definition
d5582ab1311f constdefs -> definition
huffman
parents: 29138
diff changeset
   171
  branch :: "('a \<rightarrow> 'b maybe) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c maybe)" where
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   172
  "branch p \<equiv> \<Lambda> r x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> y. return\<cdot>(r\<cdot>y))"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   173
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   174
lemma branch_rews:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   175
  "p\<cdot>x = \<bottom> \<Longrightarrow> branch p\<cdot>r\<cdot>x = \<bottom>"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   176
  "p\<cdot>x = fail \<Longrightarrow> branch p\<cdot>r\<cdot>x = fail"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   177
  "p\<cdot>x = return\<cdot>y \<Longrightarrow> branch p\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>y)"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   178
by (simp_all add: branch_def)
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   179
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   180
lemma branch_return [simp]: "branch return\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>x)"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   181
by (simp add: branch_def)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   182
28891
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   183
subsubsection {* Cases operator *}
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   184
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   185
definition
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   186
  cases :: "'a maybe \<rightarrow> 'a::pcpo" where
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   187
  "cases = maybe_when\<cdot>\<bottom>\<cdot>ID"
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   188
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   189
text {* rewrite rules for cases *}
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   190
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   191
lemma cases_strict [simp]: "cases\<cdot>\<bottom> = \<bottom>"
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   192
by (simp add: cases_def)
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   193
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   194
lemma cases_fail [simp]: "cases\<cdot>fail = \<bottom>"
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   195
by (simp add: cases_def)
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   196
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   197
lemma cases_return [simp]: "cases\<cdot>(return\<cdot>x) = x"
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   198
by (simp add: cases_def)
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   199
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   200
subsection {* Case syntax *}
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   201
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   202
nonterminals
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   203
  Case_syn  Cases_syn
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   204
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   205
syntax
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   206
  "_Case_syntax":: "['a, Cases_syn] => 'b"               ("(Case _ of/ _)" 10)
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   207
  "_Case1"      :: "['a, 'b] => Case_syn"                ("(2_ =>/ _)" 10)
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   208
  ""            :: "Case_syn => Cases_syn"               ("_")
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   209
  "_Case2"      :: "[Case_syn, Cases_syn] => Cases_syn"  ("_/ | _")
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   210
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   211
syntax (xsymbols)
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   212
  "_Case1"      :: "['a, 'b] => Case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   213
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   214
translations
28891
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   215
  "_Case_syntax x ms" == "CONST Fixrec.cases\<cdot>(ms\<cdot>x)"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   216
  "_Case2 m ms" == "m \<parallel> ms"
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   217
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   218
text {* Parsing Case expressions *}
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   219
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   220
syntax
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   221
  "_pat" :: "'a"
29322
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   222
  "_variable" :: "'a"
26046
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   223
  "_noargs" :: "'a"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   224
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   225
translations
29322
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   226
  "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_variable p r)"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   227
  "_variable (_args x y) r" => "CONST csplit\<cdot>(_variable x (_variable y r))"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   228
  "_variable _noargs r" => "CONST unit_when\<cdot>r"
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   229
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   230
parse_translation {*
35115
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   231
(* rewrite (_pat x) => (return) *)
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   232
(* rewrite (_variable x t) => (Abs_CFun (%x. t)) *)
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   233
 [(@{syntax_const "_pat"}, fn _ => Syntax.const @{const_syntax Fixrec.return}),
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   234
  mk_binder_tr (@{syntax_const "_variable"}, @{const_syntax Abs_CFun})];
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   235
*}
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   236
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   237
text {* Printing Case expressions *}
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   238
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   239
syntax
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   240
  "_match" :: "'a"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   241
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   242
print_translation {*
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   243
  let
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   244
    fun dest_LAM (Const (@{const_syntax Rep_CFun},_) $ Const (@{const_syntax unit_when},_) $ t) =
35115
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   245
          (Syntax.const @{syntax_const "_noargs"}, t)
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   246
    |   dest_LAM (Const (@{const_syntax Rep_CFun},_) $ Const (@{const_syntax csplit},_) $ t) =
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   247
          let
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   248
            val (v1, t1) = dest_LAM t;
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   249
            val (v2, t2) = dest_LAM t1;
35115
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   250
          in (Syntax.const @{syntax_const "_args"} $ v1 $ v2, t2) end
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   251
    |   dest_LAM (Const (@{const_syntax Abs_CFun},_) $ t) =
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   252
          let
35115
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   253
            val abs =
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   254
              case t of Abs abs => abs
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   255
                | _ => ("x", dummyT, incr_boundvars 1 t $ Bound 0);
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   256
            val (x, t') = atomic_abs_tr' abs;
35115
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   257
          in (Syntax.const @{syntax_const "_variable"} $ x, t') end
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   258
    |   dest_LAM _ = raise Match; (* too few vars: abort translation *)
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   259
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   260
    fun Case1_tr' [Const(@{const_syntax branch},_) $ p, r] =
35115
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   261
          let val (v, t) = dest_LAM r in
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   262
            Syntax.const @{syntax_const "_Case1"} $
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   263
              (Syntax.const @{syntax_const "_match"} $ p $ v) $ t
446c5063e4fd modernized translations;
wenzelm
parents: 33429
diff changeset
   264
          end;
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   265
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   266
  in [(@{const_syntax Rep_CFun}, Case1_tr')] end;
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   267
*}
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   268
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   269
translations
35469
6e59de61d501 fix output translation for Case syntax
huffman
parents: 35115
diff changeset
   270
  "x" <= "_match (CONST Fixrec.return) (_variable x)"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   271
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   272
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   273
subsection {* Pattern combinators for data constructors *}
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   274
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   275
types ('a, 'b) pat = "'a \<rightarrow> 'b maybe"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   276
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   277
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   278
  cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   279
  "cpair_pat p1 p2 = (\<Lambda>\<langle>x, y\<rangle>.
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   280
    bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. return\<cdot>\<langle>a, b\<rangle>)))"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   281
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   282
definition
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   283
  spair_pat ::
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   284
  "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a::pcpo \<otimes> 'b::pcpo, 'c \<times> 'd) pat" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   285
  "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>\<langle>x, y\<rangle>)"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   286
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   287
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   288
  sinl_pat :: "('a, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   289
  "sinl_pat p = sscase\<cdot>p\<cdot>(\<Lambda> x. fail)"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   290
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   291
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   292
  sinr_pat :: "('b, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   293
  "sinr_pat p = sscase\<cdot>(\<Lambda> x. fail)\<cdot>p"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   294
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   295
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   296
  up_pat :: "('a, 'b) pat \<Rightarrow> ('a u, 'b) pat" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   297
  "up_pat p = fup\<cdot>p"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   298
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   299
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   300
  TT_pat :: "(tr, unit) pat" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   301
  "TT_pat = (\<Lambda> b. If b then return\<cdot>() else fail fi)"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   302
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   303
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   304
  FF_pat :: "(tr, unit) pat" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   305
  "FF_pat = (\<Lambda> b. If b then fail else return\<cdot>() fi)"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   306
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   307
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   308
  ONE_pat :: "(one, unit) pat" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   309
  "ONE_pat = (\<Lambda> ONE. return\<cdot>())"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   310
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   311
text {* Parse translations (patterns) *}
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   312
translations
26046
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   313
  "_pat (XCONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   314
  "_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   315
  "_pat (XCONST sinl\<cdot>x)" => "CONST sinl_pat (_pat x)"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   316
  "_pat (XCONST sinr\<cdot>x)" => "CONST sinr_pat (_pat x)"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   317
  "_pat (XCONST up\<cdot>x)" => "CONST up_pat (_pat x)"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   318
  "_pat (XCONST TT)" => "CONST TT_pat"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   319
  "_pat (XCONST FF)" => "CONST FF_pat"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   320
  "_pat (XCONST ONE)" => "CONST ONE_pat"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   321
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   322
text {* CONST version is also needed for constructors with special syntax *}
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   323
translations
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   324
  "_pat (CONST cpair\<cdot>x\<cdot>y)" => "CONST cpair_pat (_pat x) (_pat y)"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   325
  "_pat (CONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   326
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   327
text {* Parse translations (variables) *}
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   328
translations
29322
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   329
  "_variable (XCONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   330
  "_variable (XCONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   331
  "_variable (XCONST sinl\<cdot>x) r" => "_variable x r"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   332
  "_variable (XCONST sinr\<cdot>x) r" => "_variable x r"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   333
  "_variable (XCONST up\<cdot>x) r" => "_variable x r"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   334
  "_variable (XCONST TT) r" => "_variable _noargs r"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   335
  "_variable (XCONST FF) r" => "_variable _noargs r"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   336
  "_variable (XCONST ONE) r" => "_variable _noargs r"
26046
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   337
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   338
translations
29322
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   339
  "_variable (CONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   340
  "_variable (CONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   341
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   342
text {* Print translations *}
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   343
translations
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   344
  "CONST cpair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)"
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   345
      <= "_match (CONST cpair_pat p1 p2) (_args v1 v2)"
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   346
  "CONST spair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)"
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   347
      <= "_match (CONST spair_pat p1 p2) (_args v1 v2)"
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   348
  "CONST sinl\<cdot>(_match p1 v1)" <= "_match (CONST sinl_pat p1) v1"
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   349
  "CONST sinr\<cdot>(_match p1 v1)" <= "_match (CONST sinr_pat p1) v1"
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   350
  "CONST up\<cdot>(_match p1 v1)" <= "_match (CONST up_pat p1) v1"
26046
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   351
  "CONST TT" <= "_match (CONST TT_pat) _noargs"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   352
  "CONST FF" <= "_match (CONST FF_pat) _noargs"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   353
  "CONST ONE" <= "_match (CONST ONE_pat) _noargs"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   354
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   355
lemma cpair_pat1:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   356
  "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = \<bottom>"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   357
apply (simp add: branch_def cpair_pat_def)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   358
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   359
done
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   360
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   361
lemma cpair_pat2:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   362
  "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = fail"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   363
apply (simp add: branch_def cpair_pat_def)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   364
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   365
done
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   366
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   367
lemma cpair_pat3:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   368
  "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow>
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   369
   branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>\<langle>x, y\<rangle> = branch q\<cdot>s\<cdot>y"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   370
apply (simp add: branch_def cpair_pat_def)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   371
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   372
apply (rule_tac p="q\<cdot>y" in maybeE, simp_all)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   373
done
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   374
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   375
lemmas cpair_pat [simp] =
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   376
  cpair_pat1 cpair_pat2 cpair_pat3
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   377
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   378
lemma spair_pat [simp]:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   379
  "branch (spair_pat p1 p2)\<cdot>r\<cdot>\<bottom> = \<bottom>"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   380
  "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk>
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   381
     \<Longrightarrow> branch (spair_pat p1 p2)\<cdot>r\<cdot>(:x, y:) =
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   382
         branch (cpair_pat p1 p2)\<cdot>r\<cdot>\<langle>x, y\<rangle>"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   383
by (simp_all add: branch_def spair_pat_def)
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   384
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   385
lemma sinl_pat [simp]:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   386
  "branch (sinl_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   387
  "x \<noteq> \<bottom> \<Longrightarrow> branch (sinl_pat p)\<cdot>r\<cdot>(sinl\<cdot>x) = branch p\<cdot>r\<cdot>x"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   388
  "y \<noteq> \<bottom> \<Longrightarrow> branch (sinl_pat p)\<cdot>r\<cdot>(sinr\<cdot>y) = fail"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   389
by (simp_all add: branch_def sinl_pat_def)
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   390
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   391
lemma sinr_pat [simp]:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   392
  "branch (sinr_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   393
  "x \<noteq> \<bottom> \<Longrightarrow> branch (sinr_pat p)\<cdot>r\<cdot>(sinl\<cdot>x) = fail"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   394
  "y \<noteq> \<bottom> \<Longrightarrow> branch (sinr_pat p)\<cdot>r\<cdot>(sinr\<cdot>y) = branch p\<cdot>r\<cdot>y"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   395
by (simp_all add: branch_def sinr_pat_def)
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   396
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   397
lemma up_pat [simp]:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   398
  "branch (up_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   399
  "branch (up_pat p)\<cdot>r\<cdot>(up\<cdot>x) = branch p\<cdot>r\<cdot>x"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   400
by (simp_all add: branch_def up_pat_def)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   401
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   402
lemma TT_pat [simp]:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   403
  "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   404
  "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = return\<cdot>r"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   405
  "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = fail"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   406
by (simp_all add: branch_def TT_pat_def)
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   407
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   408
lemma FF_pat [simp]:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   409
  "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   410
  "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = fail"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   411
  "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = return\<cdot>r"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   412
by (simp_all add: branch_def FF_pat_def)
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   413
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   414
lemma ONE_pat [simp]:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   415
  "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   416
  "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>ONE = return\<cdot>r"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   417
by (simp_all add: branch_def ONE_pat_def)
18097
d196d84c306f add case syntax stuff
huffman
parents: 18096
diff changeset
   418
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   419
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   420
subsection {* Wildcards, as-patterns, and lazy patterns *}
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   421
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   422
syntax
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   423
  "_as_pat" :: "[idt, 'a] \<Rightarrow> 'a" (infixr "\<as>" 10)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   424
  "_lazy_pat" :: "'a \<Rightarrow> 'a" ("\<lazy> _" [1000] 1000)
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   425
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   426
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   427
  wild_pat :: "'a \<rightarrow> unit maybe" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   428
  "wild_pat = (\<Lambda> x. return\<cdot>())"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   429
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   430
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   431
  as_pat :: "('a \<rightarrow> 'b maybe) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) maybe" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   432
  "as_pat p = (\<Lambda> x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> a. return\<cdot>\<langle>x, a\<rangle>))"
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   433
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   434
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   435
  lazy_pat :: "('a \<rightarrow> 'b::pcpo maybe) \<Rightarrow> ('a \<rightarrow> 'b maybe)" where
28891
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   436
  "lazy_pat p = (\<Lambda> x. return\<cdot>(cases\<cdot>(p\<cdot>x)))"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   437
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   438
text {* Parse translations (patterns) *}
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   439
translations
26046
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   440
  "_pat _" => "CONST wild_pat"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   441
  "_pat (_as_pat x y)" => "CONST as_pat (_pat y)"
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   442
  "_pat (_lazy_pat x)" => "CONST lazy_pat (_pat x)"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   443
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   444
text {* Parse translations (variables) *}
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   445
translations
29322
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   446
  "_variable _ r" => "_variable _noargs r"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   447
  "_variable (_as_pat x y) r" => "_variable (_args x y) r"
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   448
  "_variable (_lazy_pat x) r" => "_variable x r"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   449
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   450
text {* Print translations *}
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   451
translations
26046
1624b3304bb9 fix broken syntax translations
huffman
parents: 25158
diff changeset
   452
  "_" <= "_match (CONST wild_pat) _noargs"
29322
ae6f2b1559fa renamed syntax constant "_var" to "_variable", to avoid clash with internal token marker;
wenzelm
parents: 29141
diff changeset
   453
  "_as_pat x (_match p v)" <= "_match (CONST as_pat p) (_args (_variable x) v)"
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   454
  "_lazy_pat (_match p v)" <= "_match (CONST lazy_pat p) v"
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   455
19327
4565e230e6eb lazy patterns in lambda abstractions
huffman
parents: 19104
diff changeset
   456
text {* Lazy patterns in lambda abstractions *}
4565e230e6eb lazy patterns in lambda abstractions
huffman
parents: 19104
diff changeset
   457
translations
28891
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   458
  "_cabs (_lazy_pat p) r" == "CONST Fixrec.cases oo (_Case1 (_lazy_pat p) r)"
19327
4565e230e6eb lazy patterns in lambda abstractions
huffman
parents: 19104
diff changeset
   459
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   460
lemma wild_pat [simp]: "branch wild_pat\<cdot>(unit_when\<cdot>r)\<cdot>x = return\<cdot>r"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   461
by (simp add: branch_def wild_pat_def)
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   462
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   463
lemma as_pat [simp]:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   464
  "branch (as_pat p)\<cdot>(csplit\<cdot>r)\<cdot>x = branch p\<cdot>(r\<cdot>x)\<cdot>x"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   465
apply (simp add: branch_def as_pat_def)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   466
apply (rule_tac p="p\<cdot>x" in maybeE, simp_all)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   467
done
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   468
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   469
lemma lazy_pat [simp]:
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   470
  "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>\<bottom>)"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   471
  "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>(r\<cdot>\<bottom>)"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   472
  "branch p\<cdot>r\<cdot>x = return\<cdot>s \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = return\<cdot>s"
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   473
apply (simp_all add: branch_def lazy_pat_def)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   474
apply (rule_tac [!] p="p\<cdot>x" in maybeE, simp_all)
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   475
done
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18112
diff changeset
   476
18112
dc1d6f588204 reimplemented Case syntax using print/parse translations; moved as-patterns to separate section
huffman
parents: 18110
diff changeset
   477
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   478
subsection {* Match functions for built-in types *}
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   479
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   480
defaultsort pcpo
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   481
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   482
definition
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   483
  match_UU :: "'a \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   484
where
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   485
  "match_UU = strictify\<cdot>(\<Lambda> x k. fail)"
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   486
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   487
definition
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   488
  match_cpair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   489
where
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   490
  "match_cpair = (\<Lambda> x k. csplit\<cdot>k\<cdot>x)"
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   491
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   492
definition
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   493
  match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   494
where
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   495
  "match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   496
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   497
definition
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   498
  match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   499
where
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   500
  "match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)"
16551
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   501
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   502
definition
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   503
  match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   504
where
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   505
  "match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)"
16551
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   506
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   507
definition
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   508
  match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c maybe) \<rightarrow> 'c maybe"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   509
where
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   510
  "match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   511
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   512
definition
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   513
  match_ONE :: "one \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   514
where
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   515
  "match_ONE = (\<Lambda> ONE k. k)"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   516
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   517
definition
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   518
  match_TT :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   519
where
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   520
  "match_TT = (\<Lambda> x k. If x then k else fail fi)"
18094
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   521
 
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   522
definition
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   523
  match_FF :: "tr \<rightarrow> 'c maybe \<rightarrow> 'c maybe"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   524
where
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   525
  "match_FF = (\<Lambda> x k. If x then fail else k fi)"
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   526
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   527
lemma match_UU_simps [simp]:
31008
b8f4e351b5bf add proper support for bottom-patterns in fixrec package
huffman
parents: 30914
diff changeset
   528
  "match_UU\<cdot>\<bottom>\<cdot>k = \<bottom>"
b8f4e351b5bf add proper support for bottom-patterns in fixrec package
huffman
parents: 30914
diff changeset
   529
  "x \<noteq> \<bottom> \<Longrightarrow> match_UU\<cdot>x\<cdot>k = fail"
b8f4e351b5bf add proper support for bottom-patterns in fixrec package
huffman
parents: 30914
diff changeset
   530
by (simp_all add: match_UU_def)
16776
a3899ac14a1c generalized types of monadic operators to class cpo; added match function for UU
huffman
parents: 16758
diff changeset
   531
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   532
lemma match_cpair_simps [simp]:
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   533
  "match_cpair\<cdot>\<langle>x, y\<rangle>\<cdot>k = k\<cdot>x\<cdot>y"
33401
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 31738
diff changeset
   534
  "match_cpair\<cdot>(x, y)\<cdot>k = k\<cdot>x\<cdot>y"
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 31738
diff changeset
   535
by (simp_all add: match_cpair_def)
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   536
16551
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   537
lemma match_spair_simps [simp]:
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   538
  "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> match_spair\<cdot>(:x, y:)\<cdot>k = k\<cdot>x\<cdot>y"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   539
  "match_spair\<cdot>\<bottom>\<cdot>k = \<bottom>"
16551
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   540
by (simp_all add: match_spair_def)
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   541
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   542
lemma match_sinl_simps [simp]:
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   543
  "x \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinl\<cdot>x)\<cdot>k = k\<cdot>x"
30914
ceeb5f2eac75 fix too-specific types in lemmas match_{sinl,sinr}_simps
huffman
parents: 30912
diff changeset
   544
  "y \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinr\<cdot>y)\<cdot>k = fail"
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   545
  "match_sinl\<cdot>\<bottom>\<cdot>k = \<bottom>"
16551
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   546
by (simp_all add: match_sinl_def)
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   547
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   548
lemma match_sinr_simps [simp]:
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   549
  "x \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinl\<cdot>x)\<cdot>k = fail"
30914
ceeb5f2eac75 fix too-specific types in lemmas match_{sinl,sinr}_simps
huffman
parents: 30912
diff changeset
   550
  "y \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinr\<cdot>y)\<cdot>k = k\<cdot>y"
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   551
  "match_sinr\<cdot>\<bottom>\<cdot>k = \<bottom>"
16551
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   552
by (simp_all add: match_sinr_def)
7abf8a713613 added match functions for spair, sinl, sinr
huffman
parents: 16463
diff changeset
   553
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   554
lemma match_up_simps [simp]:
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   555
  "match_up\<cdot>(up\<cdot>x)\<cdot>k = k\<cdot>x"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   556
  "match_up\<cdot>\<bottom>\<cdot>k = \<bottom>"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   557
by (simp_all add: match_up_def)
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   558
16460
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   559
lemma match_ONE_simps [simp]:
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   560
  "match_ONE\<cdot>ONE\<cdot>k = k"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   561
  "match_ONE\<cdot>\<bottom>\<cdot>k = \<bottom>"
18094
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   562
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
   563
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   564
lemma match_TT_simps [simp]:
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   565
  "match_TT\<cdot>TT\<cdot>k = k"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   566
  "match_TT\<cdot>FF\<cdot>k = fail"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   567
  "match_TT\<cdot>\<bottom>\<cdot>k = \<bottom>"
18094
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   568
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
   569
72a08d509d62 added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents: 16417
diff changeset
   570
lemma match_FF_simps [simp]:
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   571
  "match_FF\<cdot>FF\<cdot>k = k"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   572
  "match_FF\<cdot>TT\<cdot>k = fail"
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30131
diff changeset
   573
  "match_FF\<cdot>\<bottom>\<cdot>k = \<bottom>"
18094
404f298220af simplify definitions
huffman
parents: 16779
diff changeset
   574
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
   575
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   576
subsection {* Mutual recursion *}
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   577
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   578
text {*
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   579
  The following rules are used to prove unfolding theorems from
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   580
  fixed-point definitions of mutually recursive functions.
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   581
*}
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   582
31095
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   583
lemma Pair_equalI: "\<lbrakk>x \<equiv> fst p; y \<equiv> snd p\<rbrakk> \<Longrightarrow> (x, y) \<equiv> p"
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   584
by simp
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   585
31095
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   586
lemma Pair_eqD1: "(x, y) = (x', y') \<Longrightarrow> x = x'"
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   587
by simp
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   588
31095
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   589
lemma Pair_eqD2: "(x, y) = (x', y') \<Longrightarrow> y = y'"
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   590
by simp
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16229
diff changeset
   591
31095
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   592
lemma def_cont_fix_eq:
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   593
  "\<lbrakk>f \<equiv> fix\<cdot>(Abs_CFun F); cont F\<rbrakk> \<Longrightarrow> f = F f"
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   594
by (simp, subst fix_eq, simp)
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   595
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   596
lemma def_cont_fix_ind:
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   597
  "\<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"
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   598
by (simp add: fix_ind)
b79d140f6d0b simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents: 31008
diff changeset
   599
16463
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16460
diff changeset
   600
text {* lemma for proving rewrite rules *}
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16460
diff changeset
   601
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16460
diff changeset
   602
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
   603
by simp
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16460
diff changeset
   604
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   605
16758
c32334d98fcd fix typo
huffman
parents: 16754
diff changeset
   606
subsection {* Initializing the fixrec package *}
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   607
35527
f4282471461d fixrec and repdef modules import holcf_library
huffman
parents: 35469
diff changeset
   608
use "Tools/holcf_library.ML"
31738
7b9b9ba532ca discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31095
diff changeset
   609
use "Tools/fixrec.ML"
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   610
31738
7b9b9ba532ca discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31095
diff changeset
   611
setup {* Fixrec.setup *}
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   612
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   613
setup {*
31738
7b9b9ba532ca discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31095
diff changeset
   614
  Fixrec.add_matchers
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   615
    [ (@{const_name up}, @{const_name match_up}),
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   616
      (@{const_name sinl}, @{const_name match_sinl}),
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   617
      (@{const_name sinr}, @{const_name match_sinr}),
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   618
      (@{const_name spair}, @{const_name match_spair}),
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   619
      (@{const_name cpair}, @{const_name match_cpair}),
33401
fc43fa403a69 add fixrec support for HOL pair constructor patterns
huffman
parents: 31738
diff changeset
   620
      (@{const_name Pair}, @{const_name match_cpair}),
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   621
      (@{const_name ONE}, @{const_name match_ONE}),
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   622
      (@{const_name TT}, @{const_name match_TT}),
31008
b8f4e351b5bf add proper support for bottom-patterns in fixrec package
huffman
parents: 30914
diff changeset
   623
      (@{const_name FF}, @{const_name match_FF}),
b8f4e351b5bf add proper support for bottom-patterns in fixrec package
huffman
parents: 30914
diff changeset
   624
      (@{const_name UU}, @{const_name match_UU}) ]
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   625
*}
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29530
diff changeset
   626
28891
f199def7a6a5 separate run and cases combinators
huffman
parents: 26046
diff changeset
   627
hide (open) const return bind fail run cases
19104
7d69b6d7b8f1 use qualified name for return
huffman
parents: 19092
diff changeset
   628
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   629
lemmas [fixrec_simp] =
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   630
  run_strict run_fail run_return
33429
42d7b6b4992b add more fixrec_simp rules
huffman
parents: 33425
diff changeset
   631
  mplus_strict mplus_fail mplus_return
33425
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   632
  spair_strict_iff
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   633
  sinl_defined_iff
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   634
  sinr_defined_iff
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   635
  up_defined
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   636
  ONE_defined
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   637
  dist_eq_tr(1,2)
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   638
  match_UU_simps
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   639
  match_cpair_simps
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   640
  match_spair_simps
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   641
  match_sinl_simps
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   642
  match_sinr_simps
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   643
  match_up_simps
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   644
  match_ONE_simps
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   645
  match_TT_simps
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   646
  match_FF_simps
7e4f3c66190d add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents: 33401
diff changeset
   647
16221
879400e029bf New theory with lemmas for the fixrec package
huffman
parents:
diff changeset
   648
end