| author | wenzelm | 
| Sat, 11 Feb 2023 22:13:55 +0100 | |
| changeset 77252 | 36c856e25b73 | 
| parent 77107 | 4c4d40913900 | 
| child 80768 | c7723cc15de8 | 
| permissions | -rw-r--r-- | 
| 
68061
 
81d90f830f99
clarified theory location and imports: avoid surprises due to Pure instead of Main (e.g. simpset operations);
 
wenzelm 
parents: 
66608 
diff
changeset
 | 
1  | 
(* Title: HOL/Library/Monad_Syntax.thy  | 
| 
 
81d90f830f99
clarified theory location and imports: avoid surprises due to Pure instead of Main (e.g. simpset operations);
 
wenzelm 
parents: 
66608 
diff
changeset
 | 
2  | 
Author: Alexander Krauss, TU Muenchen  | 
| 
 
81d90f830f99
clarified theory location and imports: avoid surprises due to Pure instead of Main (e.g. simpset operations);
 
wenzelm 
parents: 
66608 
diff
changeset
 | 
3  | 
Author: Christian Sternagel, University of Innsbruck  | 
| 37790 | 4  | 
*)  | 
5  | 
||
| 60500 | 6  | 
section \<open>Monad notation for arbitrary types\<close>  | 
| 37790 | 7  | 
|
8  | 
theory Monad_Syntax  | 
|
| 
68061
 
81d90f830f99
clarified theory location and imports: avoid surprises due to Pure instead of Main (e.g. simpset operations);
 
wenzelm 
parents: 
66608 
diff
changeset
 | 
9  | 
imports Adhoc_Overloading  | 
| 37790 | 10  | 
begin  | 
11  | 
||
| 60500 | 12  | 
text \<open>  | 
| 70433 | 13  | 
We provide a convenient do-notation for monadic expressions well-known from Haskell.  | 
| 77107 | 14  | 
\<^const>\<open>Let\<close> is printed specially in do-expressions.  | 
| 60500 | 15  | 
\<close>  | 
| 37791 | 16  | 
|
| 37790 | 17  | 
consts  | 
| 70433 | 18  | 
  bind :: "'a \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'd" (infixl "\<bind>" 54)
 | 
| 37790 | 19  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
20  | 
notation (ASCII)  | 
| 70433 | 21  | 
bind (infixl ">>=" 54)  | 
| 37790 | 22  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
23  | 
|
| 37790 | 24  | 
abbreviation (do_notation)  | 
| 70433 | 25  | 
  bind_do :: "'a \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'd"
 | 
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
26  | 
where "bind_do \<equiv> bind"  | 
| 37790 | 27  | 
|
28  | 
notation (output)  | 
|
| 70433 | 29  | 
bind_do (infixl "\<bind>" 54)  | 
| 37790 | 30  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
31  | 
notation (ASCII output)  | 
| 70433 | 32  | 
bind_do (infixl ">>=" 54)  | 
| 37790 | 33  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
34  | 
|
| 
41229
 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 
wenzelm 
parents: 
39151 
diff
changeset
 | 
35  | 
nonterminal do_binds and do_bind  | 
| 37790 | 36  | 
syntax  | 
37  | 
  "_do_block" :: "do_binds \<Rightarrow> 'a" ("do {//(2  _)//}" [12] 62)
 | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
38  | 
  "_do_bind"  :: "[pttrn, 'a] \<Rightarrow> do_bind" ("(2_ \<leftarrow>/ _)" 13)
 | 
| 37790 | 39  | 
  "_do_let" :: "[pttrn, 'a] \<Rightarrow> do_bind" ("(2let _ =/ _)" [1000, 13] 13)
 | 
40  | 
  "_do_then" :: "'a \<Rightarrow> do_bind" ("_" [14] 13)
 | 
|
41  | 
  "_do_final" :: "'a \<Rightarrow> do_binds" ("_")
 | 
|
42  | 
  "_do_cons" :: "[do_bind, do_binds] \<Rightarrow> do_binds" ("_;//_" [13, 12] 12)
 | 
|
| 70433 | 43  | 
"_thenM" :: "['a, 'b] \<Rightarrow> 'c" (infixl "\<then>" 54)  | 
| 37790 | 44  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
45  | 
syntax (ASCII)  | 
| 
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
46  | 
  "_do_bind" :: "[pttrn, 'a] \<Rightarrow> do_bind" ("(2_ <-/ _)" 13)
 | 
| 70433 | 47  | 
"_thenM" :: "['a, 'b] \<Rightarrow> 'c" (infixl ">>" 54)  | 
| 37790 | 48  | 
|
49  | 
translations  | 
|
50  | 
"_do_block (_do_cons (_do_then t) (_do_final e))"  | 
|
| 62035 | 51  | 
\<rightleftharpoons> "CONST bind_do t (\<lambda>_. e)"  | 
| 37790 | 52  | 
"_do_block (_do_cons (_do_bind p t) (_do_final e))"  | 
| 62035 | 53  | 
\<rightleftharpoons> "CONST bind_do t (\<lambda>p. e)"  | 
| 37790 | 54  | 
"_do_block (_do_cons (_do_let p t) bs)"  | 
| 62035 | 55  | 
\<rightleftharpoons> "let p = t in _do_block bs"  | 
| 37790 | 56  | 
"_do_block (_do_cons b (_do_cons c cs))"  | 
| 62035 | 57  | 
\<rightleftharpoons> "_do_block (_do_cons b (_do_final (_do_block (_do_cons c cs))))"  | 
| 37790 | 58  | 
"_do_cons (_do_let p t) (_do_final s)"  | 
| 62035 | 59  | 
\<rightleftharpoons> "_do_final (let p = t in s)"  | 
| 62189 | 60  | 
"_do_block (_do_final e)" \<rightharpoonup> "e"  | 
| 62035 | 61  | 
"(m \<then> n)" \<rightharpoonup> "(m \<bind> (\<lambda>_. n))"  | 
| 37790 | 62  | 
|
| 
52622
 
e0ff1625e96d
localized and modernized adhoc-overloading (patch by Christian Sternagel);
 
wenzelm 
parents: 
46143 
diff
changeset
 | 
63  | 
adhoc_overloading  | 
| 66608 | 64  | 
bind Set.bind Predicate.bind Option.bind List.bind  | 
| 37790 | 65  | 
|
| 66608 | 66  | 
end  |