src/HOLCF/Fixrec.thy
changeset 37109 e67760c1b851
parent 37108 00f13d3ad474
child 39807 19ddbededdd3
equal deleted inserted replaced
37108:00f13d3ad474 37109:e67760c1b851
    62 
    62 
    63 translations
    63 translations
    64   "case m of XCONST fail \<Rightarrow> t1 | XCONST succeed\<cdot>x \<Rightarrow> t2"
    64   "case m of XCONST fail \<Rightarrow> t1 | XCONST succeed\<cdot>x \<Rightarrow> t2"
    65     == "CONST match_case\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m"
    65     == "CONST match_case\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m"
    66 
    66 
    67 
       
    68 subsubsection {* Run operator *}
    67 subsubsection {* Run operator *}
    69 
    68 
    70 definition
    69 definition
    71   run :: "'a match \<rightarrow> 'a::pcpo" where
    70   run :: "'a match \<rightarrow> 'a::pcpo" where
    72   "run = match_case\<cdot>\<bottom>\<cdot>ID"
    71   "run = match_case\<cdot>\<bottom>\<cdot>ID"
   106 lemma mplus_fail2 [simp]: "m +++ fail = m"
   105 lemma mplus_fail2 [simp]: "m +++ fail = m"
   107 by (cases m, simp_all)
   106 by (cases m, simp_all)
   108 
   107 
   109 lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
   108 lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
   110 by (cases x, simp_all)
   109 by (cases x, simp_all)
   111 
       
   112 subsubsection {* Fatbar combinator *}
       
   113 
       
   114 definition
       
   115   fatbar :: "('a \<rightarrow> 'b match) \<rightarrow> ('a \<rightarrow> 'b match) \<rightarrow> ('a \<rightarrow> 'b match)" where
       
   116   "fatbar = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)"
       
   117 
       
   118 abbreviation
       
   119   fatbar_syn :: "['a \<rightarrow> 'b match, 'a \<rightarrow> 'b match] \<Rightarrow> 'a \<rightarrow> 'b match" (infixr "\<parallel>" 60)  where
       
   120   "m1 \<parallel> m2 == fatbar\<cdot>m1\<cdot>m2"
       
   121 
       
   122 lemma fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> (m \<parallel> ms)\<cdot>x = \<bottom>"
       
   123 by (simp add: fatbar_def)
       
   124 
       
   125 lemma fatbar2: "m\<cdot>x = fail \<Longrightarrow> (m \<parallel> ms)\<cdot>x = ms\<cdot>x"
       
   126 by (simp add: fatbar_def)
       
   127 
       
   128 lemma fatbar3: "m\<cdot>x = succeed\<cdot>y \<Longrightarrow> (m \<parallel> ms)\<cdot>x = succeed\<cdot>y"
       
   129 by (simp add: fatbar_def)
       
   130 
       
   131 lemmas fatbar_simps = fatbar1 fatbar2 fatbar3
       
   132 
       
   133 lemma run_fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = \<bottom>"
       
   134 by (simp add: fatbar_def)
       
   135 
       
   136 lemma run_fatbar2: "m\<cdot>x = fail \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = run\<cdot>(ms\<cdot>x)"
       
   137 by (simp add: fatbar_def)
       
   138 
       
   139 lemma run_fatbar3: "m\<cdot>x = succeed\<cdot>y \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = y"
       
   140 by (simp add: fatbar_def)
       
   141 
       
   142 lemmas run_fatbar_simps [simp] = run_fatbar1 run_fatbar2 run_fatbar3
       
   143 
       
   144 subsection {* Case branch combinator *}
       
   145 
       
   146 definition
       
   147   branch :: "('a \<rightarrow> 'b match) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c match)" where
       
   148   "branch p \<equiv> \<Lambda> r x. match_case\<cdot>fail\<cdot>(\<Lambda> y. succeed\<cdot>(r\<cdot>y))\<cdot>(p\<cdot>x)"
       
   149 
       
   150 lemma branch_simps:
       
   151   "p\<cdot>x = \<bottom> \<Longrightarrow> branch p\<cdot>r\<cdot>x = \<bottom>"
       
   152   "p\<cdot>x = fail \<Longrightarrow> branch p\<cdot>r\<cdot>x = fail"
       
   153   "p\<cdot>x = succeed\<cdot>y \<Longrightarrow> branch p\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>y)"
       
   154 by (simp_all add: branch_def)
       
   155 
       
   156 lemma branch_succeed [simp]: "branch succeed\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>x)"
       
   157 by (simp add: branch_def)
       
   158 
       
   159 subsubsection {* Cases operator *}
       
   160 
       
   161 definition
       
   162   cases :: "'a match \<rightarrow> 'a::pcpo" where
       
   163   "cases = match_case\<cdot>\<bottom>\<cdot>ID"
       
   164 
       
   165 text {* rewrite rules for cases *}
       
   166 
       
   167 lemma cases_strict [simp]: "cases\<cdot>\<bottom> = \<bottom>"
       
   168 by (simp add: cases_def)
       
   169 
       
   170 lemma cases_fail [simp]: "cases\<cdot>fail = \<bottom>"
       
   171 by (simp add: cases_def)
       
   172 
       
   173 lemma cases_succeed [simp]: "cases\<cdot>(succeed\<cdot>x) = x"
       
   174 by (simp add: cases_def)
       
   175 
       
   176 subsection {* Case syntax *}
       
   177 
       
   178 nonterminals
       
   179   Case_syn  Cases_syn
       
   180 
       
   181 syntax
       
   182   "_Case_syntax":: "['a, Cases_syn] => 'b"               ("(Case _ of/ _)" 10)
       
   183   "_Case1"      :: "['a, 'b] => Case_syn"                ("(2_ =>/ _)" 10)
       
   184   ""            :: "Case_syn => Cases_syn"               ("_")
       
   185   "_Case2"      :: "[Case_syn, Cases_syn] => Cases_syn"  ("_/ | _")
       
   186 
       
   187 syntax (xsymbols)
       
   188   "_Case1"      :: "['a, 'b] => Case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
       
   189 
       
   190 translations
       
   191   "_Case_syntax x ms" == "CONST Fixrec.cases\<cdot>(ms\<cdot>x)"
       
   192   "_Case2 m ms" == "m \<parallel> ms"
       
   193 
       
   194 text {* Parsing Case expressions *}
       
   195 
       
   196 syntax
       
   197   "_pat" :: "'a"
       
   198   "_variable" :: "'a"
       
   199   "_noargs" :: "'a"
       
   200 
       
   201 translations
       
   202   "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_variable p r)"
       
   203   "_variable (_args x y) r" => "CONST csplit\<cdot>(_variable x (_variable y r))"
       
   204   "_variable _noargs r" => "CONST unit_when\<cdot>r"
       
   205 
       
   206 parse_translation {*
       
   207 (* rewrite (_pat x) => (succeed) *)
       
   208 (* rewrite (_variable x t) => (Abs_CFun (%x. t)) *)
       
   209  [(@{syntax_const "_pat"}, fn _ => Syntax.const @{const_syntax Fixrec.succeed}),
       
   210   mk_binder_tr (@{syntax_const "_variable"}, @{const_syntax Abs_CFun})];
       
   211 *}
       
   212 
       
   213 text {* Printing Case expressions *}
       
   214 
       
   215 syntax
       
   216   "_match" :: "'a"
       
   217 
       
   218 print_translation {*
       
   219   let
       
   220     fun dest_LAM (Const (@{const_syntax Rep_CFun},_) $ Const (@{const_syntax unit_when},_) $ t) =
       
   221           (Syntax.const @{syntax_const "_noargs"}, t)
       
   222     |   dest_LAM (Const (@{const_syntax Rep_CFun},_) $ Const (@{const_syntax csplit},_) $ t) =
       
   223           let
       
   224             val (v1, t1) = dest_LAM t;
       
   225             val (v2, t2) = dest_LAM t1;
       
   226           in (Syntax.const @{syntax_const "_args"} $ v1 $ v2, t2) end
       
   227     |   dest_LAM (Const (@{const_syntax Abs_CFun},_) $ t) =
       
   228           let
       
   229             val abs =
       
   230               case t of Abs abs => abs
       
   231                 | _ => ("x", dummyT, incr_boundvars 1 t $ Bound 0);
       
   232             val (x, t') = atomic_abs_tr' abs;
       
   233           in (Syntax.const @{syntax_const "_variable"} $ x, t') end
       
   234     |   dest_LAM _ = raise Match; (* too few vars: abort translation *)
       
   235 
       
   236     fun Case1_tr' [Const(@{const_syntax branch},_) $ p, r] =
       
   237           let val (v, t) = dest_LAM r in
       
   238             Syntax.const @{syntax_const "_Case1"} $
       
   239               (Syntax.const @{syntax_const "_match"} $ p $ v) $ t
       
   240           end;
       
   241 
       
   242   in [(@{const_syntax Rep_CFun}, Case1_tr')] end;
       
   243 *}
       
   244 
       
   245 translations
       
   246   "x" <= "_match (CONST Fixrec.succeed) (_variable x)"
       
   247 
       
   248 
       
   249 subsection {* Pattern combinators for data constructors *}
       
   250 
       
   251 types ('a, 'b) pat = "'a \<rightarrow> 'b match"
       
   252 
       
   253 definition
       
   254   cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
       
   255   "cpair_pat p1 p2 = (\<Lambda>(x, y).
       
   256     match_case\<cdot>fail\<cdot>(\<Lambda> a. match_case\<cdot>fail\<cdot>(\<Lambda> b. succeed\<cdot>(a, b))\<cdot>(p2\<cdot>y))\<cdot>(p1\<cdot>x))"
       
   257 
       
   258 definition
       
   259   spair_pat ::
       
   260   "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a::pcpo \<otimes> 'b::pcpo, 'c \<times> 'd) pat" where
       
   261   "spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>(x, y))"
       
   262 
       
   263 definition
       
   264   sinl_pat :: "('a, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
       
   265   "sinl_pat p = sscase\<cdot>p\<cdot>(\<Lambda> x. fail)"
       
   266 
       
   267 definition
       
   268   sinr_pat :: "('b, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
       
   269   "sinr_pat p = sscase\<cdot>(\<Lambda> x. fail)\<cdot>p"
       
   270 
       
   271 definition
       
   272   up_pat :: "('a, 'b) pat \<Rightarrow> ('a u, 'b) pat" where
       
   273   "up_pat p = fup\<cdot>p"
       
   274 
       
   275 definition
       
   276   TT_pat :: "(tr, unit) pat" where
       
   277   "TT_pat = (\<Lambda> b. If b then succeed\<cdot>() else fail fi)"
       
   278 
       
   279 definition
       
   280   FF_pat :: "(tr, unit) pat" where
       
   281   "FF_pat = (\<Lambda> b. If b then fail else succeed\<cdot>() fi)"
       
   282 
       
   283 definition
       
   284   ONE_pat :: "(one, unit) pat" where
       
   285   "ONE_pat = (\<Lambda> ONE. succeed\<cdot>())"
       
   286 
       
   287 text {* Parse translations (patterns) *}
       
   288 translations
       
   289   "_pat (XCONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"
       
   290   "_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
       
   291   "_pat (XCONST sinl\<cdot>x)" => "CONST sinl_pat (_pat x)"
       
   292   "_pat (XCONST sinr\<cdot>x)" => "CONST sinr_pat (_pat x)"
       
   293   "_pat (XCONST up\<cdot>x)" => "CONST up_pat (_pat x)"
       
   294   "_pat (XCONST TT)" => "CONST TT_pat"
       
   295   "_pat (XCONST FF)" => "CONST FF_pat"
       
   296   "_pat (XCONST ONE)" => "CONST ONE_pat"
       
   297 
       
   298 text {* CONST version is also needed for constructors with special syntax *}
       
   299 translations
       
   300   "_pat (CONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"
       
   301   "_pat (CONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"
       
   302 
       
   303 text {* Parse translations (variables) *}
       
   304 translations
       
   305   "_variable (XCONST Pair x y) r" => "_variable (_args x y) r"
       
   306   "_variable (XCONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
       
   307   "_variable (XCONST sinl\<cdot>x) r" => "_variable x r"
       
   308   "_variable (XCONST sinr\<cdot>x) r" => "_variable x r"
       
   309   "_variable (XCONST up\<cdot>x) r" => "_variable x r"
       
   310   "_variable (XCONST TT) r" => "_variable _noargs r"
       
   311   "_variable (XCONST FF) r" => "_variable _noargs r"
       
   312   "_variable (XCONST ONE) r" => "_variable _noargs r"
       
   313 
       
   314 translations
       
   315   "_variable (CONST Pair x y) r" => "_variable (_args x y) r"
       
   316   "_variable (CONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
       
   317 
       
   318 text {* Print translations *}
       
   319 translations
       
   320   "CONST Pair (_match p1 v1) (_match p2 v2)"
       
   321       <= "_match (CONST cpair_pat p1 p2) (_args v1 v2)"
       
   322   "CONST spair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)"
       
   323       <= "_match (CONST spair_pat p1 p2) (_args v1 v2)"
       
   324   "CONST sinl\<cdot>(_match p1 v1)" <= "_match (CONST sinl_pat p1) v1"
       
   325   "CONST sinr\<cdot>(_match p1 v1)" <= "_match (CONST sinr_pat p1) v1"
       
   326   "CONST up\<cdot>(_match p1 v1)" <= "_match (CONST up_pat p1) v1"
       
   327   "CONST TT" <= "_match (CONST TT_pat) _noargs"
       
   328   "CONST FF" <= "_match (CONST FF_pat) _noargs"
       
   329   "CONST ONE" <= "_match (CONST ONE_pat) _noargs"
       
   330 
       
   331 lemma cpair_pat1:
       
   332   "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = \<bottom>"
       
   333 apply (simp add: branch_def cpair_pat_def)
       
   334 apply (cases "p\<cdot>x", simp_all)
       
   335 done
       
   336 
       
   337 lemma cpair_pat2:
       
   338   "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = fail"
       
   339 apply (simp add: branch_def cpair_pat_def)
       
   340 apply (cases "p\<cdot>x", simp_all)
       
   341 done
       
   342 
       
   343 lemma cpair_pat3:
       
   344   "branch p\<cdot>r\<cdot>x = succeed\<cdot>s \<Longrightarrow>
       
   345    branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = branch q\<cdot>s\<cdot>y"
       
   346 apply (simp add: branch_def cpair_pat_def)
       
   347 apply (cases "p\<cdot>x", simp_all)
       
   348 apply (cases "q\<cdot>y", simp_all)
       
   349 done
       
   350 
       
   351 lemmas cpair_pat [simp] =
       
   352   cpair_pat1 cpair_pat2 cpair_pat3
       
   353 
       
   354 lemma spair_pat [simp]:
       
   355   "branch (spair_pat p1 p2)\<cdot>r\<cdot>\<bottom> = \<bottom>"
       
   356   "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk>
       
   357      \<Longrightarrow> branch (spair_pat p1 p2)\<cdot>r\<cdot>(:x, y:) =
       
   358          branch (cpair_pat p1 p2)\<cdot>r\<cdot>(x, y)"
       
   359 by (simp_all add: branch_def spair_pat_def)
       
   360 
       
   361 lemma sinl_pat [simp]:
       
   362   "branch (sinl_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"
       
   363   "x \<noteq> \<bottom> \<Longrightarrow> branch (sinl_pat p)\<cdot>r\<cdot>(sinl\<cdot>x) = branch p\<cdot>r\<cdot>x"
       
   364   "y \<noteq> \<bottom> \<Longrightarrow> branch (sinl_pat p)\<cdot>r\<cdot>(sinr\<cdot>y) = fail"
       
   365 by (simp_all add: branch_def sinl_pat_def)
       
   366 
       
   367 lemma sinr_pat [simp]:
       
   368   "branch (sinr_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"
       
   369   "x \<noteq> \<bottom> \<Longrightarrow> branch (sinr_pat p)\<cdot>r\<cdot>(sinl\<cdot>x) = fail"
       
   370   "y \<noteq> \<bottom> \<Longrightarrow> branch (sinr_pat p)\<cdot>r\<cdot>(sinr\<cdot>y) = branch p\<cdot>r\<cdot>y"
       
   371 by (simp_all add: branch_def sinr_pat_def)
       
   372 
       
   373 lemma up_pat [simp]:
       
   374   "branch (up_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"
       
   375   "branch (up_pat p)\<cdot>r\<cdot>(up\<cdot>x) = branch p\<cdot>r\<cdot>x"
       
   376 by (simp_all add: branch_def up_pat_def)
       
   377 
       
   378 lemma TT_pat [simp]:
       
   379   "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
       
   380   "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = succeed\<cdot>r"
       
   381   "branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = fail"
       
   382 by (simp_all add: branch_def TT_pat_def)
       
   383 
       
   384 lemma FF_pat [simp]:
       
   385   "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
       
   386   "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = fail"
       
   387   "branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = succeed\<cdot>r"
       
   388 by (simp_all add: branch_def FF_pat_def)
       
   389 
       
   390 lemma ONE_pat [simp]:
       
   391   "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"
       
   392   "branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>ONE = succeed\<cdot>r"
       
   393 by (simp_all add: branch_def ONE_pat_def)
       
   394 
       
   395 
       
   396 subsection {* Wildcards, as-patterns, and lazy patterns *}
       
   397 
       
   398 definition
       
   399   wild_pat :: "'a \<rightarrow> unit match" where
       
   400   "wild_pat = (\<Lambda> x. succeed\<cdot>())"
       
   401 
       
   402 definition
       
   403   as_pat :: "('a \<rightarrow> 'b match) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) match" where
       
   404   "as_pat p = (\<Lambda> x. match_case\<cdot>fail\<cdot>(\<Lambda> a. succeed\<cdot>(x, a))\<cdot>(p\<cdot>x))"
       
   405 
       
   406 definition
       
   407   lazy_pat :: "('a \<rightarrow> 'b::pcpo match) \<Rightarrow> ('a \<rightarrow> 'b match)" where
       
   408   "lazy_pat p = (\<Lambda> x. succeed\<cdot>(cases\<cdot>(p\<cdot>x)))"
       
   409 
       
   410 text {* Parse translations (patterns) *}
       
   411 translations
       
   412   "_pat _" => "CONST wild_pat"
       
   413 
       
   414 text {* Parse translations (variables) *}
       
   415 translations
       
   416   "_variable _ r" => "_variable _noargs r"
       
   417 
       
   418 text {* Print translations *}
       
   419 translations
       
   420   "_" <= "_match (CONST wild_pat) _noargs"
       
   421 
       
   422 lemma wild_pat [simp]: "branch wild_pat\<cdot>(unit_when\<cdot>r)\<cdot>x = succeed\<cdot>r"
       
   423 by (simp add: branch_def wild_pat_def)
       
   424 
       
   425 lemma as_pat [simp]:
       
   426   "branch (as_pat p)\<cdot>(csplit\<cdot>r)\<cdot>x = branch p\<cdot>(r\<cdot>x)\<cdot>x"
       
   427 apply (simp add: branch_def as_pat_def)
       
   428 apply (cases "p\<cdot>x", simp_all)
       
   429 done
       
   430 
       
   431 lemma lazy_pat [simp]:
       
   432   "branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>\<bottom>)"
       
   433   "branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>\<bottom>)"
       
   434   "branch p\<cdot>r\<cdot>x = succeed\<cdot>s \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>s"
       
   435 apply (simp_all add: branch_def lazy_pat_def)
       
   436 apply (cases "p\<cdot>x", simp_all)+
       
   437 done
       
   438 
       
   439 
   110 
   440 subsection {* Match functions for built-in types *}
   111 subsection {* Match functions for built-in types *}
   441 
   112 
   442 default_sort pcpo
   113 default_sort pcpo
   443 
   114 
   582       (@{const_name TT}, @{const_name match_TT}),
   253       (@{const_name TT}, @{const_name match_TT}),
   583       (@{const_name FF}, @{const_name match_FF}),
   254       (@{const_name FF}, @{const_name match_FF}),
   584       (@{const_name UU}, @{const_name match_UU}) ]
   255       (@{const_name UU}, @{const_name match_UU}) ]
   585 *}
   256 *}
   586 
   257 
   587 hide_const (open) succeed fail run cases
   258 hide_const (open) succeed fail run
   588 
   259 
   589 end
   260 end