src/Doc/Isar_Ref/Inner_Syntax.thy
author wenzelm
Fri, 01 Jan 2016 14:44:52 +0100
changeset 62026 ea3b1b0413b4
parent 61997 4d9518c3d031
child 62106 d6af554512d7
permissions -rw-r--r--
more symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61656
cfabbc083977 more uniform jEdit properties;
wenzelm
parents: 61572
diff changeset
     1
(*:maxLineLen=78:*)
cfabbc083977 more uniform jEdit properties;
wenzelm
parents: 61572
diff changeset
     2
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     3
theory Inner_Syntax
42651
e3fdb7c96be5 formal Base theory;
wenzelm
parents: 42596
diff changeset
     4
imports Base Main
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     5
begin
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
     7
chapter \<open>Inner syntax --- the term language \label{ch:inner-syntax}\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     8
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
     9
text \<open>The inner syntax of Isabelle provides concrete notation for
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    10
  the main entities of the logical framework, notably \<open>\<lambda>\<close>-terms with types and type classes.  Applications may either
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    11
  extend existing syntactic categories by additional notation, or
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    12
  define new sub-languages that are linked to the standard term
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
    13
  language via some explicit markers.  For example \<^verbatim>\<open>FOO\<close>~\<open>foo\<close> could
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
    14
  embed the syntax corresponding for some
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    15
  user-defined nonterminal \<open>foo\<close> --- within the bounds of the
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    16
  given lexical syntax of Isabelle/Pure.
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    17
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    18
  The most basic way to specify concrete syntax for logical entities
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    19
  works via mixfix annotations (\secref{sec:mixfix}), which may be
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    20
  usually given as part of the original declaration or via explicit
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    21
  notation commands later on (\secref{sec:notation}).  This already
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    22
  covers many needs of concrete syntax without having to understand
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    23
  the full complexity of inner syntax layers.
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    24
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    25
  Further details of the syntax engine involves the classical
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    26
  distinction of lexical language versus context-free grammar (see
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
    27
  \secref{sec:pure-syntax}), and various mechanisms for \<^emph>\<open>syntax
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
    28
  transformations\<close> (see \secref{sec:syntax-transformations}).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    29
\<close>
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    30
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
    31
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    32
section \<open>Printing logical entities\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    33
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    34
subsection \<open>Diagnostic commands \label{sec:print-diag}\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    35
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    36
text \<open>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    37
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    38
    @{command_def "typ"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    39
    @{command_def "term"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    40
    @{command_def "prop"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    41
    @{command_def "thm"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    42
    @{command_def "prf"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    43
    @{command_def "full_prf"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    44
    @{command_def "print_state"}\<open>\<^sup>*\<close> & : & \<open>any \<rightarrow>\<close> \\
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    45
  \end{matharray}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    46
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    47
  These diagnostic commands assist interactive development by printing
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    48
  internal logical entities in a human-readable fashion.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    49
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55108
diff changeset
    50
  @{rail \<open>
48792
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48119
diff changeset
    51
    @@{command typ} @{syntax modes}? @{syntax type} ('::' @{syntax sort})?
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    52
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
    53
    @@{command term} @{syntax modes}? @{syntax term}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    54
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
    55
    @@{command prop} @{syntax modes}? @{syntax prop}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    56
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
    57
    @@{command thm} @{syntax modes}? @{syntax thmrefs}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    58
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
    59
    ( @@{command prf} | @@{command full_prf} ) @{syntax modes}? @{syntax thmrefs}?
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    60
    ;
52430
289e36c2870a proper diagnostic command 'print_state';
wenzelm
parents: 52414
diff changeset
    61
    @@{command print_state} @{syntax modes}?
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    62
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
    63
    @{syntax_def modes}: '(' (@{syntax name} + ) ')'
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55108
diff changeset
    64
  \<close>}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    65
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    66
  \<^descr> @{command "typ"}~\<open>\<tau>\<close> reads and prints a type expression
48792
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48119
diff changeset
    67
  according to the current context.
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48119
diff changeset
    68
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    69
  \<^descr> @{command "typ"}~\<open>\<tau> :: s\<close> uses type-inference to
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    70
  determine the most general way to make \<open>\<tau>\<close> conform to sort
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    71
  \<open>s\<close>.  For concrete \<open>\<tau>\<close> this checks if the type
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    72
  belongs to that sort.  Dummy type parameters ``\<open>_\<close>''
48792
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48119
diff changeset
    73
  (underscore) are assigned to fresh type variables with most general
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48119
diff changeset
    74
  sorts, according the the principles of type-inference.
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    75
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    76
  \<^descr> @{command "term"}~\<open>t\<close> and @{command "prop"}~\<open>\<phi>\<close>
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    77
  read, type-check and print terms or propositions according to the
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    78
  current theory or proof context; the inferred type of \<open>t\<close> is
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    79
  output as well.  Note that these commands are also useful in
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    80
  inspecting the current environment of term abbreviations.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    81
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    82
  \<^descr> @{command "thm"}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> retrieves
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    83
  theorems from the current theory or proof context.  Note that any
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    84
  attributes included in the theorem specifications are applied to a
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    85
  temporary context derived from the current theory or proof; the
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    86
  result is discarded, i.e.\ attributes involved in \<open>a\<^sub>1,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    87
  \<dots>, a\<^sub>n\<close> do not have any permanent effect.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    88
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
    89
  \<^descr> @{command "prf"} displays the (compact) proof term of the
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    90
  current proof state (if present), or of the given theorems. Note
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    91
  that this requires proof terms to be switched on for the current
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    92
  object logic (see the ``Proof terms'' section of the Isabelle
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    93
  reference manual for information on how to do this).
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    94
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
    95
  \<^descr> @{command "full_prf"} is like @{command "prf"}, but displays
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    96
  the full proof term, i.e.\ also displays information omitted in the
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    97
  compact proof term, which is denoted by ``\<open>_\<close>'' placeholders
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    98
  there.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    99
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   100
  \<^descr> @{command "print_state"} prints the current proof state (if
52430
289e36c2870a proper diagnostic command 'print_state';
wenzelm
parents: 52414
diff changeset
   101
  present), including current facts and goals.
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
   102
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   103
61997
4d9518c3d031 updated print modes;
wenzelm
parents: 61656
diff changeset
   104
  All of the diagnostic commands above admit a list of \<open>modes\<close> to be
4d9518c3d031 updated print modes;
wenzelm
parents: 61656
diff changeset
   105
  specified, which is appended to the current print mode; see also
4d9518c3d031 updated print modes;
wenzelm
parents: 61656
diff changeset
   106
  \secref{sec:print-modes}. Thus the output behavior may be modified according
4d9518c3d031 updated print modes;
wenzelm
parents: 61656
diff changeset
   107
  particular print mode features. For example, @{command
4d9518c3d031 updated print modes;
wenzelm
parents: 61656
diff changeset
   108
  "print_state"}~\<open>(latex)\<close> prints the current proof state with mathematical
4d9518c3d031 updated print modes;
wenzelm
parents: 61656
diff changeset
   109
  symbols and special characters represented in {\LaTeX} source, according to
4d9518c3d031 updated print modes;
wenzelm
parents: 61656
diff changeset
   110
  the Isabelle style @{cite "isabelle-system"}.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   111
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   112
  Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   113
  systematic way to include formal items into the printed text
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   114
  document.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   115
\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   116
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   117
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   118
subsection \<open>Details of printed content\<close>
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   119
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   120
text \<open>
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   121
  \begin{tabular}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   122
    @{attribute_def show_markup} & : & \<open>attribute\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   123
    @{attribute_def show_types} & : & \<open>attribute\<close> & default \<open>false\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   124
    @{attribute_def show_sorts} & : & \<open>attribute\<close> & default \<open>false\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   125
    @{attribute_def show_consts} & : & \<open>attribute\<close> & default \<open>false\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   126
    @{attribute_def show_abbrevs} & : & \<open>attribute\<close> & default \<open>true\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   127
    @{attribute_def show_brackets} & : & \<open>attribute\<close> & default \<open>false\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   128
    @{attribute_def names_long} & : & \<open>attribute\<close> & default \<open>false\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   129
    @{attribute_def names_short} & : & \<open>attribute\<close> & default \<open>false\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   130
    @{attribute_def names_unique} & : & \<open>attribute\<close> & default \<open>true\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   131
    @{attribute_def eta_contract} & : & \<open>attribute\<close> & default \<open>true\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   132
    @{attribute_def goals_limit} & : & \<open>attribute\<close> & default \<open>10\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   133
    @{attribute_def show_main_goal} & : & \<open>attribute\<close> & default \<open>false\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   134
    @{attribute_def show_hyps} & : & \<open>attribute\<close> & default \<open>false\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   135
    @{attribute_def show_tags} & : & \<open>attribute\<close> & default \<open>false\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   136
    @{attribute_def show_question_marks} & : & \<open>attribute\<close> & default \<open>true\<close> \\
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   137
  \end{tabular}
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   138
  \<^medskip>
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   139
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   140
  These configuration options control the detail of information that
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   141
  is displayed for types, terms, theorems, goals etc.  See also
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   142
  \secref{sec:config}.
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   143
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   144
  \<^descr> @{attribute show_markup} controls direct inlining of markup
49699
1301ed115729 some documentation of show_markup;
wenzelm
parents: 48985
diff changeset
   145
  into the printed representation of formal entities --- notably type
1301ed115729 some documentation of show_markup;
wenzelm
parents: 48985
diff changeset
   146
  and sort constraints.  This enables Prover IDE users to retrieve
1301ed115729 some documentation of show_markup;
wenzelm
parents: 48985
diff changeset
   147
  that information via tooltips or popups while hovering with the
1301ed115729 some documentation of show_markup;
wenzelm
parents: 48985
diff changeset
   148
  mouse over the output window, for example.  Consequently, this
58842
22b87ab47d3b discontinued Proof General;
wenzelm
parents: 58761
diff changeset
   149
  option is enabled by default for Isabelle/jEdit.
49699
1301ed115729 some documentation of show_markup;
wenzelm
parents: 48985
diff changeset
   150
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   151
  \<^descr> @{attribute show_types} and @{attribute show_sorts} control
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   152
  printing of type constraints for term variables, and sort
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   153
  constraints for type variables.  By default, neither of these are
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   154
  shown in output.  If @{attribute show_sorts} is enabled, types are
49699
1301ed115729 some documentation of show_markup;
wenzelm
parents: 48985
diff changeset
   155
  always shown as well.  In Isabelle/jEdit, manual setting of these
1301ed115729 some documentation of show_markup;
wenzelm
parents: 48985
diff changeset
   156
  options is normally not required thanks to @{attribute show_markup}
1301ed115729 some documentation of show_markup;
wenzelm
parents: 48985
diff changeset
   157
  above.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   158
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   159
  Note that displaying types and sorts may explain why a polymorphic
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   160
  inference rule fails to resolve with some goal, or why a rewrite
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   161
  rule does not apply as expected.
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   162
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   163
  \<^descr> @{attribute show_consts} controls printing of types of
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   164
  constants when displaying a goal state.
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   165
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   166
  Note that the output can be enormous, because polymorphic constants
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   167
  often occur at several different type instances.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   168
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   169
  \<^descr> @{attribute show_abbrevs} controls folding of constant
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   170
  abbreviations.
40879
ca132ef44944 configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents: 40255
diff changeset
   171
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   172
  \<^descr> @{attribute show_brackets} controls bracketing in pretty
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   173
  printed output.  If enabled, all sub-expressions of the pretty
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   174
  printing tree will be parenthesized, even if this produces malformed
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   175
  term syntax!  This crude way of showing the internal structure of
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   176
  pretty printed entities may occasionally help to diagnose problems
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   177
  with operator priorities, for example.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   178
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   179
  \<^descr> @{attribute names_long}, @{attribute names_short}, and
42669
04dfffda5671 more conventional naming scheme: names_long, names_short, names_unique;
wenzelm
parents: 42655
diff changeset
   180
  @{attribute names_unique} control the way of printing fully
42358
b47d41d9f4b5 Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
wenzelm
parents: 42279
diff changeset
   181
  qualified internal names in external form.  See also
b47d41d9f4b5 Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
wenzelm
parents: 42279
diff changeset
   182
  \secref{sec:antiq} for the document antiquotation options of the
b47d41d9f4b5 Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
wenzelm
parents: 42279
diff changeset
   183
  same names.
b47d41d9f4b5 Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
wenzelm
parents: 42279
diff changeset
   184
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   185
  \<^descr> @{attribute eta_contract} controls \<open>\<eta>\<close>-contracted
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   186
  printing of terms.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   187
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   188
  The \<open>\<eta>\<close>-contraction law asserts @{prop "(\<lambda>x. f x) \<equiv> f"},
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   189
  provided \<open>x\<close> is not free in \<open>f\<close>.  It asserts
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   190
  \<^emph>\<open>extensionality\<close> of functions: @{prop "f \<equiv> g"} if @{prop "f x \<equiv>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   191
  g x"} for all \<open>x\<close>.  Higher-order unification frequently puts
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   192
  terms into a fully \<open>\<eta>\<close>-expanded form.  For example, if \<open>F\<close> has type \<open>(\<tau> \<Rightarrow> \<tau>) \<Rightarrow> \<tau>\<close> then its expanded form is @{term
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   193
  "\<lambda>h. F (\<lambda>x. h x)"}.
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   194
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   195
  Enabling @{attribute eta_contract} makes Isabelle perform \<open>\<eta>\<close>-contractions before printing, so that @{term "\<lambda>h. F (\<lambda>x. h x)"}
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   196
  appears simply as \<open>F\<close>.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   197
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   198
  Note that the distinction between a term and its \<open>\<eta>\<close>-expanded
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   199
  form occasionally matters.  While higher-order resolution and
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   200
  rewriting operate modulo \<open>\<alpha>\<beta>\<eta>\<close>-conversion, some other tools
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   201
  might look at terms more discretely.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   202
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   203
  \<^descr> @{attribute goals_limit} controls the maximum number of
51960
61ac1efe02c3 option "goals_limit", with more uniform description;
wenzelm
parents: 51657
diff changeset
   204
  subgoals to be printed.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   205
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   206
  \<^descr> @{attribute show_main_goal} controls whether the main result
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   207
  to be proven should be displayed.  This information might be
39130
12dac4b58df8 updated configuration options;
wenzelm
parents: 39050
diff changeset
   208
  relevant for schematic goals, to inspect the current claim that has
12dac4b58df8 updated configuration options;
wenzelm
parents: 39050
diff changeset
   209
  been synthesized so far.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   210
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   211
  \<^descr> @{attribute show_hyps} controls printing of implicit
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   212
  hypotheses of local facts.  Normally, only those hypotheses are
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   213
  displayed that are \<^emph>\<open>not\<close> covered by the assumptions of the
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   214
  current context: this situation indicates a fault in some tool being
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   215
  used.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   216
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   217
  By enabling @{attribute show_hyps}, output of \<^emph>\<open>all\<close> hypotheses
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   218
  can be enforced, which is occasionally useful for diagnostic
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   219
  purposes.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   220
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   221
  \<^descr> @{attribute show_tags} controls printing of extra annotations
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   222
  within theorems, such as internal position information, or the case
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
   223
  names being attached by the attribute @{attribute case_names}.
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   224
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   225
  Note that the @{attribute tagged} and @{attribute untagged}
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   226
  attributes provide low-level access to the collection of tags
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   227
  associated with a theorem.
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   228
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   229
  \<^descr> @{attribute show_question_marks} controls printing of question
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   230
  marks for schematic variables, such as \<open>?x\<close>.  Only the leading
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   231
  question mark is affected, the remaining text is unchanged
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   232
  (including proper markup for schematic variables that might be
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   233
  relevant for user interfaces).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   234
\<close>
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   235
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   236
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   237
subsection \<open>Alternative print modes \label{sec:print-modes}\<close>
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   238
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   239
text \<open>
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   240
  \begin{mldecls}
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   241
    @{index_ML print_mode_value: "unit -> string list"} \\
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   242
    @{index_ML Print_Mode.with_modes: "string list -> ('a -> 'b) -> 'a -> 'b"} \\
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   243
  \end{mldecls}
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   244
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   245
  The \<^emph>\<open>print mode\<close> facility allows to modify various operations
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   246
  for printing.  Commands like @{command typ}, @{command term},
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   247
  @{command thm} (see \secref{sec:print-diag}) take additional print
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   248
  modes as optional argument.  The underlying ML operations are as
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   249
  follows.
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   250
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   251
  \<^descr> @{ML "print_mode_value ()"} yields the list of currently
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   252
  active print mode names.  This should be understood as symbolic
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   253
  representation of certain individual features for printing (with
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   254
  precedence from left to right).
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   255
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   256
  \<^descr> @{ML Print_Mode.with_modes}~\<open>modes f x\<close> evaluates
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   257
  \<open>f x\<close> in an execution context where the print mode is
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   258
  prepended by the given \<open>modes\<close>.  This provides a thread-safe
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   259
  way to augment print modes.  It is also monotonic in the set of mode
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   260
  names: it retains the default print mode that certain
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   261
  user-interfaces might have installed for their proper functioning!
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   262
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   263
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   264
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   265
  The pretty printer for inner syntax maintains alternative
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   266
  mixfix productions for any print mode name invented by the user, say
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   267
  in commands like @{command notation} or @{command abbreviation}.
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   268
  Mode names can be arbitrary, but the following ones have a specific
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   269
  meaning by convention:
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   270
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   271
  \<^item> \<^verbatim>\<open>""\<close> (the empty string): default mode;
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   272
  implicitly active as last element in the list of modes.
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   273
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   274
  \<^item> \<^verbatim>\<open>input\<close>: dummy print mode that is never active; may
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   275
  be used to specify notation that is only available for input.
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   276
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   277
  \<^item> \<^verbatim>\<open>internal\<close> dummy print mode that is never active;
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   278
  used internally in Isabelle/Pure.
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   279
61997
4d9518c3d031 updated print modes;
wenzelm
parents: 61656
diff changeset
   280
  \<^item> \<^verbatim>\<open>ASCII\<close>: prefer ASCII art over mathematical symbols.
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   281
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   282
  \<^item> \<^verbatim>\<open>latex\<close>: additional mode that is active in {\LaTeX}
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   283
  document preparation of Isabelle theory sources; allows to provide
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   284
  alternative output notation.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   285
\<close>
46284
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   286
6233d0b74d71 updated section on print modes;
wenzelm
parents: 46282
diff changeset
   287
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   288
section \<open>Mixfix annotations \label{sec:mixfix}\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   289
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   290
text \<open>Mixfix annotations specify concrete \<^emph>\<open>inner syntax\<close> of
35351
7425aece4ee3 allow general mixfix syntax for type constructors;
wenzelm
parents: 32833
diff changeset
   291
  Isabelle types and terms.  Locally fixed parameters in toplevel
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   292
  theorem statements, locale and class specifications also admit
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   293
  mixfix annotations in a fairly uniform manner.  A mixfix annotation
50635
wenzelm
parents: 49699
diff changeset
   294
  describes the concrete syntax, the translation to abstract
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   295
  syntax, and the pretty printing.  Special case annotations provide a
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   296
  simple means of specifying infix operators and binders.
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   297
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58421
diff changeset
   298
  Isabelle mixfix syntax is inspired by {\OBJ} @{cite OBJ}.  It allows
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   299
  to specify any context-free priority grammar, which is more general
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   300
  than the fixity declarations of ML and Prolog.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   301
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55108
diff changeset
   302
  @{rail \<open>
51654
8450b944e58a just one syntax category "mixfix" -- check structure annotation semantically;
wenzelm
parents: 50636
diff changeset
   303
    @{syntax_def mixfix}: '('
58761
b5ecbb1c4dc5 repaired rail diagram (cf. 8450b944e58a);
wenzelm
parents: 58726
diff changeset
   304
      (@{syntax template} prios? @{syntax nat}? |
b5ecbb1c4dc5 repaired rail diagram (cf. 8450b944e58a);
wenzelm
parents: 58726
diff changeset
   305
        (@'infix' | @'infixl' | @'infixr') @{syntax template} @{syntax nat} |
b5ecbb1c4dc5 repaired rail diagram (cf. 8450b944e58a);
wenzelm
parents: 58726
diff changeset
   306
        @'binder' @{syntax template} prios? @{syntax nat} |
b5ecbb1c4dc5 repaired rail diagram (cf. 8450b944e58a);
wenzelm
parents: 58726
diff changeset
   307
        @'structure') ')'
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   308
    ;
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   309
    template: string
46289
wenzelm
parents: 46288
diff changeset
   310
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
   311
    prios: '[' (@{syntax nat} + ',') ']'
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55108
diff changeset
   312
  \<close>}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   313
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   314
  The string given as \<open>template\<close> may include literal text,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   315
  spacing, blocks, and arguments (denoted by ``\<open>_\<close>''); the
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   316
  special symbol ``\<^verbatim>\<open>\<index>\<close>'' (printed as ``\<open>\<index>\<close>'')
51657
3db1bbc82d8d more accurate documentation of "(structure)" mixfix;
wenzelm
parents: 51654
diff changeset
   317
  represents an index argument that specifies an implicit @{keyword
3db1bbc82d8d more accurate documentation of "(structure)" mixfix;
wenzelm
parents: 51654
diff changeset
   318
  "structure"} reference (see also \secref{sec:locale}).  Only locally
3db1bbc82d8d more accurate documentation of "(structure)" mixfix;
wenzelm
parents: 51654
diff changeset
   319
  fixed variables may be declared as @{keyword "structure"}.
3db1bbc82d8d more accurate documentation of "(structure)" mixfix;
wenzelm
parents: 51654
diff changeset
   320
3db1bbc82d8d more accurate documentation of "(structure)" mixfix;
wenzelm
parents: 51654
diff changeset
   321
  Infix and binder declarations provide common abbreviations for
3db1bbc82d8d more accurate documentation of "(structure)" mixfix;
wenzelm
parents: 51654
diff changeset
   322
  particular mixfix declarations.  So in practice, mixfix templates
3db1bbc82d8d more accurate documentation of "(structure)" mixfix;
wenzelm
parents: 51654
diff changeset
   323
  mostly degenerate to literal text for concrete syntax, such as
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   324
  ``\<^verbatim>\<open>++\<close>'' for an infix symbol.
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   325
\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   326
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   327
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   328
subsection \<open>The general mixfix form\<close>
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   329
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   330
text \<open>In full generality, mixfix declarations work as follows.
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   331
  Suppose a constant \<open>c :: \<tau>\<^sub>1 \<Rightarrow> \<dots> \<tau>\<^sub>n \<Rightarrow> \<tau>\<close> is annotated by
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   332
  \<open>(mixfix [p\<^sub>1, \<dots>, p\<^sub>n] p)\<close>, where \<open>mixfix\<close> is a string
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   333
  \<open>d\<^sub>0 _ d\<^sub>1 _ \<dots> _ d\<^sub>n\<close> consisting of delimiters that surround
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   334
  argument positions as indicated by underscores.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   335
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   336
  Altogether this determines a production for a context-free priority
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   337
  grammar, where for each argument \<open>i\<close> the syntactic category
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   338
  is determined by \<open>\<tau>\<^sub>i\<close> (with priority \<open>p\<^sub>i\<close>), and the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   339
  result category is determined from \<open>\<tau>\<close> (with priority \<open>p\<close>).  Priority specifications are optional, with default 0 for
61572
ddb3ac3fef45 more antiquotations;
wenzelm
parents: 61503
diff changeset
   340
  arguments and 1000 for the result.\<^footnote>\<open>Omitting priorities is
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
   341
  prone to syntactic ambiguities unless the delimiter tokens determine
61572
ddb3ac3fef45 more antiquotations;
wenzelm
parents: 61503
diff changeset
   342
  fully bracketed notation, as in \<open>if _ then _ else _ fi\<close>.\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   343
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   344
  Since \<open>\<tau>\<close> may be again a function type, the constant
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   345
  type scheme may have more argument positions than the mixfix
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   346
  pattern.  Printing a nested application \<open>c t\<^sub>1 \<dots> t\<^sub>m\<close> for
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   347
  \<open>m > n\<close> works by attaching concrete notation only to the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   348
  innermost part, essentially by printing \<open>(c t\<^sub>1 \<dots> t\<^sub>n) \<dots> t\<^sub>m\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   349
  instead.  If a term has fewer arguments than specified in the mixfix
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   350
  template, the concrete syntax is ignored.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   351
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   352
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   353
  A mixfix template may also contain additional directives
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   354
  for pretty printing, notably spaces, blocks, and breaks.  The
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   355
  general template format is a sequence over any of the following
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   356
  entities.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   357
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   358
  \<^descr> \<open>d\<close> is a delimiter, namely a non-empty sequence of
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   359
  characters other than the following special characters:
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   360
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   361
  \<^medskip>
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   362
  \begin{tabular}{ll}
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   363
    \<^verbatim>\<open>'\<close> & single quote \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   364
    \<^verbatim>\<open>_\<close> & underscore \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   365
    \<open>\<index>\<close> & index symbol \\
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   366
    \<^verbatim>\<open>(\<close> & open parenthesis \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   367
    \<^verbatim>\<open>)\<close> & close parenthesis \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   368
    \<^verbatim>\<open>/\<close> & slash \\
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   369
  \end{tabular}
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   370
  \<^medskip>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   371
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   372
  \<^descr> \<^verbatim>\<open>'\<close> escapes the special meaning of these
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   373
  meta-characters, producing a literal version of the following
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   374
  character, unless that is a blank.
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   375
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   376
  A single quote followed by a blank separates delimiters, without
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   377
  affecting printing, but input tokens may have additional white space
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   378
  here.
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   379
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   380
  \<^descr> \<^verbatim>\<open>_\<close> is an argument position, which stands for a
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   381
  certain syntactic category in the underlying grammar.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   382
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   383
  \<^descr> \<open>\<index>\<close> is an indexed argument position; this is the place
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   384
  where implicit structure arguments can be attached.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   385
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   386
  \<^descr> \<open>s\<close> is a non-empty sequence of spaces for printing.
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   387
  This and the following specifications do not affect parsing at all.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   388
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   389
  \<^descr> \<^verbatim>\<open>(\<close>\<open>n\<close> opens a pretty printing block.  The
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   390
  optional number specifies how much indentation to add when a line
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   391
  break occurs within the block.  If the parenthesis is not followed
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   392
  by digits, the indentation defaults to 0.  A block specified via
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   393
  \<^verbatim>\<open>(00\<close> is unbreakable.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   394
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   395
  \<^descr> \<^verbatim>\<open>)\<close> closes a pretty printing block.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   396
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   397
  \<^descr> \<^verbatim>\<open>//\<close> forces a line break.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   398
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   399
  \<^descr> \<^verbatim>\<open>/\<close>\<open>s\<close> allows a line break.  Here \<open>s\<close>
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   400
  stands for the string of spaces (zero or more) right after the
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   401
  slash.  These spaces are printed if the break is \<^emph>\<open>not\<close> taken.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   402
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   403
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   404
  The general idea of pretty printing with blocks and breaks is also
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58421
diff changeset
   405
  described in @{cite "paulson-ml2"}; it goes back to @{cite "Oppen:1980"}.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   406
\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   407
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   408
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   409
subsection \<open>Infixes\<close>
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   410
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   411
text \<open>Infix operators are specified by convenient short forms that
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   412
  abbreviate general mixfix annotations as follows:
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   413
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   414
  \begin{center}
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   415
  \begin{tabular}{lll}
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   416
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   417
  \<^verbatim>\<open>(\<close>@{keyword_def "infix"}~\<^verbatim>\<open>"\<close>\<open>sy\<close>\<^verbatim>\<open>"\<close> \<open>p\<close>\<^verbatim>\<open>)\<close>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   418
  & \<open>\<mapsto>\<close> &
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   419
  \<^verbatim>\<open>("(_\<close>~\<open>sy\<close>\<^verbatim>\<open>/ _)" [\<close>\<open>p + 1\<close>\<^verbatim>\<open>,\<close>~\<open>p + 1\<close>\<^verbatim>\<open>]\<close>~\<open>p\<close>\<^verbatim>\<open>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   420
  \<^verbatim>\<open>(\<close>@{keyword_def "infixl"}~\<^verbatim>\<open>"\<close>\<open>sy\<close>\<^verbatim>\<open>"\<close> \<open>p\<close>\<^verbatim>\<open>)\<close>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   421
  & \<open>\<mapsto>\<close> &
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   422
  \<^verbatim>\<open>("(_\<close>~\<open>sy\<close>\<^verbatim>\<open>/ _)" [\<close>\<open>p\<close>\<^verbatim>\<open>,\<close>~\<open>p + 1\<close>\<^verbatim>\<open>]\<close>~\<open>p\<close>\<^verbatim>\<open>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   423
  \<^verbatim>\<open>(\<close>@{keyword_def "infixr"}~\<^verbatim>\<open>"\<close>\<open>sy\<close>\<^verbatim>\<open>"\<close>~\<open>p\<close>\<^verbatim>\<open>)\<close>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   424
  & \<open>\<mapsto>\<close> &
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   425
  \<^verbatim>\<open>("(_\<close>~\<open>sy\<close>\<^verbatim>\<open>/ _)" [\<close>\<open>p + 1\<close>\<^verbatim>\<open>,\<close>~\<open>p\<close>\<^verbatim>\<open>]\<close>~\<open>p\<close>\<^verbatim>\<open>)\<close> \\
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   426
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   427
  \end{tabular}
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   428
  \end{center}
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   429
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   430
  The mixfix template \<^verbatim>\<open>"(_\<close>~\<open>sy\<close>\<^verbatim>\<open>/ _)"\<close>
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
   431
  specifies two argument positions; the delimiter is preceded by a
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
   432
  space and followed by a space or line break; the entire phrase is a
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
   433
  pretty printing block.
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   434
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   435
  The alternative notation \<^verbatim>\<open>op\<close>~\<open>sy\<close> is introduced
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   436
  in addition.  Thus any infix operator may be written in prefix form
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   437
  (as in ML), independently of the number of arguments in the term.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   438
\<close>
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   439
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   440
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   441
subsection \<open>Binders\<close>
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   442
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   443
text \<open>A \<^emph>\<open>binder\<close> is a variable-binding construct such as a
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   444
  quantifier.  The idea to formalize \<open>\<forall>x. b\<close> as \<open>All
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   445
  (\<lambda>x. b)\<close> for \<open>All :: ('a \<Rightarrow> bool) \<Rightarrow> bool\<close> already goes back
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58421
diff changeset
   446
  to @{cite church40}.  Isabelle declarations of certain higher-order
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
   447
  operators may be annotated with @{keyword_def "binder"} annotations
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
   448
  as follows:
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   449
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   450
  \begin{center}
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   451
  \<open>c ::\<close>~\<^verbatim>\<open>"\<close>\<open>(\<tau>\<^sub>1 \<Rightarrow> \<tau>\<^sub>2) \<Rightarrow> \<tau>\<^sub>3\<close>\<^verbatim>\<open>"  (\<close>@{keyword "binder"}~\<^verbatim>\<open>"\<close>\<open>sy\<close>\<^verbatim>\<open>" [\<close>\<open>p\<close>\<^verbatim>\<open>]\<close>~\<open>q\<close>\<^verbatim>\<open>)\<close>
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   452
  \end{center}
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   453
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   454
  This introduces concrete binder syntax \<open>sy x. b\<close>, where
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   455
  \<open>x\<close> is a bound variable of type \<open>\<tau>\<^sub>1\<close>, the body \<open>b\<close> has type \<open>\<tau>\<^sub>2\<close> and the whole term has type \<open>\<tau>\<^sub>3\<close>.
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   456
  The optional integer \<open>p\<close> specifies the syntactic priority of
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   457
  the body; the default is \<open>q\<close>, which is also the priority of
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   458
  the whole construct.
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   459
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   460
  Internally, the binder syntax is expanded to something like this:
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   461
  \begin{center}
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   462
  \<open>c_binder ::\<close>~\<^verbatim>\<open>"\<close>\<open>idts \<Rightarrow> \<tau>\<^sub>2 \<Rightarrow> \<tau>\<^sub>3\<close>\<^verbatim>\<open>"  ("(3\<close>\<open>sy\<close>\<^verbatim>\<open>_./ _)" [0,\<close>~\<open>p\<close>\<^verbatim>\<open>]\<close>~\<open>q\<close>\<^verbatim>\<open>)\<close>
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   463
  \end{center}
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   464
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   465
  Here @{syntax (inner) idts} is the nonterminal symbol for a list of
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   466
  identifiers with optional type constraints (see also
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   467
  \secref{sec:pure-grammar}).  The mixfix template \<^verbatim>\<open>"(3\<close>\<open>sy\<close>\<^verbatim>\<open>_./ _)"\<close>
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   468
  defines argument positions
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   469
  for the bound identifiers and the body, separated by a dot with
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   470
  optional line break; the entire phrase is a pretty printing block of
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   471
  indentation level 3.  Note that there is no extra space after \<open>sy\<close>, so it needs to be included user specification if the binder
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   472
  syntax ends with a token that may be continued by an identifier
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   473
  token at the start of @{syntax (inner) idts}.
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   474
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   475
  Furthermore, a syntax translation to transforms \<open>c_binder x\<^sub>1
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   476
  \<dots> x\<^sub>n b\<close> into iterated application \<open>c (\<lambda>x\<^sub>1. \<dots> c (\<lambda>x\<^sub>n. b)\<dots>)\<close>.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   477
  This works in both directions, for parsing and printing.\<close>
46290
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   478
465851ba524e updated/unified section on mixfix annotations;
wenzelm
parents: 46289
diff changeset
   479
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   480
section \<open>Explicit notation \label{sec:notation}\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   481
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   482
text \<open>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   483
  \begin{matharray}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   484
    @{command_def "type_notation"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   485
    @{command_def "no_type_notation"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   486
    @{command_def "notation"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   487
    @{command_def "no_notation"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   488
    @{command_def "write"} & : & \<open>proof(state) \<rightarrow> proof(state)\<close> \\
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   489
  \end{matharray}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   490
46288
8a2c5dc0b00e more on explicit notation;
wenzelm
parents: 46287
diff changeset
   491
  Commands that introduce new logical entities (terms or types)
8a2c5dc0b00e more on explicit notation;
wenzelm
parents: 46287
diff changeset
   492
  usually allow to provide mixfix annotations on the spot, which is
8a2c5dc0b00e more on explicit notation;
wenzelm
parents: 46287
diff changeset
   493
  convenient for default notation.  Nonetheless, the syntax may be
8a2c5dc0b00e more on explicit notation;
wenzelm
parents: 46287
diff changeset
   494
  modified later on by declarations for explicit notation.  This
8a2c5dc0b00e more on explicit notation;
wenzelm
parents: 46287
diff changeset
   495
  allows to add or delete mixfix annotations for of existing logical
8a2c5dc0b00e more on explicit notation;
wenzelm
parents: 46287
diff changeset
   496
  entities within the current context.
8a2c5dc0b00e more on explicit notation;
wenzelm
parents: 46287
diff changeset
   497
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55108
diff changeset
   498
  @{rail \<open>
59783
00b62aa9f430 tuned syntax diagrams -- no duplication of "target";
wenzelm
parents: 58842
diff changeset
   499
    (@@{command type_notation} | @@{command no_type_notation}) @{syntax mode}? \<newline>
00b62aa9f430 tuned syntax diagrams -- no duplication of "target";
wenzelm
parents: 58842
diff changeset
   500
      (@{syntax nameref} @{syntax mixfix} + @'and')
35413
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   501
    ;
59783
00b62aa9f430 tuned syntax diagrams -- no duplication of "target";
wenzelm
parents: 58842
diff changeset
   502
    (@@{command notation} | @@{command no_notation}) @{syntax mode}? \<newline>
51654
8450b944e58a just one syntax category "mixfix" -- check structure annotation semantically;
wenzelm
parents: 50636
diff changeset
   503
      (@{syntax nameref} @{syntax mixfix} + @'and')
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   504
    ;
51654
8450b944e58a just one syntax category "mixfix" -- check structure annotation semantically;
wenzelm
parents: 50636
diff changeset
   505
    @@{command write} @{syntax mode}? (@{syntax nameref} @{syntax mixfix} + @'and')
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55108
diff changeset
   506
  \<close>}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   507
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   508
  \<^descr> @{command "type_notation"}~\<open>c (mx)\<close> associates mixfix
35413
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   509
  syntax with an existing type constructor.  The arity of the
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   510
  constructor is retrieved from the context.
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   511
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   512
  \<^descr> @{command "no_type_notation"} is similar to @{command
35413
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   513
  "type_notation"}, but removes the specified syntax annotation from
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   514
  the present context.
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   515
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   516
  \<^descr> @{command "notation"}~\<open>c (mx)\<close> associates mixfix
35413
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   517
  syntax with an existing constant or fixed variable.  The type
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   518
  declaration of the given entity is retrieved from the context.
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   519
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   520
  \<^descr> @{command "no_notation"} is similar to @{command "notation"},
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   521
  but removes the specified syntax annotation from the present
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   522
  context.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   523
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   524
  \<^descr> @{command "write"} is similar to @{command "notation"}, but
36508
03d2a2d0ee4a allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
wenzelm
parents: 35413
diff changeset
   525
  works within an Isar proof body.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   526
\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   527
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   528
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   529
section \<open>The Pure syntax \label{sec:pure-syntax}\<close>
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   530
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   531
subsection \<open>Lexical matters \label{sec:inner-lex}\<close>
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   532
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   533
text \<open>The inner lexical syntax vaguely resembles the outer one
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   534
  (\secref{sec:outer-lex}), but some details are different.  There are
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   535
  two main categories of inner syntax tokens:
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   536
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   537
  \<^enum> \<^emph>\<open>delimiters\<close> --- the literal tokens occurring in
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   538
  productions of the given priority grammar (cf.\
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   539
  \secref{sec:priority-grammar});
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   540
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   541
  \<^enum> \<^emph>\<open>named tokens\<close> --- various categories of identifiers etc.
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   542
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   543
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   544
  Delimiters override named tokens and may thus render certain
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   545
  identifiers inaccessible.  Sometimes the logical context admits
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   546
  alternative ways to refer to the same entity, potentially via
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   547
  qualified names.
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   548
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   549
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   550
  The categories for named tokens are defined once and for
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   551
  all as follows, reusing some categories of the outer token syntax
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   552
  (\secref{sec:outer-lex}).
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   553
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   554
  \begin{center}
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   555
  \begin{supertabular}{rcl}
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   556
    @{syntax_def (inner) id} & = & @{syntax_ref ident} \\
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   557
    @{syntax_def (inner) longid} & = & @{syntax_ref longident} \\
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   558
    @{syntax_def (inner) var} & = & @{syntax_ref var} \\
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   559
    @{syntax_def (inner) tid} & = & @{syntax_ref typefree} \\
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   560
    @{syntax_def (inner) tvar} & = & @{syntax_ref typevar} \\
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58409
diff changeset
   561
    @{syntax_def (inner) num_token} & = & @{syntax_ref nat} \\
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   562
    @{syntax_def (inner) float_token} & = & @{syntax_ref nat}\<^verbatim>\<open>.\<close>@{syntax_ref nat} \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   563
    @{syntax_def (inner) str_token} & = & \<^verbatim>\<open>''\<close> \<open>\<dots>\<close> \<^verbatim>\<open>''\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   564
    @{syntax_def (inner) string_token} & = & \<^verbatim>\<open>"\<close> \<open>\<dots>\<close> \<^verbatim>\<open>"\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   565
    @{syntax_def (inner) cartouche} & = & @{verbatim "\<open>"} \<open>\<dots>\<close> @{verbatim "\<close>"} \\
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   566
  \end{supertabular}
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   567
  \end{center}
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   568
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   569
  The token categories @{syntax (inner) num_token}, @{syntax (inner)
58421
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 58410
diff changeset
   570
  float_token}, @{syntax (inner) str_token}, @{syntax (inner) string_token},
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 58410
diff changeset
   571
  and @{syntax (inner) cartouche} are not used in Pure. Object-logics may
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 58410
diff changeset
   572
  implement numerals and string literals by adding appropriate syntax
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 58410
diff changeset
   573
  declarations, together with some translation functions (e.g.\ see @{file
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 58410
diff changeset
   574
  "~~/src/HOL/Tools/string_syntax.ML"}).
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   575
58421
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 58410
diff changeset
   576
  The derived categories @{syntax_def (inner) num_const}, and @{syntax_def
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 58410
diff changeset
   577
  (inner) float_const}, provide robust access to the respective tokens: the
37cbbd8eb460 discontinued old "xnum" token category;
wenzelm
parents: 58410
diff changeset
   578
  syntax tree holds a syntactic constant instead of a free variable.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   579
\<close>
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   580
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   581
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   582
subsection \<open>Priority grammars \label{sec:priority-grammar}\<close>
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   583
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   584
text \<open>A context-free grammar consists of a set of \<^emph>\<open>terminal
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   585
  symbols\<close>, a set of \<^emph>\<open>nonterminal symbols\<close> and a set of
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   586
  \<^emph>\<open>productions\<close>.  Productions have the form \<open>A = \<gamma>\<close>,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   587
  where \<open>A\<close> is a nonterminal and \<open>\<gamma>\<close> is a string of
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   588
  terminals and nonterminals.  One designated nonterminal is called
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   589
  the \<^emph>\<open>root symbol\<close>.  The language defined by the grammar
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   590
  consists of all strings of terminals that can be derived from the
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   591
  root symbol by applying productions as rewrite rules.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   592
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   593
  The standard Isabelle parser for inner syntax uses a \<^emph>\<open>priority
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   594
  grammar\<close>.  Each nonterminal is decorated by an integer priority:
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   595
  \<open>A\<^sup>(\<^sup>p\<^sup>)\<close>.  In a derivation, \<open>A\<^sup>(\<^sup>p\<^sup>)\<close> may be rewritten
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   596
  using a production \<open>A\<^sup>(\<^sup>q\<^sup>) = \<gamma>\<close> only if \<open>p \<le> q\<close>.  Any
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   597
  priority grammar can be translated into a normal context-free
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   598
  grammar by introducing new nonterminals and productions.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   599
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   600
  \<^medskip>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   601
  Formally, a set of context free productions \<open>G\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   602
  induces a derivation relation \<open>\<longrightarrow>\<^sub>G\<close> as follows.  Let \<open>\<alpha>\<close> and \<open>\<beta>\<close> denote strings of terminal or nonterminal symbols.
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   603
  Then \<open>\<alpha> A\<^sup>(\<^sup>p\<^sup>) \<beta> \<longrightarrow>\<^sub>G \<alpha> \<gamma> \<beta>\<close> holds if and only if \<open>G\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   604
  contains some production \<open>A\<^sup>(\<^sup>q\<^sup>) = \<gamma>\<close> for \<open>p \<le> q\<close>.
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   605
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   606
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   607
  The following grammar for arithmetic expressions
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   608
  demonstrates how binding power and associativity of operators can be
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   609
  enforced by priorities.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   610
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   611
  \begin{center}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   612
  \begin{tabular}{rclr}
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   613
  \<open>A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\<close> & \<open>=\<close> & \<^verbatim>\<open>(\<close> \<open>A\<^sup>(\<^sup>0\<^sup>)\<close> \<^verbatim>\<open>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   614
  \<open>A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\<close> & \<open>=\<close> & \<^verbatim>\<open>0\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   615
  \<open>A\<^sup>(\<^sup>0\<^sup>)\<close> & \<open>=\<close> & \<open>A\<^sup>(\<^sup>0\<^sup>)\<close> \<^verbatim>\<open>+\<close> \<open>A\<^sup>(\<^sup>1\<^sup>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   616
  \<open>A\<^sup>(\<^sup>2\<^sup>)\<close> & \<open>=\<close> & \<open>A\<^sup>(\<^sup>3\<^sup>)\<close> \<^verbatim>\<open>*\<close> \<open>A\<^sup>(\<^sup>2\<^sup>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   617
  \<open>A\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>=\<close> & \<^verbatim>\<open>-\<close> \<open>A\<^sup>(\<^sup>3\<^sup>)\<close> \\
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   618
  \end{tabular}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   619
  \end{center}
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   620
  The choice of priorities determines that \<^verbatim>\<open>-\<close> binds
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   621
  tighter than \<^verbatim>\<open>*\<close>, which binds tighter than \<^verbatim>\<open>+\<close>.
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   622
  Furthermore \<^verbatim>\<open>+\<close> associates to the left and
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   623
  \<^verbatim>\<open>*\<close> to the right.
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   624
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   625
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   626
  For clarity, grammars obey these conventions:
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   627
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   628
  \<^item> All priorities must lie between 0 and 1000.
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   629
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   630
  \<^item> Priority 0 on the right-hand side and priority 1000 on the
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   631
  left-hand side may be omitted.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   632
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   633
  \<^item> The production \<open>A\<^sup>(\<^sup>p\<^sup>) = \<alpha>\<close> is written as \<open>A = \<alpha>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   634
  (p)\<close>, i.e.\ the priority of the left-hand side actually appears in
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   635
  a column on the far right.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   636
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   637
  \<^item> Alternatives are separated by \<open>|\<close>.
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   638
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   639
  \<^item> Repetition is indicated by dots \<open>(\<dots>)\<close> in an informal
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   640
  but obvious way.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   641
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   642
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   643
  Using these conventions, the example grammar specification above
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   644
  takes the form:
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   645
  \begin{center}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   646
  \begin{tabular}{rclc}
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   647
    \<open>A\<close> & \<open>=\<close> & \<^verbatim>\<open>(\<close> \<open>A\<close> \<^verbatim>\<open>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   648
              & \<open>|\<close> & \<^verbatim>\<open>0\<close> & \qquad\qquad \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   649
              & \<open>|\<close> & \<open>A\<close> \<^verbatim>\<open>+\<close> \<open>A\<^sup>(\<^sup>1\<^sup>)\<close> & \<open>(0)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   650
              & \<open>|\<close> & \<open>A\<^sup>(\<^sup>3\<^sup>)\<close> \<^verbatim>\<open>*\<close> \<open>A\<^sup>(\<^sup>2\<^sup>)\<close> & \<open>(2)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   651
              & \<open>|\<close> & \<^verbatim>\<open>-\<close> \<open>A\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>(3)\<close> \\
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   652
  \end{tabular}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   653
  \end{center}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   654
\<close>
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   655
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   656
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   657
subsection \<open>The Pure grammar \label{sec:pure-grammar}\<close>
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   658
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   659
text \<open>The priority grammar of the \<open>Pure\<close> theory is defined
46287
0bb3d8ee5d25 more accurate Pure grammar;
wenzelm
parents: 46286
diff changeset
   660
  approximately like this:
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   661
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   662
  \begin{center}
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   663
  \begin{supertabular}{rclr}
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   664
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   665
  @{syntax_def (inner) any} & = & \<open>prop  |  logic\<close> \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   666
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   667
  @{syntax_def (inner) prop} & = & \<^verbatim>\<open>(\<close> \<open>prop\<close> \<^verbatim>\<open>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   668
    & \<open>|\<close> & \<open>prop\<^sup>(\<^sup>4\<^sup>)\<close> \<^verbatim>\<open>::\<close> \<open>type\<close> & \<open>(3)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   669
    & \<open>|\<close> & \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> \<^verbatim>\<open>==\<close> \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>(2)\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   670
    & \<open>|\<close> & \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> \<open>\<equiv>\<close> \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>(2)\<close> \\
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   671
    & \<open>|\<close> & \<open>prop\<^sup>(\<^sup>3\<^sup>)\<close> \<^verbatim>\<open>&&&\<close> \<open>prop\<^sup>(\<^sup>2\<^sup>)\<close> & \<open>(2)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   672
    & \<open>|\<close> & \<open>prop\<^sup>(\<^sup>2\<^sup>)\<close> \<^verbatim>\<open>==>\<close> \<open>prop\<^sup>(\<^sup>1\<^sup>)\<close> & \<open>(1)\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   673
    & \<open>|\<close> & \<open>prop\<^sup>(\<^sup>2\<^sup>)\<close> \<open>\<Longrightarrow>\<close> \<open>prop\<^sup>(\<^sup>1\<^sup>)\<close> & \<open>(1)\<close> \\
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   674
    & \<open>|\<close> & \<^verbatim>\<open>[|\<close> \<open>prop\<close> \<^verbatim>\<open>;\<close> \<open>\<dots>\<close> \<^verbatim>\<open>;\<close> \<open>prop\<close> \<^verbatim>\<open>|]\<close> \<^verbatim>\<open>==>\<close> \<open>prop\<^sup>(\<^sup>1\<^sup>)\<close> & \<open>(1)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   675
    & \<open>|\<close> & \<open>\<lbrakk>\<close> \<open>prop\<close> \<^verbatim>\<open>;\<close> \<open>\<dots>\<close> \<^verbatim>\<open>;\<close> \<open>prop\<close> \<open>\<rbrakk>\<close> \<open>\<Longrightarrow>\<close> \<open>prop\<^sup>(\<^sup>1\<^sup>)\<close> & \<open>(1)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   676
    & \<open>|\<close> & \<^verbatim>\<open>!!\<close> \<open>idts\<close> \<^verbatim>\<open>.\<close> \<open>prop\<close> & \<open>(0)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   677
    & \<open>|\<close> & \<open>\<And>\<close> \<open>idts\<close> \<^verbatim>\<open>.\<close> \<open>prop\<close> & \<open>(0)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   678
    & \<open>|\<close> & \<^verbatim>\<open>OFCLASS\<close> \<^verbatim>\<open>(\<close> \<open>type\<close> \<^verbatim>\<open>,\<close> \<open>logic\<close> \<^verbatim>\<open>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   679
    & \<open>|\<close> & \<^verbatim>\<open>SORT_CONSTRAINT\<close> \<^verbatim>\<open>(\<close> \<open>type\<close> \<^verbatim>\<open>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   680
    & \<open>|\<close> & \<^verbatim>\<open>TERM\<close> \<open>logic\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   681
    & \<open>|\<close> & \<^verbatim>\<open>PROP\<close> \<open>aprop\<close> \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   682
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   683
  @{syntax_def (inner) aprop} & = & \<^verbatim>\<open>(\<close> \<open>aprop\<close> \<^verbatim>\<open>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   684
    & \<open>|\<close> & \<open>id  |  longid  |  var  |\<close>~~\<^verbatim>\<open>_\<close>~~\<open>|\<close>~~\<^verbatim>\<open>...\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   685
    & \<open>|\<close> & \<^verbatim>\<open>CONST\<close> \<open>id  |\<close>~~\<^verbatim>\<open>CONST\<close> \<open>longid\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   686
    & \<open>|\<close> & \<^verbatim>\<open>XCONST\<close> \<open>id  |\<close>~~\<^verbatim>\<open>XCONST\<close> \<open>longid\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   687
    & \<open>|\<close> & \<open>logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)  any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \<dots> any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\<close> & \<open>(999)\<close> \\\\
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   688
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   689
  @{syntax_def (inner) logic} & = & \<^verbatim>\<open>(\<close> \<open>logic\<close> \<^verbatim>\<open>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   690
    & \<open>|\<close> & \<open>logic\<^sup>(\<^sup>4\<^sup>)\<close> \<^verbatim>\<open>::\<close> \<open>type\<close> & \<open>(3)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   691
    & \<open>|\<close> & \<open>id  |  longid  |  var  |\<close>~~\<^verbatim>\<open>_\<close>~~\<open>|\<close>~~\<^verbatim>\<open>...\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   692
    & \<open>|\<close> & \<^verbatim>\<open>CONST\<close> \<open>id  |\<close>~~\<^verbatim>\<open>CONST\<close> \<open>longid\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   693
    & \<open>|\<close> & \<^verbatim>\<open>XCONST\<close> \<open>id  |\<close>~~\<^verbatim>\<open>XCONST\<close> \<open>longid\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   694
    & \<open>|\<close> & \<open>logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)  any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \<dots> any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\<close> & \<open>(999)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   695
    & \<open>|\<close> & \<open>\<struct> index\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\<close> \\
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   696
    & \<open>|\<close> & \<^verbatim>\<open>%\<close> \<open>pttrns\<close> \<^verbatim>\<open>.\<close> \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>(3)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   697
    & \<open>|\<close> & \<open>\<lambda>\<close> \<open>pttrns\<close> \<^verbatim>\<open>.\<close> \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>(3)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   698
    & \<open>|\<close> & \<^verbatim>\<open>op\<close> \<^verbatim>\<open>==\<close>~~\<open>|\<close>~~\<^verbatim>\<open>op\<close> \<open>\<equiv>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>op\<close> \<^verbatim>\<open>&&&\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   699
    & \<open>|\<close> & \<^verbatim>\<open>op\<close> \<^verbatim>\<open>==>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>op\<close> \<open>\<Longrightarrow>\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   700
    & \<open>|\<close> & \<^verbatim>\<open>TYPE\<close> \<^verbatim>\<open>(\<close> \<open>type\<close> \<^verbatim>\<open>)\<close> \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   701
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   702
  @{syntax_def (inner) idt} & = & \<^verbatim>\<open>(\<close> \<open>idt\<close> \<^verbatim>\<open>)\<close>~~\<open>|  id  |\<close>~~\<^verbatim>\<open>_\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   703
    & \<open>|\<close> & \<open>id\<close> \<^verbatim>\<open>::\<close> \<open>type\<close> & \<open>(0)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   704
    & \<open>|\<close> & \<^verbatim>\<open>_\<close> \<^verbatim>\<open>::\<close> \<open>type\<close> & \<open>(0)\<close> \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   705
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   706
  @{syntax_def (inner) index} & = & \<^verbatim>\<open>\<^bsub>\<close> \<open>logic\<^sup>(\<^sup>0\<^sup>)\<close> \<^verbatim>\<open>\<^esub>\<close>~~\<open>|  |  \<index>\<close> \\\\
46287
0bb3d8ee5d25 more accurate Pure grammar;
wenzelm
parents: 46286
diff changeset
   707
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   708
  @{syntax_def (inner) idts} & = & \<open>idt  |  idt\<^sup>(\<^sup>1\<^sup>) idts\<close> & \<open>(0)\<close> \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   709
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   710
  @{syntax_def (inner) pttrn} & = & \<open>idt\<close> \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   711
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   712
  @{syntax_def (inner) pttrns} & = & \<open>pttrn  |  pttrn\<^sup>(\<^sup>1\<^sup>) pttrns\<close> & \<open>(0)\<close> \\\\
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   713
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   714
  @{syntax_def (inner) type} & = & \<^verbatim>\<open>(\<close> \<open>type\<close> \<^verbatim>\<open>)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   715
    & \<open>|\<close> & \<open>tid  |  tvar  |\<close>~~\<^verbatim>\<open>_\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   716
    & \<open>|\<close> & \<open>tid\<close> \<^verbatim>\<open>::\<close> \<open>sort  |  tvar\<close>~~\<^verbatim>\<open>::\<close> \<open>sort  |\<close>~~\<^verbatim>\<open>_\<close> \<^verbatim>\<open>::\<close> \<open>sort\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   717
    & \<open>|\<close> & \<open>type_name  |  type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) type_name\<close> \\
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   718
    & \<open>|\<close> & \<^verbatim>\<open>(\<close> \<open>type\<close> \<^verbatim>\<open>,\<close> \<open>\<dots>\<close> \<^verbatim>\<open>,\<close> \<open>type\<close> \<^verbatim>\<open>)\<close> \<open>type_name\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   719
    & \<open>|\<close> & \<open>type\<^sup>(\<^sup>1\<^sup>)\<close> \<^verbatim>\<open>=>\<close> \<open>type\<close> & \<open>(0)\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   720
    & \<open>|\<close> & \<open>type\<^sup>(\<^sup>1\<^sup>)\<close> \<open>\<Rightarrow>\<close> \<open>type\<close> & \<open>(0)\<close> \\
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   721
    & \<open>|\<close> & \<^verbatim>\<open>[\<close> \<open>type\<close> \<^verbatim>\<open>,\<close> \<open>\<dots>\<close> \<^verbatim>\<open>,\<close> \<open>type\<close> \<^verbatim>\<open>]\<close> \<^verbatim>\<open>=>\<close> \<open>type\<close> & \<open>(0)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   722
    & \<open>|\<close> & \<^verbatim>\<open>[\<close> \<open>type\<close> \<^verbatim>\<open>,\<close> \<open>\<dots>\<close> \<^verbatim>\<open>,\<close> \<open>type\<close> \<^verbatim>\<open>]\<close> \<open>\<Rightarrow>\<close> \<open>type\<close> & \<open>(0)\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   723
  @{syntax_def (inner) type_name} & = & \<open>id  |  longid\<close> \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   724
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   725
  @{syntax_def (inner) sort} & = & @{syntax class_name}~~\<open>|\<close>~~\<^verbatim>\<open>{}\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   726
    & \<open>|\<close> & \<^verbatim>\<open>{\<close> @{syntax class_name} \<^verbatim>\<open>,\<close> \<open>\<dots>\<close> \<^verbatim>\<open>,\<close> @{syntax class_name} \<^verbatim>\<open>}\<close> \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   727
  @{syntax_def (inner) class_name} & = & \<open>id  |  longid\<close> \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   728
  \end{supertabular}
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   729
  \end{center}
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   730
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   731
  \<^medskip>
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   732
  Here literal terminals are printed \<^verbatim>\<open>verbatim\<close>;
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   733
  see also \secref{sec:inner-lex} for further token categories of the
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   734
  inner syntax.  The meaning of the nonterminals defined by the above
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   735
  grammar is as follows:
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   736
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   737
  \<^descr> @{syntax_ref (inner) any} denotes any term.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   738
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   739
  \<^descr> @{syntax_ref (inner) prop} denotes meta-level propositions,
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   740
  which are terms of type @{typ prop}.  The syntax of such formulae of
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   741
  the meta-logic is carefully distinguished from usual conventions for
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   742
  object-logics.  In particular, plain \<open>\<lambda>\<close>-term notation is
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   743
  \<^emph>\<open>not\<close> recognized as @{syntax (inner) prop}.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   744
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   745
  \<^descr> @{syntax_ref (inner) aprop} denotes atomic propositions, which
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   746
  are embedded into regular @{syntax (inner) prop} by means of an
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   747
  explicit \<^verbatim>\<open>PROP\<close> token.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   748
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   749
  Terms of type @{typ prop} with non-constant head, e.g.\ a plain
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   750
  variable, are printed in this form.  Constants that yield type @{typ
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   751
  prop} are expected to provide their own concrete syntax; otherwise
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   752
  the printed version will appear like @{syntax (inner) logic} and
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   753
  cannot be parsed again as @{syntax (inner) prop}.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   754
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   755
  \<^descr> @{syntax_ref (inner) logic} denotes arbitrary terms of a
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   756
  logical type, excluding type @{typ prop}.  This is the main
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   757
  syntactic category of object-logic entities, covering plain \<open>\<lambda>\<close>-term notation (variables, abstraction, application), plus
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   758
  anything defined by the user.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   759
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   760
  When specifying notation for logical entities, all logical types
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   761
  (excluding @{typ prop}) are \<^emph>\<open>collapsed\<close> to this single category
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   762
  of @{syntax (inner) logic}.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   763
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   764
  \<^descr> @{syntax_ref (inner) index} denotes an optional index term for
51657
3db1bbc82d8d more accurate documentation of "(structure)" mixfix;
wenzelm
parents: 51654
diff changeset
   765
  indexed syntax.  If omitted, it refers to the first @{keyword_ref
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   766
  "structure"} variable in the context.  The special dummy ``\<open>\<index>\<close>'' serves as pattern variable in mixfix annotations that
46287
0bb3d8ee5d25 more accurate Pure grammar;
wenzelm
parents: 46286
diff changeset
   767
  introduce indexed notation.
0bb3d8ee5d25 more accurate Pure grammar;
wenzelm
parents: 46286
diff changeset
   768
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   769
  \<^descr> @{syntax_ref (inner) idt} denotes identifiers, possibly
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   770
  constrained by types.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   771
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   772
  \<^descr> @{syntax_ref (inner) idts} denotes a sequence of @{syntax_ref
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   773
  (inner) idt}.  This is the most basic category for variables in
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   774
  iterated binders, such as \<open>\<lambda>\<close> or \<open>\<And>\<close>.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   775
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   776
  \<^descr> @{syntax_ref (inner) pttrn} and @{syntax_ref (inner) pttrns}
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   777
  denote patterns for abstraction, cases bindings etc.  In Pure, these
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   778
  categories start as a merely copy of @{syntax (inner) idt} and
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   779
  @{syntax (inner) idts}, respectively.  Object-logics may add
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   780
  additional productions for binding forms.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   781
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   782
  \<^descr> @{syntax_ref (inner) type} denotes types of the meta-logic.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   783
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   784
  \<^descr> @{syntax_ref (inner) sort} denotes meta-level sorts.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   785
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   786
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   787
  Here are some further explanations of certain syntax features.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   788
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   789
  \<^item> In @{syntax (inner) idts}, note that \<open>x :: nat y\<close> is
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   790
  parsed as \<open>x :: (nat y)\<close>, treating \<open>y\<close> like a type
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   791
  constructor applied to \<open>nat\<close>.  To avoid this interpretation,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   792
  write \<open>(x :: nat) y\<close> with explicit parentheses.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   793
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   794
  \<^item> Similarly, \<open>x :: nat y :: nat\<close> is parsed as \<open>x ::
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   795
  (nat y :: nat)\<close>.  The correct form is \<open>(x :: nat) (y ::
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   796
  nat)\<close>, or \<open>(x :: nat) y :: nat\<close> if \<open>y\<close> is last in the
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   797
  sequence of identifiers.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   798
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   799
  \<^item> Type constraints for terms bind very weakly.  For example,
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   800
  \<open>x < y :: nat\<close> is normally parsed as \<open>(x < y) ::
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   801
  nat\<close>, unless \<open><\<close> has a very low priority, in which case the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   802
  input is likely to be ambiguous.  The correct form is \<open>x < (y
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   803
  :: nat)\<close>.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   804
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   805
  \<^item> Dummy variables (written as underscore) may occur in different
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   806
  roles.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   807
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   808
    \<^descr> A type ``\<open>_\<close>'' or ``\<open>_ :: sort\<close>'' acts like an
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   809
    anonymous inference parameter, which is filled-in according to the
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   810
    most general type produced by the type-checking phase.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   811
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   812
    \<^descr> A bound ``\<open>_\<close>'' refers to a vacuous abstraction, where
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   813
    the body does not refer to the binding introduced here.  As in the
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   814
    term @{term "\<lambda>x _. x"}, which is \<open>\<alpha>\<close>-equivalent to \<open>\<lambda>x y. x\<close>.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   815
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   816
    \<^descr> A free ``\<open>_\<close>'' refers to an implicit outer binding.
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   817
    Higher definitional packages usually allow forms like \<open>f x _
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   818
    = x\<close>.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   819
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   820
    \<^descr> A schematic ``\<open>_\<close>'' (within a term pattern, see
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   821
    \secref{sec:term-decls}) refers to an anonymous variable that is
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   822
    implicitly abstracted over its context of locally bound variables.
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   823
    For example, this allows pattern matching of \<open>{x. f x = g
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   824
    x}\<close> against \<open>{x. _ = _}\<close>, or even \<open>{_. _ = _}\<close> by
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   825
    using both bound and schematic dummies.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   826
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   827
  \<^descr> The three literal dots ``\<^verbatim>\<open>...\<close>'' may be also
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   828
  written as ellipsis symbol \<^verbatim>\<open>\<dots>\<close>.  In both cases this
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   829
  refers to a special schematic variable, which is bound in the
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   830
  context.  This special term abbreviation works nicely with
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   831
  calculational reasoning (\secref{sec:calculation}).
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   832
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   833
  \<^descr> \<^verbatim>\<open>CONST\<close> ensures that the given identifier is treated
46287
0bb3d8ee5d25 more accurate Pure grammar;
wenzelm
parents: 46286
diff changeset
   834
  as constant term, and passed through the parse tree in fully
0bb3d8ee5d25 more accurate Pure grammar;
wenzelm
parents: 46286
diff changeset
   835
  internalized form.  This is particularly relevant for translation
0bb3d8ee5d25 more accurate Pure grammar;
wenzelm
parents: 46286
diff changeset
   836
  rules (\secref{sec:syn-trans}), notably on the RHS.
0bb3d8ee5d25 more accurate Pure grammar;
wenzelm
parents: 46286
diff changeset
   837
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   838
  \<^descr> \<^verbatim>\<open>XCONST\<close> is similar to \<^verbatim>\<open>CONST\<close>, but
46287
0bb3d8ee5d25 more accurate Pure grammar;
wenzelm
parents: 46286
diff changeset
   839
  retains the constant name as given.  This is only relevant to
0bb3d8ee5d25 more accurate Pure grammar;
wenzelm
parents: 46286
diff changeset
   840
  translation rules (\secref{sec:syn-trans}), notably on the LHS.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   841
\<close>
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   842
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   843
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   844
subsection \<open>Inspecting the syntax\<close>
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   845
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   846
text \<open>
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   847
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   848
    @{command_def "print_syntax"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   849
  \end{matharray}
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   850
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   851
  \<^descr> @{command "print_syntax"} prints the inner syntax of the
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   852
  current context.  The output can be quite large; the most important
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   853
  sections are explained below.
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   854
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   855
    \<^descr> \<open>lexicon\<close> lists the delimiters of the inner token
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   856
    language; see \secref{sec:inner-lex}.
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   857
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   858
    \<^descr> \<open>prods\<close> lists the productions of the underlying
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   859
    priority grammar; see \secref{sec:priority-grammar}.
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   860
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   861
    The nonterminal \<open>A\<^sup>(\<^sup>p\<^sup>)\<close> is rendered in plain text as \<open>A[p]\<close>; delimiters are quoted.  Many productions have an extra
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   862
    \<open>\<dots> => name\<close>.  These names later become the heads of parse
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   863
    trees; they also guide the pretty printer.
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   864
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   865
    Productions without such parse tree names are called \<^emph>\<open>copy
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   866
    productions\<close>.  Their right-hand side must have exactly one
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   867
    nonterminal symbol (or named token).  The parser does not create a
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   868
    new parse tree node for copy productions, but simply returns the
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   869
    parse tree of the right-hand symbol.
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   870
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   871
    If the right-hand side of a copy production consists of a single
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   872
    nonterminal without any delimiters, then it is called a \<^emph>\<open>chain
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   873
    production\<close>.  Chain productions act as abbreviations: conceptually,
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   874
    they are removed from the grammar by adding new productions.
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   875
    Priority information attached to chain productions is ignored; only
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   876
    the dummy value \<open>-1\<close> is displayed.
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   877
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   878
    \<^descr> \<open>print modes\<close> lists the alternative print modes
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   879
    provided by this grammar; see \secref{sec:print-modes}.
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   880
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   881
    \<^descr> \<open>parse_rules\<close> and \<open>print_rules\<close> relate to
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   882
    syntax translations (macros); see \secref{sec:syn-trans}.
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
   883
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   884
    \<^descr> \<open>parse_ast_translation\<close> and \<open>print_ast_translation\<close> list sets of constants that invoke
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   885
    translation functions for abstract syntax trees, which are only
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   886
    required in very special situations; see \secref{sec:tr-funs}.
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   887
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   888
    \<^descr> \<open>parse_translation\<close> and \<open>print_translation\<close>
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   889
    list the sets of constants that invoke regular translation
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
   890
    functions; see \secref{sec:tr-funs}.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   891
\<close>
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   892
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   893
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   894
subsection \<open>Ambiguity of parsed expressions\<close>
46291
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   895
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   896
text \<open>
46291
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   897
  \begin{tabular}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   898
    @{attribute_def syntax_ambiguity_warning} & : & \<open>attribute\<close> & default \<open>true\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   899
    @{attribute_def syntax_ambiguity_limit} & : & \<open>attribute\<close> & default \<open>10\<close> \\
46291
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   900
  \end{tabular}
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   901
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   902
  Depending on the grammar and the given input, parsing may be
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   903
  ambiguous.  Isabelle lets the Earley parser enumerate all possible
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   904
  parse trees, and then tries to make the best out of the situation.
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   905
  Terms that cannot be type-checked are filtered out, which often
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   906
  leads to a unique result in the end.  Unlike regular type
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   907
  reconstruction, which is applied to the whole collection of input
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   908
  terms simultaneously, the filtering stage only treats each given
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   909
  term in isolation.  Filtering is also not attempted for individual
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   910
  types or raw ASTs (as required for @{command translations}).
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   911
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   912
  Certain warning or error messages are printed, depending on the
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   913
  situation and the given configuration options.  Parsing ultimately
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   914
  fails, if multiple results remain after the filtering phase.
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   915
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   916
  \<^descr> @{attribute syntax_ambiguity_warning} controls output of
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   917
  explicit warning messages about syntax ambiguity.
46291
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   918
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
   919
  \<^descr> @{attribute syntax_ambiguity_limit} determines the number of
46291
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   920
  resulting parse trees that are shown as part of the printed message
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   921
  in case of an ambiguity.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   922
\<close>
46291
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   923
a1827b8b45ae updated section about syntax ambiguity;
wenzelm
parents: 46290
diff changeset
   924
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   925
section \<open>Syntax transformations \label{sec:syntax-transformations}\<close>
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   926
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   927
text \<open>The inner syntax engine of Isabelle provides separate
52413
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
   928
  mechanisms to transform parse trees either via rewrite systems on
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   929
  first-order ASTs (\secref{sec:syn-trans}), or ML functions on ASTs
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   930
  or syntactic \<open>\<lambda>\<close>-terms (\secref{sec:tr-funs}).  This works
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   931
  both for parsing and printing, as outlined in
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   932
  \figref{fig:parse-print}.
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   933
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   934
  \begin{figure}[htbp]
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   935
  \begin{center}
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   936
  \begin{tabular}{cl}
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   937
  string          & \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   938
  \<open>\<down>\<close>     & lexer + parser \\
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   939
  parse tree      & \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   940
  \<open>\<down>\<close>     & parse AST translation \\
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   941
  AST             & \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   942
  \<open>\<down>\<close>     & AST rewriting (macros) \\
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   943
  AST             & \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   944
  \<open>\<down>\<close>     & parse translation \\
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   945
  --- pre-term ---    & \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   946
  \<open>\<down>\<close>     & print translation \\
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   947
  AST             & \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   948
  \<open>\<down>\<close>     & AST rewriting (macros) \\
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   949
  AST             & \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   950
  \<open>\<down>\<close>     & print AST translation \\
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   951
  string          &
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   952
  \end{tabular}
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   953
  \end{center}
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   954
  \caption{Parsing and printing with translations}\label{fig:parse-print}
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   955
  \end{figure}
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   956
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   957
  These intermediate syntax tree formats eventually lead to a pre-term
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   958
  with all names and binding scopes resolved, but most type
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   959
  information still missing.  Explicit type constraints might be given by
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   960
  the user, or implicit position information by the system --- both
48816
wenzelm
parents: 48792
diff changeset
   961
  need to be passed-through carefully by syntax transformations.
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   962
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   963
  Pre-terms are further processed by the so-called \<^emph>\<open>check\<close> and
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
   964
  \<^emph>\<open>uncheck\<close> phases that are intertwined with type-inference (see
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58421
diff changeset
   965
  also @{cite "isabelle-implementation"}).  The latter allows to operate
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   966
  on higher-order abstract syntax with proper binding and type
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   967
  information already available.
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   968
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   969
  As a rule of thumb, anything that manipulates bindings of variables
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   970
  or constants needs to be implemented as syntax transformation (see
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   971
  below).  Anything else is better done via check/uncheck: a prominent
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   972
  example application is the @{command abbreviation} concept of
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   973
  Isabelle/Pure.\<close>
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   974
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   975
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   976
subsection \<open>Abstract syntax trees \label{sec:ast}\<close>
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
   977
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   978
text \<open>The ML datatype @{ML_type Ast.ast} explicitly represents the
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   979
  intermediate AST format that is used for syntax rewriting
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   980
  (\secref{sec:syn-trans}).  It is defined in ML as follows:
61408
9020a3ba6c9a @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61143
diff changeset
   981
  @{verbatim [display]
9020a3ba6c9a @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61143
diff changeset
   982
\<open>datatype ast =
9020a3ba6c9a @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61143
diff changeset
   983
  Constant of string |
9020a3ba6c9a @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61143
diff changeset
   984
  Variable of string |
9020a3ba6c9a @{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents: 61143
diff changeset
   985
  Appl of ast list\<close>}
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   986
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   987
  An AST is either an atom (constant or variable) or a list of (at
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   988
  least two) subtrees.  Occasional diagnostic output of ASTs uses
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   989
  notation that resembles S-expression of LISP.  Constant atoms are
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   990
  shown as quoted strings, variable atoms as non-quoted strings and
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   991
  applications as a parenthesized list of subtrees.  For example, the
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   992
  AST
58724
e5f809f52f26 more antiquotations;
wenzelm
parents: 58618
diff changeset
   993
  @{ML [display] \<open>Ast.Appl [Ast.Constant "_abs", Ast.Variable "x", Ast.Variable "t"]\<close>}
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   994
  is pretty-printed as \<^verbatim>\<open>("_abs" x t)\<close>.  Note that
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
   995
  \<^verbatim>\<open>()\<close> and \<^verbatim>\<open>(x)\<close> are excluded as ASTs, because
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   996
  they have too few subtrees.
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
   997
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   998
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
   999
  AST application is merely a pro-forma mechanism to indicate
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1000
  certain syntactic structures.  Thus \<^verbatim>\<open>(c a b)\<close> could mean
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1001
  either term application or type application, depending on the
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1002
  syntactic context.
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1003
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1004
  Nested application like \<^verbatim>\<open>(("_abs" x t) u)\<close> is also
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1005
  possible, but ASTs are definitely first-order: the syntax constant
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1006
  \<^verbatim>\<open>"_abs"\<close> does not bind the \<^verbatim>\<open>x\<close> in any way.
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1007
  Proper bindings are introduced in later stages of the term syntax,
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1008
  where \<^verbatim>\<open>("_abs" x t)\<close> becomes an @{ML Abs} node and
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1009
  occurrences of \<^verbatim>\<open>x\<close> in \<^verbatim>\<open>t\<close> are replaced by bound
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1010
  variables (represented as de-Bruijn indices).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1011
\<close>
48113
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
  1012
1c4500446ba4 more explicit section "Syntax transformations";
wenzelm
parents: 46512
diff changeset
  1013
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1014
subsubsection \<open>AST constants versus variables\<close>
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1015
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1016
text \<open>Depending on the situation --- input syntax, output syntax,
56582
f05b7d6ec592 tuned spelling;
wenzelm
parents: 56451
diff changeset
  1017
  translation patterns --- the distinction of atomic ASTs as @{ML
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1018
  Ast.Constant} versus @{ML Ast.Variable} serves slightly different
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1019
  purposes.
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1020
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1021
  Input syntax of a term such as \<open>f a b = c\<close> does not yet
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1022
  indicate the scopes of atomic entities \<open>f, a, b, c\<close>: they
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1023
  could be global constants or local variables, even bound ones
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1024
  depending on the context of the term.  @{ML Ast.Variable} leaves
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1025
  this choice still open: later syntax layers (or translation
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1026
  functions) may capture such a variable to determine its role
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1027
  specifically, to make it a constant, bound variable, free variable
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1028
  etc.  In contrast, syntax translations that introduce already known
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1029
  constants would rather do it via @{ML Ast.Constant} to prevent
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1030
  accidental re-interpretation later on.
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1031
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1032
  Output syntax turns term constants into @{ML Ast.Constant} and
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1033
  variables (free or schematic) into @{ML Ast.Variable}.  This
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1034
  information is precise when printing fully formal \<open>\<lambda>\<close>-terms.
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1035
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1036
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1037
  AST translation patterns (\secref{sec:syn-trans}) that
52413
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1038
  represent terms cannot distinguish constants and variables
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1039
  syntactically.  Explicit indication of \<open>CONST c\<close> inside the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1040
  term language is required, unless \<open>c\<close> is known as special
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
  1041
  \<^emph>\<open>syntax constant\<close> (see also @{command syntax}).  It is also
52413
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1042
  possible to use @{command syntax} declarations (without mixfix
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1043
  annotation) to enforce that certain unqualified names are always
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1044
  treated as constant within the syntax machinery.
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1045
52413
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1046
  The situation is simpler for ASTs that represent types or sorts,
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1047
  since the concrete syntax already distinguishes type variables from
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1048
  type constants (constructors).  So \<open>('a, 'b) foo\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1049
  corresponds to an AST application of some constant for \<open>foo\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1050
  and variable arguments for \<open>'a\<close> and \<open>'b\<close>.  Note that
52413
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1051
  the postfix application is merely a feature of the concrete syntax,
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1052
  while in the AST the constructor occurs in head position.\<close>
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1053
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1054
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1055
subsubsection \<open>Authentic syntax names\<close>
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1056
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1057
text \<open>Naming constant entities within ASTs is another delicate
52413
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1058
  issue.  Unqualified names are resolved in the name space tables in
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1059
  the last stage of parsing, after all translations have been applied.
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1060
  Since syntax transformations do not know about this later name
52413
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1061
  resolution, there can be surprises in boundary cases.
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1062
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
  1063
  \<^emph>\<open>Authentic syntax names\<close> for @{ML Ast.Constant} avoid this
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1064
  problem: the fully-qualified constant name with a special prefix for
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1065
  its formal category (\<open>class\<close>, \<open>type\<close>, \<open>const\<close>, \<open>fixed\<close>) represents the information faithfully
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1066
  within the untyped AST format.  Accidental overlap with free or
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1067
  bound variables is excluded as well.  Authentic syntax names work
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1068
  implicitly in the following situations:
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1069
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1070
  \<^item> Input of term constants (or fixed variables) that are
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1071
  introduced by concrete syntax via @{command notation}: the
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1072
  correspondence of a particular grammar production to some known term
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1073
  entity is preserved.
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1074
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1075
  \<^item> Input of type constants (constructors) and type classes ---
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1076
  thanks to explicit syntactic distinction independently on the
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1077
  context.
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1078
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1079
  \<^item> Output of term constants, type constants, type classes ---
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1080
  this information is already available from the internal term to be
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1081
  printed.
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1082
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1083
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1084
  In other words, syntax transformations that operate on input terms
48816
wenzelm
parents: 48792
diff changeset
  1085
  written as prefix applications are difficult to make robust.
wenzelm
parents: 48792
diff changeset
  1086
  Luckily, this case rarely occurs in practice, because syntax forms
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1087
  to be translated usually correspond to some concrete notation.\<close>
48114
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1088
428e55887f24 more on "Abstract syntax trees";
wenzelm
parents: 48113
diff changeset
  1089
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1090
subsection \<open>Raw syntax and translations \label{sec:syn-trans}\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1091
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1092
text \<open>
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1093
  \begin{tabular}{rcll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1094
    @{command_def "nonterminal"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1095
    @{command_def "syntax"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1096
    @{command_def "no_syntax"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1097
    @{command_def "translations"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1098
    @{command_def "no_translations"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1099
    @{attribute_def syntax_ast_trace} & : & \<open>attribute\<close> & default \<open>false\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1100
    @{attribute_def syntax_ast_stats} & : & \<open>attribute\<close> & default \<open>false\<close> \\
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1101
  \end{tabular}
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1102
  \<^medskip>
59783
00b62aa9f430 tuned syntax diagrams -- no duplication of "target";
wenzelm
parents: 58842
diff changeset
  1103
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1104
  Unlike mixfix notation for existing formal entities
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1105
  (\secref{sec:notation}), raw syntax declarations provide full access
48115
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1106
  to the priority grammar of the inner syntax, without any sanity
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1107
  checks.  This includes additional syntactic categories (via
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1108
  @{command nonterminal}) and free-form grammar productions (via
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1109
  @{command syntax}).  Additional syntax translations (or macros, via
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1110
  @{command translations}) are required to turn resulting parse trees
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1111
  into proper representations of formal entities again.
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1112
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55108
diff changeset
  1113
  @{rail \<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
  1114
    @@{command nonterminal} (@{syntax name} + @'and')
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1115
    ;
46494
ea2ae63336f3 clarified outer syntax "constdecl", which is only local to some rail diagrams;
wenzelm
parents: 46483
diff changeset
  1116
    (@@{command syntax} | @@{command no_syntax}) @{syntax mode}? (constdecl +)
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1117
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
  1118
    (@@{command translations} | @@{command no_translations})
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
  1119
      (transpat ('==' | '=>' | '<=' | '\<rightleftharpoons>' | '\<rightharpoonup>' | '\<leftharpoondown>') transpat +)
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1120
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1121
46494
ea2ae63336f3 clarified outer syntax "constdecl", which is only local to some rail diagrams;
wenzelm
parents: 46483
diff changeset
  1122
    constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}?
ea2ae63336f3 clarified outer syntax "constdecl", which is only local to some rail diagrams;
wenzelm
parents: 46483
diff changeset
  1123
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
  1124
    mode: ('(' ( @{syntax name} | @'output' | @{syntax name} @'output' ) ')')
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1125
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
  1126
    transpat: ('(' @{syntax nameref} ')')? @{syntax string}
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55108
diff changeset
  1127
  \<close>}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1128
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1129
  \<^descr> @{command "nonterminal"}~\<open>c\<close> declares a type
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1130
  constructor \<open>c\<close> (without arguments) to act as purely syntactic
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1131
  type: a nonterminal symbol of the inner syntax.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1132
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1133
  \<^descr> @{command "syntax"}~\<open>(mode) c :: \<sigma> (mx)\<close> augments the
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1134
  priority grammar and the pretty printer table for the given print
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1135
  mode (default \<^verbatim>\<open>""\<close>). An optional keyword @{keyword_ref
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1136
  "output"} means that only the pretty printer table is affected.
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1137
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1138
  Following \secref{sec:mixfix}, the mixfix annotation \<open>mx =
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1139
  template ps q\<close> together with type \<open>\<sigma> = \<tau>\<^sub>1 \<Rightarrow> \<dots> \<tau>\<^sub>n \<Rightarrow> \<tau>\<close> and
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1140
  specify a grammar production.  The \<open>template\<close> contains
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1141
  delimiter tokens that surround \<open>n\<close> argument positions
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1142
  (\<^verbatim>\<open>_\<close>).  The latter correspond to nonterminal symbols
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1143
  \<open>A\<^sub>i\<close> derived from the argument types \<open>\<tau>\<^sub>i\<close> as
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1144
  follows:
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1145
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1146
    \<^item> \<open>prop\<close> if \<open>\<tau>\<^sub>i = prop\<close>
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1147
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1148
    \<^item> \<open>logic\<close> if \<open>\<tau>\<^sub>i = (\<dots>)\<kappa>\<close> for logical type
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1149
    constructor \<open>\<kappa> \<noteq> prop\<close>
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1150
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1151
    \<^item> \<open>any\<close> if \<open>\<tau>\<^sub>i = \<alpha>\<close> for type variables
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1152
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1153
    \<^item> \<open>\<kappa>\<close> if \<open>\<tau>\<^sub>i = \<kappa>\<close> for nonterminal \<open>\<kappa>\<close>
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
  1154
    (syntactic type constructor)
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1155
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1156
  Each \<open>A\<^sub>i\<close> is decorated by priority \<open>p\<^sub>i\<close> from the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1157
  given list \<open>ps\<close>; missing priorities default to 0.
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1158
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1159
  The resulting nonterminal of the production is determined similarly
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1160
  from type \<open>\<tau>\<close>, with priority \<open>q\<close> and default 1000.
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1161
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1162
  \<^medskip>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1163
  Parsing via this production produces parse trees \<open>t\<^sub>1, \<dots>, t\<^sub>n\<close> for the argument slots.  The resulting parse tree is
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1164
  composed as \<open>c t\<^sub>1 \<dots> t\<^sub>n\<close>, by using the syntax constant \<open>c\<close> of the syntax declaration.
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1165
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1166
  Such syntactic constants are invented on the spot, without formal
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1167
  check wrt.\ existing declarations.  It is conventional to use plain
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1168
  identifiers prefixed by a single underscore (e.g.\ \<open>_foobar\<close>).  Names should be chosen with care, to avoid clashes
48816
wenzelm
parents: 48792
diff changeset
  1169
  with other syntax declarations.
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1170
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1171
  \<^medskip>
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1172
  The special case of copy production is specified by \<open>c =\<close>~\<^verbatim>\<open>""\<close> (empty string).
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1173
  It means that the resulting parse tree \<open>t\<close> is copied directly, without any
46292
4eb48958b50f updated section on raw syntax;
wenzelm
parents: 46291
diff changeset
  1174
  further decoration.
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
  1175
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1176
  \<^descr> @{command "no_syntax"}~\<open>(mode) decls\<close> removes grammar
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1177
  declarations (and translations) resulting from \<open>decls\<close>, which
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1178
  are interpreted in the same manner as for @{command "syntax"} above.
46282
83864b045a72 clarified syntax section structure;
wenzelm
parents: 45703
diff changeset
  1179
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1180
  \<^descr> @{command "translations"}~\<open>rules\<close> specifies syntactic
48115
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1181
  translation rules (i.e.\ macros) as first-order rewrite rules on
48816
wenzelm
parents: 48792
diff changeset
  1182
  ASTs (\secref{sec:ast}).  The theory context maintains two
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1183
  independent lists translation rules: parse rules (\<^verbatim>\<open>=>\<close>
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1184
  or \<open>\<rightharpoonup>\<close>) and print rules (\<^verbatim>\<open><=\<close> or \<open>\<leftharpoondown>\<close>).
48115
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1185
  For convenience, both can be specified simultaneously as parse~/
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1186
  print rules (\<^verbatim>\<open>==\<close> or \<open>\<rightleftharpoons>\<close>).
48115
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1187
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1188
  Translation patterns may be prefixed by the syntactic category to be
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1189
  used for parsing; the default is \<open>logic\<close> which means that
48115
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1190
  regular term syntax is used.  Both sides of the syntax translation
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1191
  rule undergo parsing and parse AST translations
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1192
  \secref{sec:tr-funs}, in order to perform some fundamental
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1193
  normalization like \<open>\<lambda>x y. b \<leadsto> \<lambda>x. \<lambda>y. b\<close>, but other AST
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
  1194
  translation rules are \<^emph>\<open>not\<close> applied recursively here.
48115
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1195
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1196
  When processing AST patterns, the inner syntax lexer runs in a
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1197
  different mode that allows identifiers to start with underscore.
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1198
  This accommodates the usual naming convention for auxiliary syntax
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1199
  constants --- those that do not have a logical counter part --- by
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1200
  allowing to specify arbitrary AST applications within the term
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1201
  syntax, independently of the corresponding concrete syntax.
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1202
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1203
  Atomic ASTs are distinguished as @{ML Ast.Constant} versus @{ML
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1204
  Ast.Variable} as follows: a qualified name or syntax constant
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1205
  declared via @{command syntax}, or parse tree head of concrete
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1206
  notation becomes @{ML Ast.Constant}, anything else @{ML
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1207
  Ast.Variable}.  Note that \<open>CONST\<close> and \<open>XCONST\<close> within
48115
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1208
  the term language (\secref{sec:pure-grammar}) allow to enforce
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1209
  treatment as constants.
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1210
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1211
  AST rewrite rules \<open>(lhs, rhs)\<close> need to obey the following
48115
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1212
  side-conditions:
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1213
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1214
    \<^item> Rules must be left linear: \<open>lhs\<close> must not contain
61572
ddb3ac3fef45 more antiquotations;
wenzelm
parents: 61503
diff changeset
  1215
    repeated variables.\<^footnote>\<open>The deeper reason for this is that AST
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
  1216
    equality is not well-defined: different occurrences of the ``same''
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
  1217
    AST could be decorated differently by accidental type-constraints or
61572
ddb3ac3fef45 more antiquotations;
wenzelm
parents: 61503
diff changeset
  1218
    source position information, for example.\<close>
48115
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1219
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1220
    \<^item> Every variable in \<open>rhs\<close> must also occur in \<open>lhs\<close>.
48115
d868e4f7905b more on syntax translations;
wenzelm
parents: 48114
diff changeset
  1221
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1222
  \<^descr> @{command "no_translations"}~\<open>rules\<close> removes syntactic
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1223
  translation rules, which are interpreted in the same manner as for
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1224
  @{command "translations"} above.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1225
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
  1226
  \<^descr> @{attribute syntax_ast_trace} and @{attribute
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1227
  syntax_ast_stats} control diagnostic output in the AST normalization
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1228
  process, when translation rules are applied to concrete input or
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1229
  output.
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1230
46293
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1231
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1232
  Raw syntax and translations provides a slightly more low-level
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1233
  access to the grammar and the form of resulting parse trees.  It is
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1234
  often possible to avoid this untyped macro mechanism, and use
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1235
  type-safe @{command abbreviation} or @{command notation} instead.
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1236
  Some important situations where @{command syntax} and @{command
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1237
  translations} are really need are as follows:
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1238
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1239
  \<^item> Iterated replacement via recursive @{command translations}.
46293
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1240
  For example, consider list enumeration @{term "[a, b, c, d]"} as
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1241
  defined in theory @{theory List} in Isabelle/HOL.
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1242
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1243
  \<^item> Change of binding status of variables: anything beyond the
46293
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1244
  built-in @{keyword "binder"} mixfix annotation requires explicit
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1245
  syntax translations.  For example, consider list filter
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1246
  comprehension @{term "[x \<leftarrow> xs . P]"} as defined in theory @{theory
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1247
  List} in Isabelle/HOL.
61458
987533262fc2 Markdown support in document text;
wenzelm
parents: 61439
diff changeset
  1248
\<close>
46293
f248b5f2783a updated examples for syntax translations;
wenzelm
parents: 46292
diff changeset
  1249
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1250
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1251
subsubsection \<open>Applying translation rules\<close>
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1252
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1253
text \<open>As a term is being parsed or printed, an AST is generated as
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1254
  an intermediate form according to \figref{fig:parse-print}.  The AST
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1255
  is normalized by applying translation rules in the manner of a
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1256
  first-order term rewriting system.  We first examine how a single
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1257
  rule is applied.
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1258
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1259
  Let \<open>t\<close> be the abstract syntax tree to be normalized and
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1260
  \<open>(lhs, rhs)\<close> some translation rule.  A subtree \<open>u\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1261
  of \<open>t\<close> is called \<^emph>\<open>redex\<close> if it is an instance of \<open>lhs\<close>; in this case the pattern \<open>lhs\<close> is said to match the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1262
  object \<open>u\<close>.  A redex matched by \<open>lhs\<close> may be
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1263
  replaced by the corresponding instance of \<open>rhs\<close>, thus
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1264
  \<^emph>\<open>rewriting\<close> the AST \<open>t\<close>.  Matching requires some notion
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
  1265
  of \<^emph>\<open>place-holders\<close> in rule patterns: @{ML Ast.Variable} serves
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1266
  this purpose.
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1267
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1268
  More precisely, the matching of the object \<open>u\<close> against the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1269
  pattern \<open>lhs\<close> is performed as follows:
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1270
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1271
  \<^item> Objects of the form @{ML Ast.Variable}~\<open>x\<close> or @{ML
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1272
  Ast.Constant}~\<open>x\<close> are matched by pattern @{ML
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1273
  Ast.Constant}~\<open>x\<close>.  Thus all atomic ASTs in the object are
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1274
  treated as (potential) constants, and a successful match makes them
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1275
  actual constants even before name space resolution (see also
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1276
  \secref{sec:ast}).
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1277
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1278
  \<^item> Object \<open>u\<close> is matched by pattern @{ML
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1279
  Ast.Variable}~\<open>x\<close>, binding \<open>x\<close> to \<open>u\<close>.
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1280
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1281
  \<^item> Object @{ML Ast.Appl}~\<open>us\<close> is matched by @{ML
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1282
  Ast.Appl}~\<open>ts\<close> if \<open>us\<close> and \<open>ts\<close> have the
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1283
  same length and each corresponding subtree matches.
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1284
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1285
  \<^item> In every other case, matching fails.
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1286
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1287
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1288
  A successful match yields a substitution that is applied to \<open>rhs\<close>, generating the instance that replaces \<open>u\<close>.
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1289
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1290
  Normalizing an AST involves repeatedly applying translation rules
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1291
  until none are applicable.  This works yoyo-like: top-down,
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1292
  bottom-up, top-down, etc.  At each subtree position, rules are
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1293
  chosen in order of appearance in the theory definitions.
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1294
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1295
  The configuration options @{attribute syntax_ast_trace} and
48816
wenzelm
parents: 48792
diff changeset
  1296
  @{attribute syntax_ast_stats} might help to understand this process
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1297
  and diagnose problems.
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1298
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1299
  \begin{warn}
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1300
  If syntax translation rules work incorrectly, the output of
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
  1301
  @{command_ref print_syntax} with its \<^emph>\<open>rules\<close> sections reveals the
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1302
  actual internal forms of AST pattern, without potentially confusing
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1303
  concrete syntax.  Recall that AST constants appear as quoted strings
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1304
  and variables without quotes.
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1305
  \end{warn}
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1306
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1307
  \begin{warn}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1308
  If @{attribute_ref eta_contract} is set to \<open>true\<close>, terms
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1309
  will be \<open>\<eta>\<close>-contracted \<^emph>\<open>before\<close> the AST rewriter sees
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1310
  them.  Thus some abstraction nodes needed for print rules to match
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1311
  may vanish.  For example, \<open>Ball A (\<lambda>x. P x)\<close> would contract
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1312
  to \<open>Ball A P\<close> and the standard print rule would fail to
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1313
  apply.  This problem can be avoided by hand-written ML translation
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1314
  functions (see also \secref{sec:tr-funs}), which is in fact the same
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1315
  mechanism used in built-in @{keyword "binder"} declarations.
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1316
  \end{warn}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1317
\<close>
48117
e21f4d5b9636 more on "Applying translation rules";
wenzelm
parents: 48115
diff changeset
  1318
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1319
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1320
subsection \<open>Syntax translation functions \label{sec:tr-funs}\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1321
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1322
text \<open>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1323
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1324
    @{command_def "parse_ast_translation"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1325
    @{command_def "parse_translation"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1326
    @{command_def "print_translation"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1327
    @{command_def "typed_print_translation"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1328
    @{command_def "print_ast_translation"} & : & \<open>theory \<rightarrow> theory\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1329
    @{ML_antiquotation_def "class_syntax"} & : & \<open>ML antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1330
    @{ML_antiquotation_def "type_syntax"} & : & \<open>ML antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1331
    @{ML_antiquotation_def "const_syntax"} & : & \<open>ML antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1332
    @{ML_antiquotation_def "syntax_const"} & : & \<open>ML antiquotation\<close> \\
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1333
  \end{matharray}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1334
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1335
  Syntax translation functions written in ML admit almost arbitrary
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1336
  manipulations of inner syntax, at the expense of some complexity and
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1337
  obscurity in the implementation.
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1338
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55108
diff changeset
  1339
  @{rail \<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
  1340
  ( @@{command parse_ast_translation} | @@{command parse_translation} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 42358
diff changeset
  1341
    @@{command print_translation} | @@{command typed_print_translation} |
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1342
    @@{command print_ast_translation}) @{syntax text}
48119
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1343
  ;
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1344
  (@@{ML_antiquotation class_syntax} |
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1345
   @@{ML_antiquotation type_syntax} |
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1346
   @@{ML_antiquotation const_syntax} |
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1347
   @@{ML_antiquotation syntax_const}) name
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55108
diff changeset
  1348
  \<close>}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1349
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
  1350
  \<^descr> @{command parse_translation} etc. declare syntax translation
48119
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1351
  functions to the theory.  Any of these commands have a single
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1352
  @{syntax text} argument that refers to an ML expression of
52413
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1353
  appropriate type as follows:
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1354
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1355
  \<^medskip>
48119
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1356
  {\footnotesize
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1357
  \begin{tabular}{l}
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1358
  @{command parse_ast_translation} : \\
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1359
  \quad @{ML_type "(string * (Proof.context -> Ast.ast list -> Ast.ast)) list"} \\
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1360
  @{command parse_translation} : \\
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1361
  \quad @{ML_type "(string * (Proof.context -> term list -> term)) list"} \\
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1362
  @{command print_translation} : \\
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1363
  \quad @{ML_type "(string * (Proof.context -> term list -> term)) list"} \\
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1364
  @{command typed_print_translation} : \\
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1365
  \quad @{ML_type "(string * (Proof.context -> typ -> term list -> term)) list"} \\
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1366
  @{command print_ast_translation} : \\
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51960
diff changeset
  1367
  \quad @{ML_type "(string * (Proof.context -> Ast.ast list -> Ast.ast)) list"} \\
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1368
  \end{tabular}}
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1369
  \<^medskip>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1370
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1371
  The argument list consists of \<open>(c, tr)\<close> pairs, where \<open>c\<close> is the syntax name of the formal entity involved, and \<open>tr\<close> a function that translates a syntax form \<open>c args\<close> into
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1372
  \<open>tr ctxt args\<close> (depending on the context).  The Isabelle/ML
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1373
  naming convention for parse translations is \<open>c_tr\<close> and for
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1374
  print translations \<open>c_tr'\<close>.
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1375
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1376
  The @{command_ref print_syntax} command displays the sets of names
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1377
  associated with the translation functions of a theory under \<open>parse_ast_translation\<close> etc.
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1378
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1379
  \<^descr> \<open>@{class_syntax c}\<close>, \<open>@{type_syntax c}\<close>,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1380
  \<open>@{const_syntax c}\<close> inline the authentic syntax name of the
48119
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1381
  given formal entities into the ML source.  This is the
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1382
  fully-qualified logical name prefixed by a special marker to
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1383
  indicate its kind: thus different logical name spaces are properly
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1384
  distinguished within parse trees.
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1385
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1386
  \<^descr> \<open>@{const_syntax c}\<close> inlines the name \<open>c\<close> of
48119
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1387
  the given syntax constant, having checked that it has been declared
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1388
  via some @{command syntax} commands within the theory context.  Note
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1389
  that the usual naming convention makes syntax constants start with
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1390
  underscore, to reduce the chance of accidental clashes with other
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1391
  names occurring in parse trees (unqualified constants etc.).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1392
\<close>
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1393
48119
55c305e29f4b cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
wenzelm
parents: 48118
diff changeset
  1394
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1395
subsubsection \<open>The translation strategy\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1396
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1397
text \<open>The different kinds of translation functions are invoked during
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1398
  the transformations between parse trees, ASTs and syntactic terms
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1399
  (cf.\ \figref{fig:parse-print}).  Whenever a combination of the form
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1400
  \<open>c x\<^sub>1 \<dots> x\<^sub>n\<close> is encountered, and a translation function
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1401
  \<open>f\<close> of appropriate kind is declared for \<open>c\<close>, the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1402
  result is produced by evaluation of \<open>f [x\<^sub>1, \<dots>, x\<^sub>n]\<close> in ML.
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1403
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1404
  For AST translations, the arguments \<open>x\<^sub>1, \<dots>, x\<^sub>n\<close> are ASTs.  A
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1405
  combination has the form @{ML "Ast.Constant"}~\<open>c\<close> or @{ML
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1406
  "Ast.Appl"}~\<open>[\<close>@{ML Ast.Constant}~\<open>c, x\<^sub>1, \<dots>, x\<^sub>n]\<close>.
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1407
  For term translations, the arguments are terms and a combination has
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1408
  the form @{ML Const}~\<open>(c, \<tau>)\<close> or @{ML Const}~\<open>(c, \<tau>)
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1409
  $ x\<^sub>1 $ \<dots> $ x\<^sub>n\<close>.  Terms allow more sophisticated transformations
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1410
  than ASTs do, typically involving abstractions and bound
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
  1411
  variables. \<^emph>\<open>Typed\<close> print translations may even peek at the type
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1412
  \<open>\<tau>\<close> of the constant they are invoked on, although some
52413
a59ba6de9687 misc tuning and clarification;
wenzelm
parents: 52143
diff changeset
  1413
  information might have been suppressed for term output already.
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1414
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1415
  Regardless of whether they act on ASTs or terms, translation
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1416
  functions called during the parsing process differ from those for
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1417
  printing in their overall behaviour:
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1418
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
  1419
  \<^descr>[Parse translations] are applied bottom-up.  The arguments are
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1420
  already in translated form.  The translations must not fail;
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1421
  exceptions trigger an error message.  There may be at most one
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1422
  function associated with any syntactic name.
46294
wenzelm
parents: 46293
diff changeset
  1423
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61421
diff changeset
  1424
  \<^descr>[Print translations] are applied top-down.  They are supplied
48118
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1425
  with arguments that are partly still in internal form.  The result
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1426
  again undergoes translation; therefore a print translation should
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1427
  not introduce as head the very constant that invoked it.  The
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1428
  function may raise exception @{ML Match} to indicate failure; in
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1429
  this event it has no effect.  Multiple functions associated with
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1430
  some syntactic name are tried in the order of declaration in the
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1431
  theory.
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1432
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1433
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1434
  Only constant atoms --- constructor @{ML Ast.Constant} for ASTs and
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1435
  @{ML Const} for terms --- can invoke translation functions.  This
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1436
  means that parse translations can only be associated with parse tree
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1437
  heads of concrete syntax, or syntactic constants introduced via
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1438
  other translations.  For plain identifiers within the term language,
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1439
  the status of constant versus variable is not yet know during
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1440
  parsing.  This is in contrast to print translations, where constants
8537313dd261 more on "Syntax translation functions";
wenzelm
parents: 48117
diff changeset
  1441
  are explicitly known from the given term in its fully internal form.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1442
\<close>
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1443
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1444
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1445
subsection \<open>Built-in syntax transformations\<close>
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1446
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1447
text \<open>
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1448
  Here are some further details of the main syntax transformation
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1449
  phases of \figref{fig:parse-print}.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1450
\<close>
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1451
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1452
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1453
subsubsection \<open>Transforming parse trees to ASTs\<close>
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1454
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1455
text \<open>The parse tree is the raw output of the parser.  It is
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1456
  transformed into an AST according to some basic scheme that may be
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1457
  augmented by AST translation functions as explained in
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1458
  \secref{sec:tr-funs}.
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1459
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1460
  The parse tree is constructed by nesting the right-hand sides of the
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1461
  productions used to recognize the input.  Such parse trees are
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1462
  simply lists of tokens and constituent parse trees, the latter
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1463
  representing the nonterminals of the productions.  Ignoring AST
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1464
  translation functions, parse trees are transformed to ASTs by
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1465
  stripping out delimiters and copy productions, while retaining some
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1466
  source position information from input tokens.
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1467
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1468
  The Pure syntax provides predefined AST translations to make the
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1469
  basic \<open>\<lambda>\<close>-term structure more apparent within the
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1470
  (first-order) AST representation, and thus facilitate the use of
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1471
  @{command translations} (see also \secref{sec:syn-trans}).  This
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1472
  covers ordinary term application, type application, nested
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1473
  abstraction, iterated meta implications and function types.  The
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1474
  effect is illustrated on some representative input strings is as
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1475
  follows:
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1476
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1477
  \begin{center}
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1478
  \begin{tabular}{ll}
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1479
  input source & AST \\
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1480
  \hline
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1481
  \<open>f x y z\<close> & \<^verbatim>\<open>(f x y z)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1482
  \<open>'a ty\<close> & \<^verbatim>\<open>(ty 'a)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1483
  \<open>('a, 'b)ty\<close> & \<^verbatim>\<open>(ty 'a 'b)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1484
  \<open>\<lambda>x y z. t\<close> & \<^verbatim>\<open>("_abs" x ("_abs" y ("_abs" z t)))\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1485
  \<open>\<lambda>x :: 'a. t\<close> & \<^verbatim>\<open>("_abs" ("_constrain" x 'a) t)\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1486
  \<open>\<lbrakk>P; Q; R\<rbrakk> \<Longrightarrow> S\<close> & \<^verbatim>\<open>("Pure.imp" P ("Pure.imp" Q ("Pure.imp" R S)))\<close> \\
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1487
   \<open>['a, 'b, 'c] \<Rightarrow> 'd\<close> & \<^verbatim>\<open>("fun" 'a ("fun" 'b ("fun" 'c 'd)))\<close> \\
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1488
  \end{tabular}
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1489
  \end{center}
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1490
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1491
  Note that type and sort constraints may occur in further places ---
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1492
  translations need to be ready to cope with them.  The built-in
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1493
  syntax transformation from parse trees to ASTs insert additional
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1494
  constraints that represent source positions.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1495
\<close>
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1496
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1497
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1498
subsubsection \<open>Transforming ASTs to terms\<close>
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1499
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1500
text \<open>After application of macros (\secref{sec:syn-trans}), the AST
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1501
  is transformed into a term.  This term still lacks proper type
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1502
  information, but it might contain some constraints consisting of
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1503
  applications with head \<^verbatim>\<open>_constrain\<close>, where the second
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1504
  argument is a type encoded as a pre-term within the syntax.  Type
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1505
  inference later introduces correct types, or indicates type errors
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1506
  in the input.
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1507
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1508
  Ignoring parse translations, ASTs are transformed to terms by
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1509
  mapping AST constants to term constants, AST variables to term
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1510
  variables or constants (according to the name space), and AST
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1511
  applications to iterated term applications.
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1512
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1513
  The outcome is still a first-order term.  Proper abstractions and
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1514
  bound variables are introduced by parse translations associated with
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1515
  certain syntax constants.  Thus \<^verbatim>\<open>("_abs" x x)\<close> eventually
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1516
  becomes a de-Bruijn term \<^verbatim>\<open>Abs ("x", _, Bound 0)\<close>.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1517
\<close>
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1518
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1519
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1520
subsubsection \<open>Printing of terms\<close>
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1521
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1522
text \<open>The output phase is essentially the inverse of the input
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1523
  phase.  Terms are translated via abstract syntax trees into
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1524
  pretty-printed text.
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1525
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1526
  Ignoring print translations, the transformation maps term constants,
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1527
  variables and applications to the corresponding constructs on ASTs.
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1528
  Abstractions are mapped to applications of the special constant
61503
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1529
  \<^verbatim>\<open>_abs\<close> as seen before.  Type constraints are represented
28e788ca2c5d more control symbols;
wenzelm
parents: 61494
diff changeset
  1530
  via special \<^verbatim>\<open>_constrain\<close> forms, according to various
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1531
  policies of type annotation determined elsewhere.  Sort constraints
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1532
  of type variables are handled in a similar fashion.
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1533
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1534
  After application of macros (\secref{sec:syn-trans}), the AST is
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1535
  finally pretty-printed.  The built-in print AST translations reverse
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1536
  the corresponding parse AST translations.
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1537
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1538
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1539
  For the actual printing process, the priority grammar
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1540
  (\secref{sec:priority-grammar}) plays a vital role: productions are
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1541
  used as templates for pretty printing, with argument slots stemming
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1542
  from nonterminals, and syntactic sugar stemming from literal tokens.
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1543
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1544
  Each AST application with constant head \<open>c\<close> and arguments
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1545
  \<open>t\<^sub>1\<close>, \dots, \<open>t\<^sub>n\<close> (for \<open>n = 0\<close> the AST is
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1546
  just the constant \<open>c\<close> itself) is printed according to the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1547
  first grammar production of result name \<open>c\<close>.  The required
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1548
  syntax priority of the argument slot is given by its nonterminal
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1549
  \<open>A\<^sup>(\<^sup>p\<^sup>)\<close>.  The argument \<open>t\<^sub>i\<close> that corresponds to the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1550
  position of \<open>A\<^sup>(\<^sup>p\<^sup>)\<close> is printed recursively, and then put in
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1551
  parentheses \<^emph>\<open>if\<close> its priority \<open>p\<close> requires this.  The
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1552
  resulting output is concatenated with the syntactic sugar according
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1553
  to the grammar production.
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1554
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1555
  If an AST application \<open>(c x\<^sub>1 \<dots> x\<^sub>m)\<close> has more arguments than
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1556
  the corresponding production, it is first split into \<open>((c x\<^sub>1
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1557
  \<dots> x\<^sub>n) x\<^sub>n\<^sub>+\<^sub>1 \<dots> x\<^sub>m)\<close> and then printed recursively as above.
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1558
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1559
  Applications with too few arguments or with non-constant head or
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1560
  without a corresponding production are printed in prefix-form like
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1561
  \<open>f t\<^sub>1 \<dots> t\<^sub>n\<close> for terms.
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1562
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1563
  Multiple productions associated with some name \<open>c\<close> are tried
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1564
  in order of appearance within the grammar.  An occurrence of some
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1565
  AST variable \<open>x\<close> is printed as \<open>x\<close> outright.
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1566
61421
e0825405d398 more symbols;
wenzelm
parents: 61408
diff changeset
  1567
  \<^medskip>
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61459
diff changeset
  1568
  White space is \<^emph>\<open>not\<close> inserted automatically.  If
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1569
  blanks (or breaks) are required to separate tokens, they need to be
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1570
  specified in the mixfix declaration (\secref{sec:mixfix}).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1571
\<close>
52414
8429123bc58a more on built-in syntax transformations, based on reduced version of old material;
wenzelm
parents: 52413
diff changeset
  1572
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
  1573
end