doc-src/IsarRef/Thy/Inner_Syntax.thy
author blanchet
Mon, 19 Apr 2010 18:14:45 +0200
changeset 36230 43d10a494c91
parent 35413 4c7cba1f7ce9
child 36508 03d2a2d0ee4a
permissions -rw-r--r--
added warning about inconsistent context to Metis; it makes more sense here than in Sledgehammer, because Sledgehammer is unsound and there's no point in having people panicking about the consistency of their context when their context is in fact consistent
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     1
theory Inner_Syntax
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     2
imports Main
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     3
begin
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     4
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
     5
chapter {* Inner syntax --- the term language \label{ch:inner-syntax} *}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     6
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     7
section {* Printing logical entities *}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     8
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
     9
subsection {* Diagnostic commands *}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    10
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    11
text {*
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    12
  \begin{matharray}{rcl}
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    13
    @{command_def "typ"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    14
    @{command_def "term"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    15
    @{command_def "prop"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    16
    @{command_def "thm"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    17
    @{command_def "prf"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    18
    @{command_def "full_prf"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    19
    @{command_def "pr"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    20
  \end{matharray}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    21
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    22
  These diagnostic commands assist interactive development by printing
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    23
  internal logical entities in a human-readable fashion.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    24
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    25
  \begin{rail}
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    26
    'typ' modes? type
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    27
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    28
    'term' modes? term
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    29
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    30
    'prop' modes? prop
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    31
    ;
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    32
    'thm' modes? thmrefs
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    33
    ;
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    34
    ( 'prf' | 'full\_prf' ) modes? thmrefs?
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    35
    ;
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    36
    'pr' modes? nat? (',' nat)?
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    37
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    38
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    39
    modes: '(' (name + ) ')'
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    40
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    41
  \end{rail}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    42
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    43
  \begin{description}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    44
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    45
  \item @{command "typ"}~@{text \<tau>} reads and prints types of the
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    46
  meta-logic according to the current theory or proof context.
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    47
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    48
  \item @{command "term"}~@{text t} and @{command "prop"}~@{text \<phi>}
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    49
  read, type-check and print terms or propositions according to the
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    50
  current theory or proof context; the inferred type of @{text t} is
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    51
  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
    52
  inspecting the current environment of term abbreviations.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    53
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    54
  \item @{command "thm"}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} retrieves
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    55
  theorems from the current theory or proof context.  Note that any
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    56
  attributes included in the theorem specifications are applied to a
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    57
  temporary context derived from the current theory or proof; the
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    58
  result is discarded, i.e.\ attributes involved in @{text "a\<^sub>1,
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    59
  \<dots>, a\<^sub>n"} do not have any permanent effect.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    60
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    61
  \item @{command "prf"} displays the (compact) proof term of the
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    62
  current proof state (if present), or of the given theorems. Note
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    63
  that this requires proof terms to be switched on for the current
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    64
  object logic (see the ``Proof terms'' section of the Isabelle
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    65
  reference manual for information on how to do this).
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    66
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    67
  \item @{command "full_prf"} is like @{command "prf"}, but displays
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    68
  the full proof term, i.e.\ also displays information omitted in the
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    69
  compact proof term, which is denoted by ``@{text _}'' placeholders
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    70
  there.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    71
28766
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    72
  \item @{command "pr"}~@{text "goals, prems"} prints the current
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    73
  proof state (if present), including the proof context, current facts
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    74
  and goals.  The optional limit arguments affect the number of goals
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    75
  and premises to be displayed, which is initially 10 for both.
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    76
  Omitting limit values leaves the current setting unchanged.
accab7594b8e misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28765
diff changeset
    77
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    78
  \end{description}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    79
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    80
  All of the diagnostic commands above admit a list of @{text modes}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    81
  to be specified, which is appended to the current print mode (see
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    82
  also \cite{isabelle-ref}).  Thus the output behavior may be modified
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    83
  according particular print mode features.  For example, @{command
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    84
  "pr"}~@{text "(latex xsymbols)"} would print the current proof state
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    85
  with mathematical symbols and special characters represented in
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    86
  {\LaTeX} source, according to the Isabelle style
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    87
  \cite{isabelle-sys}.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    88
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    89
  Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    90
  systematic way to include formal items into the printed text
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    91
  document.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    92
*}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    93
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
    94
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
    95
subsection {* Details of printed content *}
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
    96
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
    97
text {*
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
    98
  \begin{mldecls} 
32833
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
    99
    @{index_ML show_types: "bool Unsynchronized.ref"} & default @{ML false} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   100
    @{index_ML show_sorts: "bool Unsynchronized.ref"} & default @{ML false} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   101
    @{index_ML show_consts: "bool Unsynchronized.ref"} & default @{ML false} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   102
    @{index_ML long_names: "bool Unsynchronized.ref"} & default @{ML false} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   103
    @{index_ML short_names: "bool Unsynchronized.ref"} & default @{ML false} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   104
    @{index_ML unique_names: "bool Unsynchronized.ref"} & default @{ML true} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   105
    @{index_ML show_brackets: "bool Unsynchronized.ref"} & default @{ML false} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   106
    @{index_ML eta_contract: "bool Unsynchronized.ref"} & default @{ML true} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   107
    @{index_ML goals_limit: "int Unsynchronized.ref"} & default @{ML 10} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   108
    @{index_ML Proof.show_main_goal: "bool Unsynchronized.ref"} & default @{ML false} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   109
    @{index_ML show_hyps: "bool Unsynchronized.ref"} & default @{ML false} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   110
    @{index_ML show_tags: "bool Unsynchronized.ref"} & default @{ML false} \\
f3716d1a2e48 explicitly Unsynchronized;
wenzelm
parents: 30397
diff changeset
   111
    @{index_ML show_question_marks: "bool Unsynchronized.ref"} & default @{ML true} \\
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   112
  \end{mldecls}
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   113
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   114
  These global ML variables control the detail of information that is
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   115
  displayed for types, terms, theorems, goals etc.
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   116
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   117
  In interactive sessions, the user interface usually manages these
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   118
  global parameters of the Isabelle process, even with some concept of
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   119
  persistence.  Nonetheless it is occasionally useful to manipulate ML
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   120
  variables directly, e.g.\ using @{command "ML_val"} or @{command
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   121
  "ML_command"}.
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   122
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   123
  Batch-mode logic sessions may be configured by putting appropriate
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   124
  ML text directly into the @{verbatim ROOT.ML} file.
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   125
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   126
  \begin{description}
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   127
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   128
  \item @{ML show_types} and @{ML show_sorts} control printing of type
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   129
  constraints for term variables, and sort constraints for type
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   130
  variables.  By default, neither of these are shown in output.  If
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   131
  @{ML show_sorts} is set to @{ML true}, types are always shown as
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   132
  well.
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   133
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   134
  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
   135
  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
   136
  rule does not apply as expected.
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   137
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   138
  \item @{ML show_consts} controls printing of types of constants when
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   139
  displaying a goal state.
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   140
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   141
  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
   142
  often occur at several different type instances.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   143
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   144
  \item @{ML long_names}, @{ML short_names}, and @{ML unique_names}
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   145
  control the way of printing fully qualified internal names in
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   146
  external form.  See also \secref{sec:antiq} for the document
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   147
  antiquotation options of the same names.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   148
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   149
  \item @{ML show_brackets} controls bracketing in pretty printed
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   150
  output.  If set to @{ML true}, all sub-expressions of the pretty
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   151
  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
   152
  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
   153
  pretty printed entities may occasionally help to diagnose problems
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   154
  with operator priorities, for example.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   155
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   156
  \item @{ML eta_contract} controls @{text "\<eta>"}-contracted printing of
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   157
  terms.
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
  The @{text \<eta>}-contraction law asserts @{prop "(\<lambda>x. f x) \<equiv> f"},
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   160
  provided @{text x} is not free in @{text f}.  It asserts
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   161
  \emph{extensionality} of functions: @{prop "f \<equiv> g"} if @{prop "f x \<equiv>
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   162
  g x"} for all @{text x}.  Higher-order unification frequently puts
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   163
  terms into a fully @{text \<eta>}-expanded form.  For example, if @{text
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   164
  F} has type @{text "(\<tau> \<Rightarrow> \<tau>) \<Rightarrow> \<tau>"} then its expanded form is @{term
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   165
  "\<lambda>h. F (\<lambda>x. h x)"}.
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   166
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   167
  Setting @{ML eta_contract} makes Isabelle perform @{text
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   168
  \<eta>}-contractions before printing, so that @{term "\<lambda>h. F (\<lambda>x. h x)"}
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   169
  appears simply as @{text F}.
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   170
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   171
  Note that the distinction between a term and its @{text \<eta>}-expanded
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   172
  form occasionally matters.  While higher-order resolution and
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   173
  rewriting operate modulo @{text "\<alpha>\<beta>\<eta>"}-conversion, some other tools
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   174
  might look at terms more discretely.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   175
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   176
  \item @{ML goals_limit} controls the maximum number of subgoals to
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   177
  be shown in goal output.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   178
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   179
  \item @{ML Proof.show_main_goal} controls whether the main result to
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   180
  be proven should be displayed.  This information might be relevant
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   181
  for schematic goals, to inspect the current claim that has been
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   182
  synthesized so far.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   183
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   184
  \item @{ML show_hyps} controls printing of implicit hypotheses of
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   185
  local facts.  Normally, only those hypotheses are displayed that are
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   186
  \emph{not} covered by the assumptions of the current context: this
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   187
  situation indicates a fault in some tool being used.
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   188
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   189
  By setting @{ML show_hyps} to @{ML true}, output of \emph{all}
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   190
  hypotheses can be enforced, which is occasionally useful for
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   191
  diagnostic purposes.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   192
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   193
  \item @{ML show_tags} controls printing of extra annotations within
28765
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   194
  theorems, such as internal position information, or the case names
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   195
  being attached by the attribute @{attribute case_names}.
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   196
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   197
  Note that the @{attribute tagged} and @{attribute untagged}
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   198
  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
   199
  associated with a theorem.
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   200
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   201
  \item @{ML show_question_marks} controls printing of question marks
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   202
  for schematic variables, such as @{text ?x}.  Only the leading
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   203
  question mark is affected, the remaining text is unchanged
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   204
  (including proper markup for schematic variables that might be
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   205
  relevant for user interfaces).
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   206
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   207
  \end{description}
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   208
*}
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   209
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   210
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   211
subsection {* Printing limits *}
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   212
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   213
text {*
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   214
  \begin{mldecls}
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   215
    @{index_ML Pretty.setdepth: "int -> unit"} \\
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   216
    @{index_ML Pretty.setmargin: "int -> unit"} \\
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   217
    @{index_ML print_depth: "int -> unit"} \\
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   218
  \end{mldecls}
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   219
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   220
  These ML functions set limits for pretty printed text.
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   221
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   222
  \begin{description}
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   223
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   224
  \item @{ML Pretty.setdepth}~@{text d} tells the pretty printer to
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   225
  limit the printing depth to @{text d}.  This affects the display of
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   226
  types, terms, theorems etc.  The default value is 0, which permits
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   227
  printing to an arbitrary depth.  Other useful values for @{text d}
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   228
  are 10 and 20.
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   229
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   230
  \item @{ML Pretty.setmargin}~@{text m} tells the pretty printer to
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   231
  assume a right margin (page width) of @{text m}.  The initial margin
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   232
  is 76, but user interfaces might adapt the margin automatically when
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   233
  resizing windows.
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   234
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   235
  \item @{ML print_depth}~@{text n} limits the printing depth of the
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   236
  ML toplevel pretty printer; the precise effect depends on the ML
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   237
  compiler and run-time system.  Typically @{text n} should be less
da8f6f4a74be misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents: 28763
diff changeset
   238
  than 10.  Bigger values such as 100--1000 are useful for debugging.
28763
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   239
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   240
  \end{description}
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   241
*}
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   242
b5e6122ff575 added pretty printing options (from old ref manual);
wenzelm
parents: 28762
diff changeset
   243
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   244
section {* Mixfix annotations *}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   245
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   246
text {* Mixfix annotations specify concrete \emph{inner syntax} of
35351
7425aece4ee3 allow general mixfix syntax for type constructors;
wenzelm
parents: 32833
diff changeset
   247
  Isabelle types and terms.  Locally fixed parameters in toplevel
7425aece4ee3 allow general mixfix syntax for type constructors;
wenzelm
parents: 32833
diff changeset
   248
  theorem statements, locale specifications etc.\ also admit mixfix
7425aece4ee3 allow general mixfix syntax for type constructors;
wenzelm
parents: 32833
diff changeset
   249
  annotations.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   250
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   251
  \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   252
  \begin{rail}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   253
    infix: '(' ('infix' | 'infixl' | 'infixr') string nat ')'
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   254
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   255
    mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')'
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   256
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   257
    structmixfix: mixfix | '(' 'structure' ')'
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   258
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   259
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   260
    prios: '[' (nat + ',') ']'
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   261
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   262
  \end{rail}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   263
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   264
  Here the \railtok{string} specifications refer to the actual mixfix
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   265
  template, which may include literal text, spacing, blocks, and
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   266
  arguments (denoted by ``@{text _}''); the special symbol
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   267
  ``@{verbatim "\<index>"}'' (printed as ``@{text "\<index>"}'') represents an index
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   268
  argument that specifies an implicit structure reference (see also
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   269
  \secref{sec:locale}).  Infix and binder declarations provide common
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   270
  abbreviations for particular mixfix declarations.  So in practice,
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   271
  mixfix templates mostly degenerate to literal text for concrete
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   272
  syntax, such as ``@{verbatim "++"}'' for an infix symbol.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   273
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   274
  \medskip In full generality, mixfix declarations work as follows.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   275
  Suppose a constant @{text "c :: \<tau>\<^sub>1 \<Rightarrow> \<dots> \<tau>\<^sub>n \<Rightarrow> \<tau>"} is
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   276
  annotated by @{text "(mixfix [p\<^sub>1, \<dots>, p\<^sub>n] p)"}, where @{text
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   277
  "mixfix"} is a string @{text "d\<^sub>0 _ d\<^sub>1 _ \<dots> _ d\<^sub>n"} consisting of
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   278
  delimiters that surround argument positions as indicated by
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   279
  underscores.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   280
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   281
  Altogether this determines a production for a context-free priority
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   282
  grammar, where for each argument @{text "i"} the syntactic category
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   283
  is determined by @{text "\<tau>\<^sub>i"} (with priority @{text "p\<^sub>i"}), and
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   284
  the result category is determined from @{text "\<tau>"} (with
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   285
  priority @{text "p"}).  Priority specifications are optional, with
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   286
  default 0 for arguments and 1000 for the result.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   287
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   288
  Since @{text "\<tau>"} may be again a function type, the constant
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   289
  type scheme may have more argument positions than the mixfix
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   290
  pattern.  Printing a nested application @{text "c t\<^sub>1 \<dots> t\<^sub>m"} for
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   291
  @{text "m > n"} works by attaching concrete notation only to the
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   292
  innermost part, essentially by printing @{text "(c t\<^sub>1 \<dots> t\<^sub>n) \<dots> t\<^sub>m"}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   293
  instead.  If a term has fewer arguments than specified in the mixfix
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   294
  template, the concrete syntax is ignored.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   295
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   296
  \medskip A mixfix template may also contain additional directives
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   297
  for pretty printing, notably spaces, blocks, and breaks.  The
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   298
  general template format is a sequence over any of the following
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   299
  entities.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   300
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   301
  \begin{description}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   302
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   303
  \item @{text "d"} is a delimiter, namely a non-empty sequence of
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   304
  characters other than the following special characters:
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   305
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   306
  \smallskip
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   307
  \begin{tabular}{ll}
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   308
    @{verbatim "'"} & single quote \\
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   309
    @{verbatim "_"} & underscore \\
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   310
    @{text "\<index>"} & index symbol \\
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   311
    @{verbatim "("} & open parenthesis \\
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   312
    @{verbatim ")"} & close parenthesis \\
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   313
    @{verbatim "/"} & slash \\
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   314
  \end{tabular}
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   315
  \medskip
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   316
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   317
  \item @{verbatim "'"} escapes the special meaning of these
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   318
  meta-characters, producing a literal version of the following
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   319
  character, unless that is a blank.
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   320
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   321
  A single quote followed by a blank separates delimiters, without
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   322
  affecting printing, but input tokens may have additional white space
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   323
  here.
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   324
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   325
  \item @{verbatim "_"} is an argument position, which stands for a
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   326
  certain syntactic category in the underlying grammar.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   327
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   328
  \item @{text "\<index>"} is an indexed argument position; this is the place
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   329
  where implicit structure arguments can be attached.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   330
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   331
  \item @{text "s"} is a non-empty sequence of spaces for printing.
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   332
  This and the following specifications do not affect parsing at all.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   333
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   334
  \item @{verbatim "("}@{text n} opens a pretty printing block.  The
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   335
  optional number specifies how much indentation to add when a line
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   336
  break occurs within the block.  If the parenthesis is not followed
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   337
  by digits, the indentation defaults to 0.  A block specified via
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   338
  @{verbatim "(00"} is unbreakable.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   339
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   340
  \item @{verbatim ")"} closes a pretty printing block.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   341
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   342
  \item @{verbatim "//"} forces a line break.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   343
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   344
  \item @{verbatim "/"}@{text s} allows a line break.  Here @{text s}
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   345
  stands for the string of spaces (zero or more) right after the
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   346
  slash.  These spaces are printed if the break is \emph{not} taken.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   347
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   348
  \end{description}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   349
28771
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   350
  For example, the template @{verbatim "(_ +/ _)"} specifies an infix
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   351
  operator.  There are two argument positions; the delimiter
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   352
  @{verbatim "+"} is preceded by a space and followed by a space or
4510201c6aaf mixfix annotations: verbatim for special symbols;
wenzelm
parents: 28770
diff changeset
   353
  line break; the entire phrase is a pretty printing block.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   354
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   355
  The general idea of pretty printing with blocks and breaks is also
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   356
  described in \cite{paulson-ml2}.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   357
*}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   358
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   359
35413
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   360
section {* Explicit notation *}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   361
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   362
text {*
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   363
  \begin{matharray}{rcll}
35413
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   364
    @{command_def "type_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   365
    @{command_def "no_type_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   366
    @{command_def "notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   367
    @{command_def "no_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   368
  \end{matharray}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   369
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   370
  \begin{rail}
35413
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   371
    ('type\_notation' | 'no\_type\_notation') target? mode? \\ (nameref mixfix + 'and')
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   372
    ;
29741
wenzelm
parents: 29157
diff changeset
   373
    ('notation' | 'no\_notation') target? mode? \\ (nameref structmixfix + 'and')
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   374
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   375
  \end{rail}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   376
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   377
  \begin{description}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   378
35413
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   379
  \item @{command "type_notation"}~@{text "c (mx)"} associates mixfix
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   380
  syntax with an existing type constructor.  The arity of the
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   381
  constructor is retrieved from the context.
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   382
  
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   383
  \item @{command "no_type_notation"} is similar to @{command
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   384
  "type_notation"}, but removes the specified syntax annotation from
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   385
  the present context.
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   386
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   387
  \item @{command "notation"}~@{text "c (mx)"} associates mixfix
35413
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   388
  syntax with an existing constant or fixed variable.  The type
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   389
  declaration of the given entity is retrieved from the context.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   390
  
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   391
  \item @{command "no_notation"} is similar to @{command "notation"},
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   392
  but removes the specified syntax annotation from the present
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   393
  context.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   394
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   395
  \end{description}
35413
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   396
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   397
  Compared to the underlying @{command "syntax"} and @{command
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   398
  "no_syntax"} primitives (\secref{sec:syn-trans}), the above commands
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   399
  provide explicit checking wrt.\ the logical context, and work within
4c7cba1f7ce9 added type_notation command;
wenzelm
parents: 35351
diff changeset
   400
  general local theory targets, not just the global theory.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   401
*}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   402
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   403
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   404
section {* The Pure syntax \label{sec:pure-syntax} *}
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   405
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   406
subsection {* Priority grammars \label{sec:priority-grammar} *}
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   407
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   408
text {* A context-free grammar consists of a set of \emph{terminal
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   409
  symbols}, a set of \emph{nonterminal symbols} and a set of
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   410
  \emph{productions}.  Productions have the form @{text "A = \<gamma>"},
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   411
  where @{text A} is a nonterminal and @{text \<gamma>} is a string of
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   412
  terminals and nonterminals.  One designated nonterminal is called
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   413
  the \emph{root symbol}.  The language defined by the grammar
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   414
  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
   415
  root symbol by applying productions as rewrite rules.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   416
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   417
  The standard Isabelle parser for inner syntax uses a \emph{priority
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   418
  grammar}.  Each nonterminal is decorated by an integer priority:
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   419
  @{text "A\<^sup>(\<^sup>p\<^sup>)"}.  In a derivation, @{text "A\<^sup>(\<^sup>p\<^sup>)"} may be rewritten
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   420
  using a production @{text "A\<^sup>(\<^sup>q\<^sup>) = \<gamma>"} only if @{text "p \<le> q"}.  Any
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   421
  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
   422
  grammar by introducing new nonterminals and productions.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   423
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   424
  \medskip Formally, a set of context free productions @{text G}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   425
  induces a derivation relation @{text "\<longrightarrow>\<^sub>G"} as follows.  Let @{text
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   426
  \<alpha>} and @{text \<beta>} denote strings of terminal or nonterminal symbols.
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   427
  Then @{text "\<alpha> A\<^sup>(\<^sup>p\<^sup>) \<beta> \<longrightarrow>\<^sub>G \<alpha> \<gamma> \<beta>"} holds if and only if @{text G}
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   428
  contains some production @{text "A\<^sup>(\<^sup>q\<^sup>) = \<gamma>"} for @{text "p \<le> q"}.
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   429
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   430
  \medskip The following grammar for arithmetic expressions
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   431
  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
   432
  enforced by priorities.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   433
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   434
  \begin{center}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   435
  \begin{tabular}{rclr}
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   436
  @{text "A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "="} & @{verbatim "("} @{text "A\<^sup>(\<^sup>0\<^sup>)"} @{verbatim ")"} \\
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   437
  @{text "A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "="} & @{verbatim 0} \\
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   438
  @{text "A\<^sup>(\<^sup>0\<^sup>)"} & @{text "="} & @{text "A\<^sup>(\<^sup>0\<^sup>)"} @{verbatim "+"} @{text "A\<^sup>(\<^sup>1\<^sup>)"} \\
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   439
  @{text "A\<^sup>(\<^sup>2\<^sup>)"} & @{text "="} & @{text "A\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "*"} @{text "A\<^sup>(\<^sup>2\<^sup>)"} \\
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   440
  @{text "A\<^sup>(\<^sup>3\<^sup>)"} & @{text "="} & @{verbatim "-"} @{text "A\<^sup>(\<^sup>3\<^sup>)"} \\
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   441
  \end{tabular}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   442
  \end{center}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   443
  The choice of priorities determines that @{verbatim "-"} binds
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   444
  tighter than @{verbatim "*"}, which binds tighter than @{verbatim
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   445
  "+"}.  Furthermore @{verbatim "+"} associates to the left and
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   446
  @{verbatim "*"} to the right.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   447
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   448
  \medskip For clarity, grammars obey these conventions:
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   449
  \begin{itemize}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   450
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   451
  \item All priorities must lie between 0 and 1000.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   452
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   453
  \item Priority 0 on the right-hand side and priority 1000 on the
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   454
  left-hand side may be omitted.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   455
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   456
  \item The production @{text "A\<^sup>(\<^sup>p\<^sup>) = \<alpha>"} is written as @{text "A = \<alpha>
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   457
  (p)"}, i.e.\ the priority of the left-hand side actually appears in
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   458
  a column on the far right.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   459
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   460
  \item Alternatives are separated by @{text "|"}.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   461
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   462
  \item Repetition is indicated by dots @{text "(\<dots>)"} in an informal
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   463
  but obvious way.
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   464
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   465
  \end{itemize}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   466
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   467
  Using these conventions, the example grammar specification above
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   468
  takes the form:
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   469
  \begin{center}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   470
  \begin{tabular}{rclc}
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   471
    @{text A} & @{text "="} & @{verbatim "("} @{text A} @{verbatim ")"} \\
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   472
              & @{text "|"} & @{verbatim 0} & \qquad\qquad \\
28769
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   473
              & @{text "|"} & @{text A} @{verbatim "+"} @{text "A\<^sup>(\<^sup>1\<^sup>)"} & @{text "(0)"} \\
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   474
              & @{text "|"} & @{text "A\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "*"} @{text "A\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   475
              & @{text "|"} & @{verbatim "-"} @{text "A\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   476
  \end{tabular}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   477
  \end{center}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   478
*}
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   479
8fc228f21861 added section "Priority grammars" (variant from old ref manual);
wenzelm
parents: 28767
diff changeset
   480
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   481
subsection {* The Pure grammar *}
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   482
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   483
text {*
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   484
  The priority grammar of the @{text "Pure"} theory is defined as follows:
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   485
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   486
  %FIXME syntax for "index" (?)
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   487
  %FIXME "op" versions of ==> etc. (?)
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   488
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   489
  \begin{center}
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   490
  \begin{supertabular}{rclr}
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   491
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   492
  @{syntax_def (inner) any} & = & @{text "prop  |  logic"} \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   493
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   494
  @{syntax_def (inner) prop} & = & @{verbatim "("} @{text prop} @{verbatim ")"} \\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   495
    & @{text "|"} & @{text "prop\<^sup>(\<^sup>4\<^sup>)"} @{verbatim "::"} @{text type} & @{text "(3)"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   496
    & @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "=?="} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   497
    & @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "=="} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   498
    & @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{text "\<equiv>"} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28779
diff changeset
   499
    & @{text "|"} & @{text "prop\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "&&&"} @{text "prop\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   500
    & @{text "|"} & @{text "prop\<^sup>(\<^sup>2\<^sup>)"} @{verbatim "==>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   501
    & @{text "|"} & @{text "prop\<^sup>(\<^sup>2\<^sup>)"} @{text "\<Longrightarrow>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   502
    & @{text "|"} & @{verbatim "[|"} @{text prop} @{verbatim ";"} @{text "\<dots>"} @{verbatim ";"} @{text prop} @{verbatim "|]"} @{verbatim "==>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   503
    & @{text "|"} & @{text "\<lbrakk>"} @{text prop} @{verbatim ";"} @{text "\<dots>"} @{verbatim ";"} @{text prop} @{text "\<rbrakk>"} @{text "\<Longrightarrow>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   504
    & @{text "|"} & @{verbatim "!!"} @{text idts} @{verbatim "."} @{text prop} & @{text "(0)"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   505
    & @{text "|"} & @{text "\<And>"} @{text idts} @{verbatim "."} @{text prop} & @{text "(0)"} \\
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   506
    & @{text "|"} & @{verbatim OFCLASS} @{verbatim "("} @{text type} @{verbatim ","} @{text logic} @{verbatim ")"} \\
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   507
    & @{text "|"} & @{verbatim SORT_CONSTRAINT} @{verbatim "("} @{text type} @{verbatim ")"} \\
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28779
diff changeset
   508
    & @{text "|"} & @{verbatim TERM} @{text logic} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   509
    & @{text "|"} & @{verbatim PROP} @{text aprop} \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   510
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28779
diff changeset
   511
  @{syntax_def (inner) aprop} & = & @{verbatim "("} @{text aprop} @{verbatim ")"} \\
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28779
diff changeset
   512
    & @{text "|"} & @{text "id  |  longid  |  var  |  "}@{verbatim "_"}@{text "  |  "}@{verbatim "..."} \\
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28779
diff changeset
   513
    & @{text "|"} & @{verbatim CONST} @{text "id  |  "}@{verbatim CONST} @{text "longid"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   514
    & @{text "|"} & @{text "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>)"} & @{text "(999)"} \\\\
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   515
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   516
  @{syntax_def (inner) logic} & = & @{verbatim "("} @{text logic} @{verbatim ")"} \\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   517
    & @{text "|"} & @{text "logic\<^sup>(\<^sup>4\<^sup>)"} @{verbatim "::"} @{text type} & @{text "(3)"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   518
    & @{text "|"} & @{text "id  |  longid  |  var  |  "}@{verbatim "_"}@{text "  |  "}@{verbatim "..."} \\
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28779
diff changeset
   519
    & @{text "|"} & @{verbatim CONST} @{text "id  |  "}@{verbatim CONST} @{text "longid"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   520
    & @{text "|"} & @{text "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>)"} & @{text "(999)"} \\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   521
    & @{text "|"} & @{verbatim "%"} @{text pttrns} @{verbatim "."} @{text "any\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   522
    & @{text "|"} & @{text \<lambda>} @{text pttrns} @{verbatim "."} @{text "any\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   523
    & @{text "|"} & @{verbatim TYPE} @{verbatim "("} @{text type} @{verbatim ")"} \\\\
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   524
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   525
  @{syntax_def (inner) idt} & = & @{verbatim "("} @{text idt} @{verbatim ")"}@{text "  |  id  |  "}@{verbatim "_"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   526
    & @{text "|"} & @{text id} @{verbatim "::"} @{text type} & @{text "(0)"} \\
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   527
    & @{text "|"} & @{verbatim "_"} @{verbatim "::"} @{text type} & @{text "(0)"} \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   528
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   529
  @{syntax_def (inner) idts} & = & @{text "idt  |  idt\<^sup>(\<^sup>1\<^sup>) idts"} & @{text "(0)"} \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   530
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   531
  @{syntax_def (inner) pttrn} & = & @{text idt} \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   532
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   533
  @{syntax_def (inner) pttrns} & = & @{text "pttrn  |  pttrn\<^sup>(\<^sup>1\<^sup>) pttrns"} & @{text "(0)"} \\\\
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   534
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   535
  @{syntax_def (inner) type} & = & @{verbatim "("} @{text type} @{verbatim ")"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   536
    & @{text "|"} & @{text "tid  |  tvar  |  "}@{verbatim "_"} \\
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   537
    & @{text "|"} & @{text "tid"} @{verbatim "::"} @{text "sort  |  tvar  "}@{verbatim "::"} @{text "sort  |  "}@{verbatim "_"} @{verbatim "::"} @{text "sort"} \\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   538
    & @{text "|"} & @{text "id  |  type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) id  |  "}@{verbatim "("} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim ")"} @{text id} \\
29741
wenzelm
parents: 29157
diff changeset
   539
    & @{text "|"} & @{text "longid  |  type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) longid"} \\
wenzelm
parents: 29157
diff changeset
   540
    & @{text "|"} & @{verbatim "("} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim ")"} @{text longid} \\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   541
    & @{text "|"} & @{text "type\<^sup>(\<^sup>1\<^sup>)"} @{verbatim "=>"} @{text type} & @{text "(0)"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   542
    & @{text "|"} & @{text "type\<^sup>(\<^sup>1\<^sup>)"} @{text "\<Rightarrow>"} @{text type} & @{text "(0)"} \\
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   543
    & @{text "|"} & @{verbatim "["} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim "]"} @{verbatim "=>"} @{text type} & @{text "(0)"} \\
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   544
    & @{text "|"} & @{verbatim "["} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim "]"} @{text "\<Rightarrow>"} @{text type} & @{text "(0)"} \\\\
28772
3f6bc48ebb9b added Pure grammer (from old ref manual);
wenzelm
parents: 28771
diff changeset
   545
29741
wenzelm
parents: 29157
diff changeset
   546
  @{syntax_def (inner) sort} & = & @{text "id  |  longid  |  "}@{verbatim "{}"} \\
wenzelm
parents: 29157
diff changeset
   547
    & @{text "|"} & @{verbatim "{"} @{text "(id | longid)"} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text "(id | longid)"} @{verbatim "}"} \\
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   548
  \end{supertabular}
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   549
  \end{center}
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   550
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   551
  \medskip Here literal terminals are printed @{verbatim "verbatim"};
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   552
  see also \secref{sec:inner-lex} for further token categories of the
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   553
  inner syntax.  The meaning of the nonterminals defined by the above
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   554
  grammar is as follows:
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   555
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   556
  \begin{description}
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   557
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   558
  \item @{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
   559
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   560
  \item @{syntax_ref (inner) prop} denotes meta-level propositions,
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   561
  which are terms of type @{typ prop}.  The syntax of such formulae of
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   562
  the meta-logic is carefully distinguished from usual conventions for
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   563
  object-logics.  In particular, plain @{text "\<lambda>"}-term notation is
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   564
  \emph{not} recognized as @{syntax (inner) prop}.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   565
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   566
  \item @{syntax_ref (inner) aprop} denotes atomic propositions, which
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   567
  are embedded into regular @{syntax (inner) prop} by means of an
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   568
  explicit @{verbatim PROP} token.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   569
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   570
  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
   571
  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
   572
  prop} are expected to provide their own concrete syntax; otherwise
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   573
  the printed version will appear like @{syntax (inner) logic} and
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   574
  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
   575
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   576
  \item @{syntax_ref (inner) logic} denotes arbitrary terms of a
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   577
  logical type, excluding type @{typ prop}.  This is the main
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   578
  syntactic category of object-logic entities, covering plain @{text
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   579
  \<lambda>}-term notation (variables, abstraction, application), plus
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   580
  anything defined by the user.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   581
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   582
  When specifying notation for logical entities, all logical types
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   583
  (excluding @{typ prop}) are \emph{collapsed} to this single category
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   584
  of @{syntax (inner) logic}.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   585
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   586
  \item @{syntax_ref (inner) idt} denotes identifiers, possibly
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   587
  constrained by types.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   588
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   589
  \item @{syntax_ref (inner) idts} denotes a sequence of @{syntax_ref
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   590
  (inner) idt}.  This is the most basic category for variables in
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   591
  iterated binders, such as @{text "\<lambda>"} or @{text "\<And>"}.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   592
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   593
  \item @{syntax_ref (inner) pttrn} and @{syntax_ref (inner) pttrns}
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   594
  denote patterns for abstraction, cases bindings etc.  In Pure, these
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   595
  categories start as a merely copy of @{syntax (inner) idt} and
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   596
  @{syntax (inner) idts}, respectively.  Object-logics may add
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   597
  additional productions for binding forms.
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   598
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   599
  \item @{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
   600
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   601
  \item @{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
   602
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   603
  \end{description}
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   604
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   605
  Here are some further explanations of certain syntax features.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   606
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   607
  \begin{itemize}
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   608
28778
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   609
  \item In @{syntax (inner) idts}, note that @{text "x :: nat y"} is
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   610
  parsed as @{text "x :: (nat y)"}, treating @{text y} like a type
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   611
  constructor applied to @{text nat}.  To avoid this interpretation,
a25630deacaf misc tuning of inner syntax;
wenzelm
parents: 28777
diff changeset
   612
  write @{text "(x :: nat) y"} with explicit parentheses.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   613
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   614
  \item Similarly, @{text "x :: nat y :: nat"} is parsed as @{text "x ::
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   615
  (nat y :: nat)"}.  The correct form is @{text "(x :: nat) (y ::
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   616
  nat)"}, or @{text "(x :: nat) y :: nat"} if @{text y} is last in the
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   617
  sequence of identifiers.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   618
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   619
  \item Type constraints for terms bind very weakly.  For example,
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   620
  @{text "x < y :: nat"} is normally parsed as @{text "(x < y) ::
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   621
  nat"}, unless @{text "<"} has a very low priority, in which case the
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   622
  input is likely to be ambiguous.  The correct form is @{text "x < (y
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   623
  :: nat)"}.
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   624
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   625
  \item Constraints may be either written with two literal colons
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   626
  ``@{verbatim "::"}'' or the double-colon symbol @{verbatim "\<Colon>"},
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   627
  which actually looks exactly the same in some {\LaTeX} styles.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   628
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   629
  \item Dummy variables (written as underscore) may occur in different
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   630
  roles.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   631
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   632
  \begin{description}
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   633
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   634
  \item A type ``@{text "_"}'' or ``@{text "_ :: sort"}'' acts like an
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   635
  anonymous inference parameter, which is filled-in according to the
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   636
  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
   637
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   638
  \item A bound ``@{text "_"}'' refers to a vacuous abstraction, where
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   639
  the body does not refer to the binding introduced here.  As in the
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   640
  term @{term "\<lambda>x _. x"}, which is @{text "\<alpha>"}-equivalent to @{text
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   641
  "\<lambda>x y. x"}.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   642
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   643
  \item A free ``@{text "_"}'' refers to an implicit outer binding.
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   644
  Higher definitional packages usually allow forms like @{text "f x _
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   645
  = x"}.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   646
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   647
  \item A schematic ``@{text "_"}'' (within a term pattern, see
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   648
  \secref{sec:term-decls}) refers to an anonymous variable that is
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   649
  implicitly abstracted over its context of locally bound variables.
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   650
  For example, this allows pattern matching of @{text "{x. f x = g
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   651
  x}"} against @{text "{x. _ = _}"}, or even @{text "{_. _ = _}"} by
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   652
  using both bound and schematic dummies.
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   653
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   654
  \end{description}
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   655
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   656
  \item The three literal dots ``@{verbatim "..."}'' may be also
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   657
  written as ellipsis symbol @{verbatim "\<dots>"}.  In both cases this
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   658
  refers to a special schematic variable, which is bound in the
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   659
  context.  This special term abbreviation works nicely with
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   660
  calculational reasoning (\secref{sec:calculation}).
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   661
28773
39b4cedb8433 updated and elaborated Pure grammer;
wenzelm
parents: 28772
diff changeset
   662
  \end{itemize}
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   663
*}
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   664
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   665
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   666
section {* Lexical matters \label{sec:inner-lex} *}
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   667
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   668
text {* The inner lexical syntax vaguely resembles the outer one
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   669
  (\secref{sec:outer-lex}), but some details are different.  There are
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   670
  two main categories of inner syntax tokens:
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   671
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   672
  \begin{enumerate}
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   673
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   674
  \item \emph{delimiters} --- the literal tokens occurring in
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   675
  productions of the given priority grammar (cf.\
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   676
  \secref{sec:priority-grammar});
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   677
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   678
  \item \emph{named tokens} --- various categories of identifiers etc.
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   679
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   680
  \end{enumerate}
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   681
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   682
  Delimiters override named tokens and may thus render certain
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   683
  identifiers inaccessible.  Sometimes the logical context admits
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   684
  alternative ways to refer to the same entity, potentially via
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   685
  qualified names.
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   686
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   687
  \medskip The categories for named tokens are defined once and for
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   688
  all as follows, reusing some categories of the outer token syntax
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   689
  (\secref{sec:outer-lex}).
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   690
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   691
  \begin{center}
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   692
  \begin{supertabular}{rcl}
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   693
    @{syntax_def (inner) id} & = & @{syntax_ref ident} \\
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   694
    @{syntax_def (inner) longid} & = & @{syntax_ref longident} \\
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   695
    @{syntax_def (inner) var} & = & @{syntax_ref var} \\
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   696
    @{syntax_def (inner) tid} & = & @{syntax_ref typefree} \\
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   697
    @{syntax_def (inner) tvar} & = & @{syntax_ref typevar} \\
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   698
    @{syntax_def (inner) num} & = & @{syntax_ref nat}@{text "  |  "}@{verbatim "-"}@{syntax_ref nat} \\
29157
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   699
    @{syntax_def (inner) float_token} & = & @{syntax_ref nat}@{verbatim "."}@{syntax_ref nat}@{text "  |  "}@{verbatim "-"}@{syntax_ref nat}@{verbatim "."}@{syntax_ref nat} \\
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   700
    @{syntax_def (inner) xnum} & = & @{verbatim "#"}@{syntax_ref nat}@{text "  |  "}@{verbatim "#-"}@{syntax_ref nat} \\
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   701
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   702
    @{syntax_def (inner) xstr} & = & @{verbatim "''"} @{text "\<dots>"} @{verbatim "''"} \\
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   703
  \end{supertabular}
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   704
  \end{center}
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   705
29157
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   706
  The token categories @{syntax (inner) num}, @{syntax (inner)
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   707
  float_token}, @{syntax (inner) xnum}, and @{syntax (inner) xstr} are
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   708
  not used in Pure.  Object-logics may implement numerals and string
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   709
  constants by adding appropriate syntax declarations, together with
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   710
  some translation functions (e.g.\ see Isabelle/HOL).
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   711
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   712
  The derived categories @{syntax_def (inner) num_const} and
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   713
  @{syntax_def (inner) float_const} provide robust access to @{syntax
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   714
  (inner) num}, and @{syntax (inner) float_token}, respectively: the
461f34ed79ec added float_token, and num_const, float_const;
wenzelm
parents: 28856
diff changeset
   715
  syntax tree holds a syntactic constant instead of a free variable.
28777
2eeeced17228 added inner lexical syntax, reusing outer one;
wenzelm
parents: 28774
diff changeset
   716
*}
28774
0e25ef17b06b more tuning of Pure grammer;
wenzelm
parents: 28773
diff changeset
   717
28770
93a372e2dc7a added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents: 28769
diff changeset
   718
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   719
section {* Syntax and translations \label{sec:syn-trans} *}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   720
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   721
text {*
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   722
  \begin{matharray}{rcl}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   723
    @{command_def "nonterminals"} & : & @{text "theory \<rightarrow> theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   724
    @{command_def "syntax"} & : & @{text "theory \<rightarrow> theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   725
    @{command_def "no_syntax"} & : & @{text "theory \<rightarrow> theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   726
    @{command_def "translations"} & : & @{text "theory \<rightarrow> theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   727
    @{command_def "no_translations"} & : & @{text "theory \<rightarrow> theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   728
  \end{matharray}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   729
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   730
  \begin{rail}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   731
    'nonterminals' (name +)
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   732
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   733
    ('syntax' | 'no\_syntax') mode? (constdecl +)
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   734
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   735
    ('translations' | 'no\_translations') (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +)
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   736
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   737
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   738
    mode: ('(' ( name | 'output' | name 'output' ) ')')
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   739
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   740
    transpat: ('(' nameref ')')? string
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   741
    ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   742
  \end{rail}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   743
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   744
  \begin{description}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   745
  
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   746
  \item @{command "nonterminals"}~@{text c} declares a type
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   747
  constructor @{text c} (without arguments) to act as purely syntactic
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   748
  type: a nonterminal symbol of the inner syntax.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   749
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   750
  \item @{command "syntax"}~@{text "(mode) decls"} is similar to
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   751
  @{command "consts"}~@{text decls}, except that the actual logical
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   752
  signature extension is omitted.  Thus the context free grammar of
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   753
  Isabelle's inner syntax may be augmented in arbitrary ways,
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   754
  independently of the logic.  The @{text mode} argument refers to the
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   755
  print mode that the grammar rules belong; unless the @{keyword_ref
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   756
  "output"} indicator is given, all productions are added both to the
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   757
  input and output grammar.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   758
  
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   759
  \item @{command "no_syntax"}~@{text "(mode) decls"} removes grammar
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   760
  declarations (and translations) resulting from @{text decls}, which
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   761
  are interpreted in the same manner as for @{command "syntax"} above.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   762
  
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   763
  \item @{command "translations"}~@{text rules} specifies syntactic
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   764
  translation rules (i.e.\ macros): parse~/ print rules (@{text "\<rightleftharpoons>"}),
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   765
  parse rules (@{text "\<rightharpoonup>"}), or print rules (@{text "\<leftharpoondown>"}).
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   766
  Translation patterns may be prefixed by the syntactic category to be
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   767
  used for parsing; the default is @{text logic}.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   768
  
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   769
  \item @{command "no_translations"}~@{text rules} removes syntactic
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   770
  translation rules, which are interpreted in the same manner as for
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   771
  @{command "translations"} above.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   772
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   773
  \end{description}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   774
*}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   775
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   776
28779
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   777
section {* Syntax translation functions \label{sec:tr-funs} *}
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   778
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   779
text {*
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   780
  \begin{matharray}{rcl}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   781
    @{command_def "parse_ast_translation"} & : & @{text "theory \<rightarrow> theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   782
    @{command_def "parse_translation"} & : & @{text "theory \<rightarrow> theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   783
    @{command_def "print_translation"} & : & @{text "theory \<rightarrow> theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   784
    @{command_def "typed_print_translation"} & : & @{text "theory \<rightarrow> theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   785
    @{command_def "print_ast_translation"} & : & @{text "theory \<rightarrow> theory"} \\
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   786
  \end{matharray}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   787
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   788
  \begin{rail}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   789
  ( 'parse\_ast\_translation' | 'parse\_translation' | 'print\_translation' |
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   790
    'typed\_print\_translation' | 'print\_ast\_translation' ) ('(advanced)')? text
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   791
  ;
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   792
  \end{rail}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   793
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   794
  Syntax translation functions written in ML admit almost arbitrary
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   795
  manipulations of Isabelle's inner syntax.  Any of the above commands
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   796
  have a single \railqtok{text} argument that refers to an ML
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   797
  expression of appropriate type, which are as follows by default:
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   798
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   799
%FIXME proper antiquotations
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   800
\begin{ttbox}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   801
val parse_ast_translation   : (string * (ast list -> ast)) list
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   802
val parse_translation       : (string * (term list -> term)) list
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   803
val print_translation       : (string * (term list -> term)) list
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   804
val typed_print_translation :
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   805
  (string * (bool -> typ -> term list -> term)) list
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   806
val print_ast_translation   : (string * (ast list -> ast)) list
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   807
\end{ttbox}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   808
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   809
  If the @{text "(advanced)"} option is given, the corresponding
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   810
  translation functions may depend on the current theory or proof
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   811
  context.  This allows to implement advanced syntax mechanisms, as
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   812
  translations functions may refer to specific theory declarations or
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   813
  auxiliary proof data.
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   814
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 29741
diff changeset
   815
  See also \cite{isabelle-ref} for more information on the general
b6212ae21656 markup antiquotation options;
wenzelm
parents: 29741
diff changeset
   816
  concept of syntax transformations in Isabelle.
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   817
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   818
%FIXME proper antiquotations
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   819
\begin{ttbox}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   820
val parse_ast_translation:
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   821
  (string * (Proof.context -> ast list -> ast)) list
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   822
val parse_translation:
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   823
  (string * (Proof.context -> term list -> term)) list
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   824
val print_translation:
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   825
  (string * (Proof.context -> term list -> term)) list
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   826
val typed_print_translation:
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   827
  (string * (Proof.context -> bool -> typ -> term list -> term)) list
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   828
val print_ast_translation:
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   829
  (string * (Proof.context -> ast list -> ast)) list
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   830
\end{ttbox}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   831
*}
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   832
28779
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   833
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   834
section {* Inspecting the syntax *}
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   835
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   836
text {*
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   837
  \begin{matharray}{rcl}
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   838
    @{command_def "print_syntax"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   839
  \end{matharray}
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   840
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   841
  \begin{description}
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   842
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   843
  \item @{command "print_syntax"} prints the inner syntax of the
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   844
  current context.  The output can be quite large; the most important
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   845
  sections are explained below.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   846
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   847
  \begin{description}
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   848
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   849
  \item @{text "lexicon"} lists the delimiters of the inner token
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   850
  language; see \secref{sec:inner-lex}.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   851
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   852
  \item @{text "prods"} lists the productions of the underlying
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   853
  priority grammar; see \secref{sec:priority-grammar}.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   854
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   855
  The nonterminal @{text "A\<^sup>(\<^sup>p\<^sup>)"} is rendered in plain text as @{text
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   856
  "A[p]"}; delimiters are quoted.  Many productions have an extra
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   857
  @{text "\<dots> => name"}.  These names later become the heads of parse
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   858
  trees; they also guide the pretty printer.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   859
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   860
  Productions without such parse tree names are called \emph{copy
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   861
  productions}.  Their right-hand side must have exactly one
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   862
  nonterminal symbol (or named token).  The parser does not create a
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   863
  new parse tree node for copy productions, but simply returns the
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   864
  parse tree of the right-hand symbol.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   865
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   866
  If the right-hand side of a copy production consists of a single
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   867
  nonterminal without any delimiters, then it is called a \emph{chain
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   868
  production}.  Chain productions act as abbreviations: conceptually,
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   869
  they are removed from the grammar by adding new productions.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   870
  Priority information attached to chain productions is ignored; only
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   871
  the dummy value @{text "-1"} is displayed.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   872
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28779
diff changeset
   873
  \item @{text "print modes"} lists the alternative print modes
28779
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   874
  provided by this grammar; see \secref{sec:print-modes}.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   875
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   876
  \item @{text "parse_rules"} and @{text "print_rules"} relate to
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   877
  syntax translations (macros); see \secref{sec:syn-trans}.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   878
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   879
  \item @{text "parse_ast_translation"} and @{text
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   880
  "print_ast_translation"} list sets of constants that invoke
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   881
  translation functions for abstract syntax trees, which are only
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   882
  required in very special situations; see \secref{sec:tr-funs}.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   883
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   884
  \item @{text "parse_translation"} and @{text "print_translation"}
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   885
  list the sets of constants that invoke regular translation
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   886
  functions; see \secref{sec:tr-funs}.
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   887
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   888
  \end{description}
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   889
  
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   890
  \end{description}
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   891
*}
698960f08652 separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents: 28778
diff changeset
   892
28762
f5d79aeffd81 separate chapter "Inner syntax --- the term language";
wenzelm
parents:
diff changeset
   893
end