src/HOL/HOLCF/Fixrec.thy
changeset 62175 8ffc4d0e652d
parent 58880 0baae4311a9f
child 65380 ae93953746fc
equal deleted inserted replaced
62174:fae6233c5f37 62175:8ffc4d0e652d
     7 theory Fixrec
     7 theory Fixrec
     8 imports Plain_HOLCF
     8 imports Plain_HOLCF
     9 keywords "fixrec" :: thy_decl
     9 keywords "fixrec" :: thy_decl
    10 begin
    10 begin
    11 
    11 
    12 subsection {* Pattern-match monad *}
    12 subsection \<open>Pattern-match monad\<close>
    13 
    13 
    14 default_sort cpo
    14 default_sort cpo
    15 
    15 
    16 pcpodef 'a match = "UNIV::(one ++ 'a u) set"
    16 pcpodef 'a match = "UNIV::(one ++ 'a u) set"
    17 by simp_all
    17 by simp_all
    44 
    44 
    45 lemma succeed_neq_fail [simp]:
    45 lemma succeed_neq_fail [simp]:
    46   "succeed\<cdot>x \<noteq> fail" "fail \<noteq> succeed\<cdot>x"
    46   "succeed\<cdot>x \<noteq> fail" "fail \<noteq> succeed\<cdot>x"
    47 by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject)
    47 by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject)
    48 
    48 
    49 subsubsection {* Run operator *}
    49 subsubsection \<open>Run operator\<close>
    50 
    50 
    51 definition
    51 definition
    52   run :: "'a match \<rightarrow> 'a::pcpo" where
    52   run :: "'a match \<rightarrow> 'a::pcpo" where
    53   "run = (\<Lambda> m. sscase\<cdot>\<bottom>\<cdot>(fup\<cdot>ID)\<cdot>(Rep_match m))"
    53   "run = (\<Lambda> m. sscase\<cdot>\<bottom>\<cdot>(fup\<cdot>ID)\<cdot>(Rep_match m))"
    54 
    54 
    55 text {* rewrite rules for run *}
    55 text \<open>rewrite rules for run\<close>
    56 
    56 
    57 lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>"
    57 lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>"
    58 unfolding run_def
    58 unfolding run_def
    59 by (simp add: cont_Rep_match Rep_match_strict)
    59 by (simp add: cont_Rep_match Rep_match_strict)
    60 
    60 
    64 
    64 
    65 lemma run_succeed [simp]: "run\<cdot>(succeed\<cdot>x) = x"
    65 lemma run_succeed [simp]: "run\<cdot>(succeed\<cdot>x) = x"
    66 unfolding run_def succeed_def
    66 unfolding run_def succeed_def
    67 by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse)
    67 by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse)
    68 
    68 
    69 subsubsection {* Monad plus operator *}
    69 subsubsection \<open>Monad plus operator\<close>
    70 
    70 
    71 definition
    71 definition
    72   mplus :: "'a match \<rightarrow> 'a match \<rightarrow> 'a match" where
    72   mplus :: "'a match \<rightarrow> 'a match \<rightarrow> 'a match" where
    73   "mplus = (\<Lambda> m1 m2. sscase\<cdot>(\<Lambda> _. m2)\<cdot>(\<Lambda> _. m1)\<cdot>(Rep_match m1))"
    73   "mplus = (\<Lambda> m1 m2. sscase\<cdot>(\<Lambda> _. m2)\<cdot>(\<Lambda> _. m1)\<cdot>(Rep_match m1))"
    74 
    74 
    75 abbreviation
    75 abbreviation
    76   mplus_syn :: "['a match, 'a match] \<Rightarrow> 'a match"  (infixr "+++" 65)  where
    76   mplus_syn :: "['a match, 'a match] \<Rightarrow> 'a match"  (infixr "+++" 65)  where
    77   "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2"
    77   "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2"
    78 
    78 
    79 text {* rewrite rules for mplus *}
    79 text \<open>rewrite rules for mplus\<close>
    80 
    80 
    81 lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>"
    81 lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>"
    82 unfolding mplus_def
    82 unfolding mplus_def
    83 by (simp add: cont_Rep_match Rep_match_strict)
    83 by (simp add: cont_Rep_match Rep_match_strict)
    84 
    84 
    94 by (cases m, simp_all)
    94 by (cases m, simp_all)
    95 
    95 
    96 lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
    96 lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
    97 by (cases x, simp_all)
    97 by (cases x, simp_all)
    98 
    98 
    99 subsection {* Match functions for built-in types *}
    99 subsection \<open>Match functions for built-in types\<close>
   100 
   100 
   101 default_sort pcpo
   101 default_sort pcpo
   102 
   102 
   103 definition
   103 definition
   104   match_bottom :: "'a \<rightarrow> 'c match \<rightarrow> 'c match"
   104   match_bottom :: "'a \<rightarrow> 'c match \<rightarrow> 'c match"
   190   "match_FF\<cdot>FF\<cdot>k = k"
   190   "match_FF\<cdot>FF\<cdot>k = k"
   191   "match_FF\<cdot>TT\<cdot>k = fail"
   191   "match_FF\<cdot>TT\<cdot>k = fail"
   192   "match_FF\<cdot>\<bottom>\<cdot>k = \<bottom>"
   192   "match_FF\<cdot>\<bottom>\<cdot>k = \<bottom>"
   193 by (simp_all add: match_FF_def)
   193 by (simp_all add: match_FF_def)
   194 
   194 
   195 subsection {* Mutual recursion *}
   195 subsection \<open>Mutual recursion\<close>
   196 
   196 
   197 text {*
   197 text \<open>
   198   The following rules are used to prove unfolding theorems from
   198   The following rules are used to prove unfolding theorems from
   199   fixed-point definitions of mutually recursive functions.
   199   fixed-point definitions of mutually recursive functions.
   200 *}
   200 \<close>
   201 
   201 
   202 lemma Pair_equalI: "\<lbrakk>x \<equiv> fst p; y \<equiv> snd p\<rbrakk> \<Longrightarrow> (x, y) \<equiv> p"
   202 lemma Pair_equalI: "\<lbrakk>x \<equiv> fst p; y \<equiv> snd p\<rbrakk> \<Longrightarrow> (x, y) \<equiv> p"
   203 by simp
   203 by simp
   204 
   204 
   205 lemma Pair_eqD1: "(x, y) = (x', y') \<Longrightarrow> x = x'"
   205 lemma Pair_eqD1: "(x, y) = (x', y') \<Longrightarrow> x = x'"
   214 
   214 
   215 lemma def_cont_fix_ind:
   215 lemma def_cont_fix_ind:
   216   "\<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"
   216   "\<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"
   217 by (simp add: fix_ind)
   217 by (simp add: fix_ind)
   218 
   218 
   219 text {* lemma for proving rewrite rules *}
   219 text \<open>lemma for proving rewrite rules\<close>
   220 
   220 
   221 lemma ssubst_lhs: "\<lbrakk>t = s; P s = Q\<rbrakk> \<Longrightarrow> P t = Q"
   221 lemma ssubst_lhs: "\<lbrakk>t = s; P s = Q\<rbrakk> \<Longrightarrow> P t = Q"
   222 by simp
   222 by simp
   223 
   223 
   224 
   224 
   225 subsection {* Initializing the fixrec package *}
   225 subsection \<open>Initializing the fixrec package\<close>
   226 
   226 
   227 ML_file "Tools/holcf_library.ML"
   227 ML_file "Tools/holcf_library.ML"
   228 ML_file "Tools/fixrec.ML"
   228 ML_file "Tools/fixrec.ML"
   229 
   229 
   230 method_setup fixrec_simp = {*
   230 method_setup fixrec_simp = \<open>
   231   Scan.succeed (SIMPLE_METHOD' o Fixrec.fixrec_simp_tac)
   231   Scan.succeed (SIMPLE_METHOD' o Fixrec.fixrec_simp_tac)
   232 *} "pattern prover for fixrec constants"
   232 \<close> "pattern prover for fixrec constants"
   233 
   233 
   234 setup {*
   234 setup \<open>
   235   Fixrec.add_matchers
   235   Fixrec.add_matchers
   236     [ (@{const_name up}, @{const_name match_up}),
   236     [ (@{const_name up}, @{const_name match_up}),
   237       (@{const_name sinl}, @{const_name match_sinl}),
   237       (@{const_name sinl}, @{const_name match_sinl}),
   238       (@{const_name sinr}, @{const_name match_sinr}),
   238       (@{const_name sinr}, @{const_name match_sinr}),
   239       (@{const_name spair}, @{const_name match_spair}),
   239       (@{const_name spair}, @{const_name match_spair}),
   240       (@{const_name Pair}, @{const_name match_Pair}),
   240       (@{const_name Pair}, @{const_name match_Pair}),
   241       (@{const_name ONE}, @{const_name match_ONE}),
   241       (@{const_name ONE}, @{const_name match_ONE}),
   242       (@{const_name TT}, @{const_name match_TT}),
   242       (@{const_name TT}, @{const_name match_TT}),
   243       (@{const_name FF}, @{const_name match_FF}),
   243       (@{const_name FF}, @{const_name match_FF}),
   244       (@{const_name bottom}, @{const_name match_bottom}) ]
   244       (@{const_name bottom}, @{const_name match_bottom}) ]
   245 *}
   245 \<close>
   246 
   246 
   247 hide_const (open) succeed fail run
   247 hide_const (open) succeed fail run
   248 
   248 
   249 end
   249 end