author | berghofe |
Thu, 12 Jan 2012 00:13:37 +0100 | |
changeset 46191 | a88546428c2a |
parent 41959 | b460124855b8 |
child 52143 | 36ffe23b25f8 |
permissions | -rw-r--r-- |
41959 | 1 |
(* Title: HOL/Statespace/StateSpaceSyntax.thy |
25171 | 2 |
Author: Norbert Schirmer, TU Muenchen |
3 |
*) |
|
4 |
||
5 |
header {* Syntax for State Space Lookup and Update \label{sec:StateSpaceSyntax}*} |
|
6 |
theory StateSpaceSyntax |
|
7 |
imports StateSpaceLocale |
|
8 |
begin |
|
9 |
||
10 |
text {* The state space syntax is kept in an extra theory so that you |
|
11 |
can choose if you want to use it or not. *} |
|
12 |
||
13 |
syntax |
|
35114
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
14 |
"_statespace_lookup" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" ("_\<cdot>_" [60, 60] 60) |
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
15 |
"_statespace_update" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c \<Rightarrow> ('a \<Rightarrow> 'b)" |
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
16 |
"_statespace_updates" :: "('a \<Rightarrow> 'b) \<Rightarrow> updbinds \<Rightarrow> ('a \<Rightarrow> 'b)" ("_<_>" [900, 0] 900) |
25171 | 17 |
|
18 |
translations |
|
35114
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
19 |
"_statespace_updates f (_updbinds b bs)" == |
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
20 |
"_statespace_updates (_statespace_updates f b) bs" |
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
21 |
"s<x:=y>" == "_statespace_update s x y" |
25171 | 22 |
|
23 |
||
24 |
parse_translation (advanced) |
|
25 |
{* |
|
35114
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
26 |
[(@{syntax_const "_statespace_lookup"}, StateSpace.lookup_tr), |
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
27 |
(@{syntax_const "_statespace_update"}, StateSpace.update_tr)] |
25171 | 28 |
*} |
29 |
||
30 |
||
31 |
print_translation (advanced) |
|
32 |
{* |
|
35114
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
33 |
[(@{const_syntax lookup}, StateSpace.lookup_tr'), |
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
34 |
(@{const_syntax update}, StateSpace.update_tr')] |
25171 | 35 |
*} |
36 |
||
35114
b1fd1d756e20
formal markup of @{syntax_const} and @{const_syntax};
wenzelm
parents:
25171
diff
changeset
|
37 |
end |