doc-src/IsarRef/Thy/pure.thy
author wenzelm
Mon, 02 Jun 2008 23:38:25 +0200
changeset 27050 cd8d99b9ef09
parent 27040 3d3e6e07b931
permissions -rw-r--r--
tuned structure;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26767
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
     1
(* $Id$ *)
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
     2
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
     3
theory pure
26957
e3f04fdd994d eliminated theory CPure;
wenzelm
parents: 26894
diff changeset
     4
imports Pure
26767
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
     5
begin
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
     6
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
     7
chapter {* Basic language elements \label{ch:pure-syntax} *}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
     8
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
     9
section {* Other commands *}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    10
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    11
subsection {* Diagnostics *}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    12
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    13
text {*
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    14
  \begin{matharray}{rcl}
26866
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
    15
    @{command_def "pr"}@{text "\<^sup>*"} & : & \isarkeep{\cdot} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
    16
    @{command_def "thm"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
    17
    @{command_def "term"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
    18
    @{command_def "prop"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
    19
    @{command_def "typ"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
    20
    @{command_def "prf"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
    21
    @{command_def "full_prf"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
26767
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    22
  \end{matharray}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    23
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    24
  These diagnostic commands assist interactive development.  Note that
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    25
  @{command undo} does not apply here, the theory or proof
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    26
  configuration is not changed.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    27
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    28
  \begin{rail}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    29
    'pr' modes? nat? (',' nat)?
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    30
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    31
    'thm' modes? thmrefs
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    32
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    33
    'term' modes? term
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    34
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    35
    'prop' modes? prop
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    36
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    37
    'typ' modes? type
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    38
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    39
    'prf' modes? thmrefs?
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    40
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    41
    'full\_prf' modes? thmrefs?
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    42
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    43
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    44
    modes: '(' (name + ) ')'
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    45
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    46
  \end{rail}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    47
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    48
  \begin{descr}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    49
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    50
  \item [@{command "pr"}~@{text "goals, prems"}] prints the current
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    51
  proof state (if present), including the proof context, current facts
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    52
  and goals.  The optional limit arguments affect the number of goals
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    53
  and premises to be displayed, which is initially 10 for both.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    54
  Omitting limit values leaves the current setting unchanged.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    55
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    56
  \item [@{command "thm"}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}] retrieves
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    57
  theorems from the current theory or proof context.  Note that any
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    58
  attributes included in the theorem specifications are applied to a
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    59
  temporary context derived from the current theory or proof; the
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    60
  result is discarded, i.e.\ attributes involved in @{text "a\<^sub>1,
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    61
  \<dots>, a\<^sub>n"} do not have any permanent effect.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    62
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    63
  \item [@{command "term"}~@{text t} and @{command "prop"}~@{text \<phi>}]
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    64
  read, type-check and print terms or propositions according to the
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    65
  current theory or proof context; the inferred type of @{text t} is
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    66
  output as well.  Note that these commands are also useful in
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    67
  inspecting the current environment of term abbreviations.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    68
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    69
  \item [@{command "typ"}~@{text \<tau>}] reads and prints types of the
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    70
  meta-logic according to the current theory or proof context.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    71
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    72
  \item [@{command "prf"}] displays the (compact) proof term of the
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    73
  current proof state (if present), or of the given theorems. Note
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    74
  that this requires proof terms to be switched on for the current
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    75
  object logic (see the ``Proof terms'' section of the Isabelle
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    76
  reference manual for information on how to do this).
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    77
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    78
  \item [@{command "full_prf"}] is like @{command "prf"}, but displays
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    79
  the full proof term, i.e.\ also displays information omitted in the
26777
134529bc72db misc tuning;
wenzelm
parents: 26767
diff changeset
    80
  compact proof term, which is denoted by ``@{text _}'' placeholders
134529bc72db misc tuning;
wenzelm
parents: 26767
diff changeset
    81
  there.
26767
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    82
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    83
  \end{descr}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    84
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    85
  All of the diagnostic commands above admit a list of @{text modes}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    86
  to be specified, which is appended to the current print mode (see
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    87
  also \cite{isabelle-ref}).  Thus the output behavior may be modified
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    88
  according particular print mode features.  For example, @{command
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    89
  "pr"}~@{text "(latex xsymbols symbols)"} would print the current
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    90
  proof state with mathematical symbols and special characters
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    91
  represented in {\LaTeX} source, according to the Isabelle style
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    92
  \cite{isabelle-sys}.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    93
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    94
  Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    95
  systematic way to include formal items into the printed text
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    96
  document.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    97
*}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    98
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
    99
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   100
subsection {* Inspecting the context *}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   101
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   102
text {*
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   103
  \begin{matharray}{rcl}
26866
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   104
    @{command_def "print_commands"}@{text "\<^sup>*"} & : & \isarkeep{\cdot} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   105
    @{command_def "print_theory"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   106
    @{command_def "print_syntax"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   107
    @{command_def "print_methods"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   108
    @{command_def "print_attributes"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   109
    @{command_def "print_theorems"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   110
    @{command_def "find_theorems"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
26894
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26870
diff changeset
   111
    @{command_def "thm_deps"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
26866
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   112
    @{command_def "print_facts"}@{text "\<^sup>*"} & : & \isarkeep{proof} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   113
    @{command_def "print_binds"}@{text "\<^sup>*"} & : & \isarkeep{proof} \\
26767
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   114
  \end{matharray}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   115
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   116
  \begin{rail}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   117
    'print\_theory' ( '!'?)
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   118
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   119
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   120
    'find\_theorems' (('(' (nat)? ('with\_dups')? ')')?) (criterion *)
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   121
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   122
    criterion: ('-'?) ('name' ':' nameref | 'intro' | 'elim' | 'dest' |
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   123
      'simp' ':' term | term)
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   124
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   125
    'thm\_deps' thmrefs
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   126
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   127
  \end{rail}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   128
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   129
  These commands print certain parts of the theory and proof context.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   130
  Note that there are some further ones available, such as for the set
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   131
  of rules declared for simplifications.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   132
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   133
  \begin{descr}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   134
  
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   135
  \item [@{command "print_commands"}] prints Isabelle's outer theory
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   136
  syntax, including keywords and command.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   137
  
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   138
  \item [@{command "print_theory"}] prints the main logical content of
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   139
  the theory context; the ``@{text "!"}'' option indicates extra
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   140
  verbosity.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   141
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   142
  \item [@{command "print_syntax"}] prints the inner syntax of types
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   143
  and terms, depending on the current context.  The output can be very
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   144
  verbose, including grammar tables and syntax translation rules.  See
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   145
  \cite[\S7, \S8]{isabelle-ref} for further information on Isabelle's
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   146
  inner syntax.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   147
  
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   148
  \item [@{command "print_methods"}] prints all proof methods
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   149
  available in the current theory context.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   150
  
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   151
  \item [@{command "print_attributes"}] prints all attributes
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   152
  available in the current theory context.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   153
  
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   154
  \item [@{command "print_theorems"}] prints theorems resulting from
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   155
  the last command.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   156
  
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   157
  \item [@{command "find_theorems"}~@{text criteria}] retrieves facts
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   158
  from the theory or proof context matching all of given search
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   159
  criteria.  The criterion @{text "name: p"} selects all theorems
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   160
  whose fully qualified name matches pattern @{text p}, which may
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   161
  contain ``@{text "*"}'' wildcards.  The criteria @{text intro},
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   162
  @{text elim}, and @{text dest} select theorems that match the
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   163
  current goal as introduction, elimination or destruction rules,
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   164
  respectively.  The criterion @{text "simp: t"} selects all rewrite
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   165
  rules whose left-hand side matches the given term.  The criterion
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   166
  term @{text t} selects all theorems that contain the pattern @{text
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   167
  t} -- as usual, patterns may contain occurrences of the dummy
26777
134529bc72db misc tuning;
wenzelm
parents: 26767
diff changeset
   168
  ``@{text _}'', schematic variables, and type constraints.
26767
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   169
  
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   170
  Criteria can be preceded by ``@{text "-"}'' to select theorems that
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   171
  do \emph{not} match. Note that giving the empty list of criteria
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   172
  yields \emph{all} currently known facts.  An optional limit for the
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   173
  number of printed facts may be given; the default is 40.  By
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   174
  default, duplicates are removed from the search result. Use
26894
1120f6cc10b0 proper checking of various Isar elements;
wenzelm
parents: 26870
diff changeset
   175
  @{text with_dups} to display duplicates.
26767
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   176
  
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   177
  \item [@{command "thm_deps"}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}]
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   178
  visualizes dependencies of facts, using Isabelle's graph browser
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   179
  tool (see also \cite{isabelle-sys}).
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   180
  
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   181
  \item [@{command "print_facts"}] prints all local facts of the
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   182
  current context, both named and unnamed ones.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   183
  
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   184
  \item [@{command "print_binds"}] prints all term abbreviations
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   185
  present in the context.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   186
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   187
  \end{descr}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   188
*}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   189
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   190
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   191
subsection {* History commands \label{sec:history} *}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   192
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   193
text {*
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   194
  \begin{matharray}{rcl}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   195
    @{command_def "undo"}^{{ * }{ * }} & : & \isarkeep{\cdot} \\
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   196
    @{command_def "redo"}^{{ * }{ * }} & : & \isarkeep{\cdot} \\
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   197
    @{command_def "kill"}^{{ * }{ * }} & : & \isarkeep{\cdot} \\
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   198
  \end{matharray}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   199
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   200
  The Isabelle/Isar top-level maintains a two-stage history, for
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   201
  theory and proof state transformation.  Basically, any command can
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   202
  be undone using @{command "undo"}, excluding mere diagnostic
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   203
  elements.  Its effect may be revoked via @{command "redo"}, unless
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   204
  the corresponding @{command "undo"} step has crossed the beginning
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   205
  of a proof or theory.  The @{command "kill"} command aborts the
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   206
  current history node altogether, discontinuing a proof or even the
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   207
  whole theory.  This operation is \emph{not} undo-able.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   208
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   209
  \begin{warn}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   210
    History commands should never be used with user interfaces such as
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   211
    Proof~General \cite{proofgeneral,Aspinall:TACAS:2000}, which takes
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   212
    care of stepping forth and back itself.  Interfering by manual
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   213
    @{command "undo"}, @{command "redo"}, or even @{command "kill"}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   214
    commands would quickly result in utter confusion.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   215
  \end{warn}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   216
*}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   217
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   218
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   219
subsection {* System operations *}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   220
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   221
text {*
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   222
  \begin{matharray}{rcl}
26866
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   223
    @{command_def "cd"}@{text "\<^sup>*"} & : & \isarkeep{\cdot} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   224
    @{command_def "pwd"}@{text "\<^sup>*"} & : & \isarkeep{\cdot} \\
3cff7b336c75 proper antiquotations for commands;
wenzelm
parents: 26852
diff changeset
   225
    @{command_def "use_thy"}@{text "\<^sup>*"} & : & \isarkeep{\cdot} \\
26767
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   226
  \end{matharray}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   227
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   228
  \begin{rail}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   229
    ('cd' | 'use\_thy' | 'update\_thy') name
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   230
    ;
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   231
  \end{rail}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   232
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   233
  \begin{descr}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   234
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   235
  \item [@{command "cd"}~@{text path}] changes the current directory
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   236
  of the Isabelle process.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   237
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   238
  \item [@{command "pwd"}] prints the current working directory.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   239
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   240
  \item [@{command "use_thy"}~@{text A}] preload theory @{text A}.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   241
  These system commands are scarcely used when working interactively,
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   242
  since loading of theories is done automatically as required.
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   243
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   244
  \end{descr}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   245
*}
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   246
cc127cc0951b converted pure.tex to Thy/pure.thy;
wenzelm
parents:
diff changeset
   247
end