| author | Lukas Stevens <mail@lukas-stevens.de> | 
| Fri, 30 Sep 2022 12:41:32 +0200 | |
| changeset 76226 | 2aad8698f82f | 
| parent 72581 | de581f98a3a1 | 
| child 80783 | 14ed085de388 | 
| permissions | -rw-r--r-- | 
| 66270 
403d84138c5c
State_Monad ~> Open_State_Syntax
 Lars Hupel <lars.hupel@mytum.de> parents: 
66258diff
changeset | 1 | (* Title: HOL/Library/Open_State_Syntax.thy | 
| 21192 | 2 | Author: Florian Haftmann, TU Muenchen | 
| 3 | *) | |
| 4 | ||
| 60500 | 5 | section \<open>Combinator syntax for generic, open state monads (single-threaded monads)\<close> | 
| 21192 | 6 | |
| 66270 
403d84138c5c
State_Monad ~> Open_State_Syntax
 Lars Hupel <lars.hupel@mytum.de> parents: 
66258diff
changeset | 7 | theory Open_State_Syntax | 
| 66258 | 8 | imports Main | 
| 21192 | 9 | begin | 
| 10 | ||
| 72581 | 11 | context | 
| 12 | includes state_combinator_syntax | |
| 13 | begin | |
| 14 | ||
| 60500 | 15 | subsection \<open>Motivation\<close> | 
| 21192 | 16 | |
| 60500 | 17 | text \<open> | 
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 18 | The logic HOL has no notion of constructor classes, so it is not | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 19 | possible to model monads the Haskell way in full genericity in | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 20 | Isabelle/HOL. | 
| 21192 | 21 | |
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 22 | However, this theory provides substantial support for a very common | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 23 |   class of monads: \emph{state monads} (or \emph{single-threaded
 | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 24 | monads}, since a state is transformed single-threadedly). | 
| 21192 | 25 | |
| 26 | To enter from the Haskell world, | |
| 68224 | 27 | \<^url>\<open>https://www.engr.mun.ca/~theo/Misc/haskell_and_monads.htm\<close> makes | 
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 28 | a good motivating start. Here we just sketch briefly how those | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 29 | monads enter the game of Isabelle/HOL. | 
| 60500 | 30 | \<close> | 
| 21192 | 31 | |
| 60500 | 32 | subsection \<open>State transformations and combinators\<close> | 
| 21192 | 33 | |
| 60500 | 34 | text \<open> | 
| 21192 | 35 | We classify functions operating on states into two categories: | 
| 36 | ||
| 37 |   \begin{description}
 | |
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 38 | |
| 61585 | 39 | \item[transformations] with type signature \<open>\<sigma> \<Rightarrow> \<sigma>'\<close>, | 
| 21192 | 40 | transforming a state. | 
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 41 | |
| 61585 | 42 | \item[``yielding'' transformations] with type signature \<open>\<sigma> | 
| 43 | \<Rightarrow> \<alpha> \<times> \<sigma>'\<close>, ``yielding'' a side result while transforming a | |
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 44 | state. | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 45 | |
| 61585 | 46 | \item[queries] with type signature \<open>\<sigma> \<Rightarrow> \<alpha>\<close>, computing a | 
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 47 | result dependent on a state. | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 48 | |
| 21192 | 49 |   \end{description}
 | 
| 50 | ||
| 61585 | 51 | By convention we write \<open>\<sigma>\<close> for types representing states and | 
| 52 | \<open>\<alpha>\<close>, \<open>\<beta>\<close>, \<open>\<gamma>\<close>, \<open>\<dots>\<close> for types | |
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 53 | representing side results. Type changes due to transformations are | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 54 | not excluded in our scenario. | 
| 21192 | 55 | |
| 61585 | 56 | We aim to assert that values of any state type \<open>\<sigma>\<close> are used | 
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 57 | in a single-threaded way: after application of a transformation on a | 
| 61585 | 58 | value of type \<open>\<sigma>\<close>, the former value should not be used | 
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 59 | again. To achieve this, we use a set of monad combinators: | 
| 60500 | 60 | \<close> | 
| 21192 | 61 | |
| 60500 | 62 | text \<open> | 
| 69593 | 63 | Given two transformations \<^term>\<open>f\<close> and \<^term>\<open>g\<close>, they may be | 
| 64 | directly composed using the \<^term>\<open>(\<circ>>)\<close> combinator, forming a | |
| 65 | forward composition: \<^prop>\<open>(f \<circ>> g) s = f (g s)\<close>. | |
| 21192 | 66 | |
| 67 | After any yielding transformation, we bind the side result | |
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 68 | immediately using a lambda abstraction. This is the purpose of the | 
| 69593 | 69 | \<^term>\<open>(\<circ>\<rightarrow>)\<close> combinator: \<^prop>\<open>(f \<circ>\<rightarrow> (\<lambda>x. g)) s = (let (x, s') | 
| 70 | = f s in g s')\<close>. | |
| 21192 | 71 | |
| 69593 | 72 | For queries, the existing \<^term>\<open>Let\<close> is appropriate. | 
| 21192 | 73 | |
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 74 | Naturally, a computation may yield a side result by pairing it to | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 75 | the state from the left; we introduce the suggestive abbreviation | 
| 69593 | 76 | \<^term>\<open>return\<close> for this purpose. | 
| 21192 | 77 | |
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 78 | The most crucial distinction to Haskell is that we do not need to | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 79 | introduce distinguished type constructors for different kinds of | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 80 | state. This has two consequences: | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 81 | |
| 21192 | 82 |   \begin{itemize}
 | 
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 83 | |
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 84 | \item The monad model does not state anything about the kind of | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 85 | state; the model for the state is completely orthogonal and may | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 86 | be specified completely independently. | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 87 | |
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 88 | \item There is no distinguished type constructor encapsulating | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 89 | away the state transformation, i.e.~transformations may be | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 90 | applied directly without using any lifting or providing and | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 91 | dropping units (``open monad''). | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 92 | |
| 21192 | 93 | \item The type of states may change due to a transformation. | 
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 94 | |
| 21192 | 95 |   \end{itemize}
 | 
| 60500 | 96 | \<close> | 
| 21192 | 97 | |
| 98 | ||
| 60500 | 99 | subsection \<open>Monad laws\<close> | 
| 21192 | 100 | |
| 60500 | 101 | text \<open> | 
| 40359 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 102 | The common monadic laws hold and may also be used as normalization | 
| 
84388bba911d
dropped return abbreviation for Pair -- confusing if combined with monad syntax; tuned document
 haftmann parents: 
38345diff
changeset | 103 | rules for monadic expressions: | 
| 60500 | 104 | \<close> | 
| 21192 | 105 | |
| 28145 | 106 | lemmas monad_simp = Pair_scomp scomp_Pair id_fcomp fcomp_id | 
| 107 | scomp_scomp scomp_fcomp fcomp_scomp fcomp_assoc | |
| 21192 | 108 | |
| 60500 | 109 | text \<open> | 
| 21192 | 110 | Evaluation of monadic expressions by force: | 
| 60500 | 111 | \<close> | 
| 21192 | 112 | |
| 28145 | 113 | lemmas monad_collapse = monad_simp fcomp_apply scomp_apply split_beta | 
| 26260 | 114 | |
| 72581 | 115 | end | 
| 116 | ||
| 37817 | 117 | |
| 60500 | 118 | subsection \<open>Do-syntax\<close> | 
| 37817 | 119 | |
| 41229 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
40361diff
changeset | 120 | nonterminal sdo_binds and sdo_bind | 
| 37932 | 121 | |
| 122 | syntax | |
| 123 |   "_sdo_block" :: "sdo_binds \<Rightarrow> 'a" ("exec {//(2  _)//}" [12] 62)
 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61585diff
changeset | 124 |   "_sdo_bind"  :: "[pttrn, 'a] \<Rightarrow> sdo_bind" ("(_ \<leftarrow>/ _)" 13)
 | 
| 37932 | 125 |   "_sdo_let" :: "[pttrn, 'a] \<Rightarrow> sdo_bind" ("(2let _ =/ _)" [1000, 13] 13)
 | 
| 126 |   "_sdo_then" :: "'a \<Rightarrow> sdo_bind" ("_" [14] 13)
 | |
| 127 |   "_sdo_final" :: "'a \<Rightarrow> sdo_binds" ("_")
 | |
| 128 |   "_sdo_cons" :: "[sdo_bind, sdo_binds] \<Rightarrow> sdo_binds" ("_;//_" [13, 12] 12)
 | |
| 21192 | 129 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61585diff
changeset | 130 | syntax (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61585diff
changeset | 131 |   "_sdo_bind" :: "[pttrn, 'a] \<Rightarrow> sdo_bind" ("(_ <-/ _)" 13)
 | 
| 37932 | 132 | |
| 133 | translations | |
| 134 | "_sdo_block (_sdo_cons (_sdo_bind p t) (_sdo_final e))" | |
| 135 | == "CONST scomp t (\<lambda>p. e)" | |
| 136 | "_sdo_block (_sdo_cons (_sdo_then t) (_sdo_final e))" | |
| 38345 
8b8fc27c1872
print fcomp combinator only monadic in connection with other monadic expressions
 haftmann parents: 
37932diff
changeset | 137 | => "CONST fcomp t e" | 
| 
8b8fc27c1872
print fcomp combinator only monadic in connection with other monadic expressions
 haftmann parents: 
37932diff
changeset | 138 | "_sdo_final (_sdo_block (_sdo_cons (_sdo_then t) (_sdo_final e)))" | 
| 
8b8fc27c1872
print fcomp combinator only monadic in connection with other monadic expressions
 haftmann parents: 
37932diff
changeset | 139 | <= "_sdo_final (CONST fcomp t e)" | 
| 
8b8fc27c1872
print fcomp combinator only monadic in connection with other monadic expressions
 haftmann parents: 
37932diff
changeset | 140 | "_sdo_block (_sdo_cons (_sdo_then t) e)" | 
| 
8b8fc27c1872
print fcomp combinator only monadic in connection with other monadic expressions
 haftmann parents: 
37932diff
changeset | 141 | <= "CONST fcomp t (_sdo_block e)" | 
| 37932 | 142 | "_sdo_block (_sdo_cons (_sdo_let p t) bs)" | 
| 143 | == "let p = t in _sdo_block bs" | |
| 144 | "_sdo_block (_sdo_cons b (_sdo_cons c cs))" | |
| 145 | == "_sdo_block (_sdo_cons b (_sdo_final (_sdo_block (_sdo_cons c cs))))" | |
| 146 | "_sdo_cons (_sdo_let p t) (_sdo_final s)" | |
| 147 | == "_sdo_final (let p = t in s)" | |
| 148 | "_sdo_block (_sdo_final e)" => "e" | |
| 37817 | 149 | |
| 60500 | 150 | text \<open> | 
| 63680 | 151 | For an example, see \<^file>\<open>~~/src/HOL/Proofs/Extraction/Higman_Extraction.thy\<close>. | 
| 60500 | 152 | \<close> | 
| 21192 | 153 | |
| 22664 | 154 | end |