author | wenzelm |
Fri, 20 Sep 2024 19:51:08 +0200 | |
changeset 80914 | d97fdabd9e2b |
parent 80784 | 3d9e7746d9db |
child 81187 | c66e24eae281 |
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 |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
18 |
bind :: "'a \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'd" (infixl \<open>\<bind>\<close> 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) |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
21 |
bind (infixl \<open>>>=\<close> 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) |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
29 |
bind_do (infixl \<open>\<bind>\<close> 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) |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
32 |
bind_do (infixl \<open>>>=\<close> 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 |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
37 |
"_do_block" :: "do_binds \<Rightarrow> 'a" (\<open>do {//(2 _)//}\<close> [12] 62) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
38 |
"_do_bind" :: "[pttrn, 'a] \<Rightarrow> do_bind" (\<open>(2_ \<leftarrow>/ _)\<close> 13) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
39 |
"_do_let" :: "[pttrn, 'a] \<Rightarrow> do_bind" (\<open>(2let _ =/ _)\<close> [1000, 13] 13) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
40 |
"_do_then" :: "'a \<Rightarrow> do_bind" (\<open>_\<close> [14] 13) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
41 |
"_do_final" :: "'a \<Rightarrow> do_binds" (\<open>_\<close>) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
42 |
"_do_cons" :: "[do_bind, do_binds] \<Rightarrow> do_binds" (\<open>_;//_\<close> [13, 12] 12) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
43 |
"_thenM" :: "['a, 'b] \<Rightarrow> 'c" (infixl \<open>\<then>\<close> 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) |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
46 |
"_do_bind" :: "[pttrn, 'a] \<Rightarrow> do_bind" (\<open>(2_ <-/ _)\<close> 13) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80784
diff
changeset
|
47 |
"_thenM" :: "['a, 'b] \<Rightarrow> 'c" (infixl \<open>>>\<close> 54) |
37790 | 48 |
|
80768 | 49 |
syntax_consts |
80784 | 50 |
"_do_block" "_do_cons" "_do_bind" "_do_then" \<rightleftharpoons> bind and |
80768 | 51 |
"_do_let" \<rightleftharpoons> Let |
52 |
||
37790 | 53 |
translations |
54 |
"_do_block (_do_cons (_do_then t) (_do_final e))" |
|
62035 | 55 |
\<rightleftharpoons> "CONST bind_do t (\<lambda>_. e)" |
37790 | 56 |
"_do_block (_do_cons (_do_bind p t) (_do_final e))" |
62035 | 57 |
\<rightleftharpoons> "CONST bind_do t (\<lambda>p. e)" |
37790 | 58 |
"_do_block (_do_cons (_do_let p t) bs)" |
62035 | 59 |
\<rightleftharpoons> "let p = t in _do_block bs" |
37790 | 60 |
"_do_block (_do_cons b (_do_cons c cs))" |
62035 | 61 |
\<rightleftharpoons> "_do_block (_do_cons b (_do_final (_do_block (_do_cons c cs))))" |
37790 | 62 |
"_do_cons (_do_let p t) (_do_final s)" |
62035 | 63 |
\<rightleftharpoons> "_do_final (let p = t in s)" |
62189 | 64 |
"_do_block (_do_final e)" \<rightharpoonup> "e" |
62035 | 65 |
"(m \<then> n)" \<rightharpoonup> "(m \<bind> (\<lambda>_. n))" |
37790 | 66 |
|
52622
e0ff1625e96d
localized and modernized adhoc-overloading (patch by Christian Sternagel);
wenzelm
parents:
46143
diff
changeset
|
67 |
adhoc_overloading |
66608 | 68 |
bind Set.bind Predicate.bind Option.bind List.bind |
37790 | 69 |
|
66608 | 70 |
end |