CookBook/Package/Ind_Interface.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 20 Feb 2009 23:19:41 +0000
changeset 127 74846cb0fff9
parent 126 fcc0e6e54dca
child 129 e0d368a45537
permissions -rw-r--r--
updated and added two tentative recipes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     1
theory Ind_Interface
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
     2
imports "../Base" "../Parsing" Simple_Inductive_Package
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     3
begin
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     4
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
     5
section {* Parsing the Specification *}
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
     6
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
     7
text {* 
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
     8
  To be able to write down the specification in Isabelle, we have to introduce
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
     9
  a new command (see Section~\ref{sec:newcommand}).  As the keyword for the new
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    10
  command we chose \simpleinductive{}. We want that the package can cope with
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    11
  specifications inside locales. For example it should be possible to declare
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    12
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    13
*}
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    14
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    15
locale rel =
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    16
  fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    17
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    18
text {*
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    19
  and then define the transitive closure and the accessible part as follows:
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    20
*}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    21
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    22
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    23
simple_inductive (in rel) trcl' 
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    24
where
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    25
  base: "trcl' x x"
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    26
| step: "trcl' x y \<Longrightarrow> R y z \<Longrightarrow> trcl' x z"
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    27
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    28
simple_inductive (in rel) accpart'
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    29
where
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    30
  accpartI: "(\<And>y. R y x \<Longrightarrow> accpart' y) \<Longrightarrow> accpart' x"
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    31
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    32
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    33
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    34
text {*
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    35
  After the keyword we expect a constant (or constants) with possible typing 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    36
  annotations and a
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    37
  list of introduction rules. While these specifications are all
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    38
  straightforward, there is a technicality we like to deal with to do with
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    39
  fixed parameters and locales. Remember we pointed out that the parameter
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    40
  @{text R} is fixed throughout the specifications of @{text trcl} and @{text
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    41
  accpart}. The point is that they might be fixed in a locale and we like to
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    42
  support this. Accordingly we treat some parameters of the inductive
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    43
  definition specially; see Figure~\ref{fig:inddefsfixed} where the transitive
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    44
  closure and accessible part are defined with a fixed parameter @{text R} and
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    45
  also inside a locale fixing @{text R}.
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    46
*}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    47
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    48
text_raw {*
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    49
  \begin{figure}[p]
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    50
  \begin{isabelle}
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    51
*}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    52
simple_inductive
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    53
  trcl'' for R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    54
where
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    55
  base: "trcl'' R x x"
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    56
| step: "trcl'' R x y \<Longrightarrow> R y z \<Longrightarrow> trcl'' R x z"
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    57
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    58
simple_inductive
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    59
  accpart'' for R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    60
where
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    61
  accpartI: "(\<And>y. R y x \<Longrightarrow> accpart'' R y) \<Longrightarrow> accpart'' R x"
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    62
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    63
text_raw {*
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    64
  \end{isabelle}
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    65
  \caption{The first definition is for the transitive closure where the
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    66
  relation @{text R} is explicitly fixed. Simiraly the second definition
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    67
  of the accessible part of the relation @{text R}. The last two definitions
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    68
  specify the same inductive predicates, but this time defined inside
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    69
  a locale.\label{fig:inddefsfixed}}
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    70
  \end{figure}
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    71
*}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    72
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    73
text {*
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    74
  \begin{figure}[p]
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    75
  \begin{isabelle}
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    76
  \railnontermfont{\rmfamily\itshape}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    77
  \railterm{simpleinductive,where,for}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    78
  \railalias{simpleinductive}{\simpleinductive{}}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    79
  \railalias{where}{\isacommand{where}}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    80
  \railalias{for}{\isacommand{for}}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    81
  \begin{rail}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    82
  simpleinductive target? fixes (for fixes)? \\
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    83
  (where (thmdecl? prop + '|'))?
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    84
  ;
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    85
  \end{rail}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    86
  \end{isabelle}
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    87
  \caption{A railroad diagram describing the syntax of \simpleinductive{}. 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    88
  The \emph{target} indicates an optional locale; the \emph{fixes} are an 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    89
  \isacommand{and}-separated list of names for the inductive predicates (they
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    90
  can also contain typing- and syntax anotations); similarly the \emph{fixes} 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    91
  after \isacommand{for} to indicate fixed parameters; \emph{prop} stands for a 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    92
  introduction rule with an optional theorem declaration (\emph{thmdecl}).
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    93
  \label{fig:railroad}}
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    94
  \end{figure}
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    95
*}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    96
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    97
text {*
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    98
  For the first subtask, the syntax of the \simpleinductive{} command can be
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    99
  described by the railroad diagram in Figure~\ref{fig:railroad}. This diagram
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   100
  more or less translates directly into the parser:
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   101
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 117
diff changeset
   102
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 117
diff changeset
   103
  @{ML_chunk [display,gray] parser}
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 117
diff changeset
   104
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   105
  which we described in Section~\ref{sec:parsingspecs}. If we feed into the 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   106
  parser the string (which corresponds to our definition of @{term even} and 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   107
  @{term odd}):
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   108
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   109
  @{ML_response [display,gray]
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   110
"let
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   111
  val input = filtered_input
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   112
     (\"even and odd \" ^  
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   113
      \"where \" ^
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   114
      \"  even0[intro]: \\\"even 0\\\" \" ^ 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   115
      \"| evenS[intro]: \\\"odd n \<Longrightarrow> even (Suc n)\\\" \" ^ 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   116
      \"| oddS[intro]:  \\\"even n \<Longrightarrow> odd (Suc n)\\\"\")
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   117
in
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   118
  parse spec_parser input
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   119
end"
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   120
"((((NONE, [(even, NONE, NoSyn), (odd, NONE, NoSyn)]), []),
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   121
     [((even0,\<dots>), \"\\^E\\^Ftoken\\^Eeven 0\\^E\\^F\\^E\"),
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   122
      ((evenS,\<dots>), \"\\^E\\^Ftoken\\^Eodd n \<Longrightarrow> even (Suc n)\\^E\\^F\\^E\"),
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   123
      ((oddS,\<dots>), \"\\^E\\^Ftoken\\^Eeven n \<Longrightarrow> odd (Suc n)\\^E\\^F\\^E\")]), [])"}
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   124
*}
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   125
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   126
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   127
text {*
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   128
  then we get back a locale (in this case @{ML NONE}), the predicates (with type
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   129
  and syntax annotations), the parameters (similar as the predicates) and
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   130
  the specifications of the introduction rules. 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   131
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   132
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   133
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   134
  This is all the information we
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   135
  need for calling the package and setting up the keyword. The latter is
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   136
  done in Lines 6 and 7 in the code below.
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   137
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 117
diff changeset
   138
  @{ML_chunk [display,gray,linenos] syntax}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   139
  
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   140
  We call @{ML OuterSyntax.command} with the kind-indicator @{ML
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   141
  OuterKeyword.thy_decl} since the package does not need to open up any goal
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   142
  state (see Section~\ref{sec:newcommand}). Note that the predicates and
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   143
  parameters are at the moment only some ``naked'' variables: they have no
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   144
  type yet (even if we annotate them with types) and they are also no defined
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   145
  constants yet (which the predicates will eventually be).  In Lines 1 to 4 we
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   146
  gather the information from the parser to be processed further. The locale
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   147
  is passed as argument to the function @{ML
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   148
  Toplevel.local_theory}.\footnote{FIXME Is this already described?} The other
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   149
  arguments, i.e.~the predicates, parameters and intro rule specifications,
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   150
  are passed to the function @{ML add_inductive in SimpleInductivePackage}
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   151
  (Line 4).
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   152
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   153
  We now come to the second subtask of the package, namely transforming the 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   154
  parser output into some internal datastructures that can be processed further. 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   155
  Remember that at the moment the introduction rules are just strings, and even
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   156
  if the predicates and parameters can contain some typing annotations, they
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   157
  are not yet in any way reflected in the introduction rules. So the task of
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   158
  @{ML add_inductive in SimpleInductivePackage} is to transform the strings
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   159
  into properly typed terms. For this it can use the function 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   160
  @{ML read_specification in Specification}. This function takes some constants
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   161
  with possible typing annotations and some rule specifications and attempts to
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   162
  find a type according to the given type constraints and the type constraints
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   163
  by the surrounding (local theory). However this function is a bit
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   164
  too general for our purposes: we want that each introduction rule has only 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   165
  name (for example @{text even0} or @{text evenS}), if a name is given at all.
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   166
  The function @{ML read_specification in Specification} however allows more
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   167
  than one rule. Since it is quite convenient to rely on this function (instead of
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   168
  building your own) we just quick ly write a wrapper function that translates
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   169
  between our specific format and the general format expected by 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   170
  @{ML read_specification in Specification}. The code of this wrapper is as follows:
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   171
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   172
  @{ML_chunk [display,gray,linenos] read_specification}
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   173
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   174
  It takes a list of constants, a list of rule specifications and a local theory 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   175
  as input. Does the transformation of the rule specifications in Line 3; calls
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   176
  the function and transforms the now typed rule specifications back into our
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   177
  format and returns the type parameter and typed rule specifications. 
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   178
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   179
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   180
   @{ML_chunk [display,gray,linenos] add_inductive}
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   181
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   182
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   183
  In order to add a new inductive predicate to a theory with the help of our
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   184
  package, the user must \emph{invoke} it. For every package, there are
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   185
  essentially two different ways of invoking it, which we will refer to as
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   186
  \emph{external} and \emph{internal}. By external invocation we mean that the
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   187
  package is called from within a theory document. In this case, the
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   188
  specification of the inductive predicate, including type annotations and
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   189
  introduction rules, are given as strings by the user. Before the package can
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   190
  actually make the definition, the type and introduction rules have to be
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   191
  parsed. In contrast, internal invocation means that the package is called by
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   192
  some other package. For example, the function definition package
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   193
  calls the inductive definition package to define the
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   194
  graph of the function. However, it is not a good idea for the function
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   195
  definition package to pass the introduction rules for the function graph to
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   196
  the inductive definition package as strings. In this case, it is better to
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   197
  directly pass the rules to the package as a list of terms, which is more
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   198
  robust than handling strings that are lacking the additional structure of
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   199
  terms. These two ways of invoking the package are reflected in its ML
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   200
  programming interface, which consists of two functions:
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   201
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   202
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   203
  @{ML_chunk [display,gray] SIMPLE_INDUCTIVE_PACKAGE}
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   204
*}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   205
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   206
text {*
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   207
  (FIXME: explain Binding.binding; Attrib.binding somewhere else)
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   208
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   209
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   210
  The function for external invocation of the package is called @{ML
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   211
  add_inductive in SimpleInductivePackage}, whereas the one for internal
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   212
  invocation is called @{ML add_inductive_i in SimpleInductivePackage}. Both
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   213
  of these functions take as arguments the names and types of the inductive
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   214
  predicates, the names and types of their parameters, the actual introduction
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   215
  rules and a \emph{local theory}.  They return a local theory containing the
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   216
  definition and the induction principle as well introduction rules. 
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   217
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   218
  Note that @{ML add_inductive_i in SimpleInductivePackage} expects
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   219
  the types of the predicates and parameters to be specified using the
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   220
  datatype @{ML_type typ} of Isabelle's logical framework, whereas @{ML
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   221
  add_inductive in SimpleInductivePackage} expects them to be given as
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   222
  optional strings. If no string is given for a particular predicate or
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   223
  parameter, this means that the type should be inferred by the
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   224
  package. 
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   225
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   226
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   227
  Additional \emph{mixfix syntax} may be associated with the
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   228
  predicates and parameters as well. Note that @{ML add_inductive_i in
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   229
  SimpleInductivePackage} does not allow mixfix syntax to be associated with
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   230
  parameters, since it can only be used for parsing.\footnote{FIXME: why ist it there then?} 
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   231
  The names of the
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   232
  predicates, parameters and rules are represented by the type @{ML_type
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   233
  Binding.binding}. Strings can be turned into elements of the type @{ML_type
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   234
  Binding.binding} using the function @{ML [display] "Binding.name : string ->
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   235
  Binding.binding"} Each introduction rule is given as a tuple containing its
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   236
  name, a list of \emph{attributes} and a logical formula. Note that the type
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   237
  @{ML_type Attrib.binding} used in the list of introduction rules is just a
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   238
  shorthand for the type @{ML_type "Binding.binding * Attrib.src list"}.  The
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   239
  function @{ML add_inductive_i in SimpleInductivePackage} expects the formula
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   240
  to be specified using the datatype @{ML_type term}, whereas @{ML
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   241
  add_inductive in SimpleInductivePackage} expects it to be given as a string.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   242
  An attribute specifies additional actions and transformations that should be
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   243
  applied to a theorem, such as storing it in the rule databases used by
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   244
  automatic tactics like the simplifier. The code of the package, which will
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   245
  be described in the following section, will mostly treat attributes as a
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   246
  black box and just forward them to other functions for storing theorems in
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   247
  local theories.  The implementation of the function @{ML add_inductive in
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   248
  SimpleInductivePackage} for external invocation of the package is quite
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   249
  simple. Essentially, it just parses the introduction rules and then passes
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   250
  them on to @{ML add_inductive_i in SimpleInductivePackage}:
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   251
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   252
  @{ML_chunk [display] add_inductive}
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   253
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   254
  For parsing and type checking the introduction rules, we use the function
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   255
  
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   256
  @{ML [display] "Specification.read_specification:
76
b99fa5fa63fc adapted to changes in binding.ML
Christian Urban <urbanc@in.tum.de>
parents: 71
diff changeset
   257
  (Binding.binding * string option * mixfix) list ->  (*{variables}*)
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   258
  (Attrib.binding * string list) list list ->  (*{rules}*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   259
  local_theory ->
76
b99fa5fa63fc adapted to changes in binding.ML
Christian Urban <urbanc@in.tum.de>
parents: 71
diff changeset
   260
  (((Binding.binding * typ) * mixfix) list *
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   261
   (Attrib.binding * term list) list) *
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   262
  local_theory"}
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   263
*}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   264
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   265
text {*
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   266
  During parsing, both predicates and parameters are treated as variables, so
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   267
  the lists \verb!preds_syn! and \verb!params_syn! are just appended
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   268
  before being passed to @{ML read_specification in Specification}. Note that the format
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   269
  for rules supported by @{ML read_specification in Specification} is more general than
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   270
  what is required for our package. It allows several rules to be associated
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   271
  with one name, and the list of rules can be partitioned into several
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   272
  sublists. In order for the list \verb!intro_srcs! of introduction rules
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   273
  to be acceptable as an input for @{ML read_specification in Specification}, we first
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   274
  have to turn it into a list of singleton lists. This transformation
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   275
  has to be reversed later on by applying the function
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   276
  @{ML [display] "the_single: 'a list -> 'a"}
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   277
  to the list \verb!specs! containing the parsed introduction rules.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   278
  The function @{ML read_specification in Specification} also returns the list \verb!vars!
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   279
  of predicates and parameters that contains the inferred types as well.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   280
  This list has to be chopped into the two lists \verb!preds_syn'! and
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   281
  \verb!params_syn'! for predicates and parameters, respectively.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   282
  All variables occurring in a rule but not in the list of variables passed to
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   283
  @{ML read_specification in Specification} will be bound by a meta-level universal
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   284
  quantifier.
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   285
*}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   286
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   287
text {*
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   288
  Finally, @{ML read_specification in Specification} also returns another local theory,
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   289
  but we can safely discard it. As an example, let us look at how we can use this
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   290
  function to parse the introduction rules of the @{text trcl} predicate:
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   291
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   292
  @{ML_response [display]
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   293
"Specification.read_specification
55
0b55402ae95e Adapted to changes in binding module.
berghofe
parents: 53
diff changeset
   294
  [(Binding.name \"trcl\", NONE, NoSyn),
0b55402ae95e Adapted to changes in binding module.
berghofe
parents: 53
diff changeset
   295
   (Binding.name \"r\", SOME \"'a \<Rightarrow> 'a \<Rightarrow> bool\", NoSyn)]
0b55402ae95e Adapted to changes in binding module.
berghofe
parents: 53
diff changeset
   296
  [[((Binding.name \"base\", []), [\"trcl r x x\"])],
0b55402ae95e Adapted to changes in binding module.
berghofe
parents: 53
diff changeset
   297
   [((Binding.name \"step\", []), [\"trcl r x y \<Longrightarrow> r y z \<Longrightarrow> trcl r x z\"])]]
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   298
  @{context}"
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   299
"((\<dots>,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   300
  [(\<dots>,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   301
    [Const (\"all\", \<dots>) $ Abs (\"x\", TFree (\"'a\", \<dots>),
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   302
       Const (\"Trueprop\", \<dots>) $
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   303
         (Free (\"trcl\", \<dots>) $ Free (\"r\", \<dots>) $ Bound 0 $ Bound 0))]),
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   304
   (\<dots>,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   305
    [Const (\"all\", \<dots>) $ Abs (\"x\", TFree (\"'a\", \<dots>),
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   306
       Const (\"all\", \<dots>) $ Abs (\"y\", TFree (\"'a\", \<dots>),
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   307
         Const (\"all\", \<dots>) $ Abs (\"z\", TFree (\"'a\", \<dots>),
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   308
           Const (\"==>\", \<dots>) $
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   309
             (Const (\"Trueprop\", \<dots>) $
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 35
diff changeset
   310
               (Free (\"trcl\", \<dots>) $ Free (\"r\", \<dots>) $ Bound 2 $ Bound 1)) $
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   311
             (Const (\"==>\", \<dots>) $ \<dots> $ \<dots>))))])]),
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   312
 \<dots>)
76
b99fa5fa63fc adapted to changes in binding.ML
Christian Urban <urbanc@in.tum.de>
parents: 71
diff changeset
   313
: (((Binding.binding * typ) * mixfix) list *
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   314
   (Attrib.binding * term list) list) * local_theory"}
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   315
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   316
  In the list of variables passed to @{ML read_specification in Specification}, we have
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   317
  used the mixfix annotation @{ML NoSyn} to indicate that we do not want to associate any
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   318
  mixfix syntax with the variable. Moreover, we have only specified the type of \texttt{r},
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   319
  whereas the type of \texttt{trcl} is computed using type inference.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   320
  The local variables \texttt{x}, \texttt{y} and \texttt{z} of the introduction rules
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   321
  are turned into bound variables with the de Bruijn indices,
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   322
  whereas \texttt{trcl} and \texttt{r} remain free variables.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   323
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   324
*}
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   325
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   326
text {*
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   327
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   328
  \paragraph{Parsers for theory syntax}
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   329
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   330
  Although the function @{ML add_inductive in SimpleInductivePackage} parses terms and types, it still
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   331
  cannot be used to invoke the package directly from within a theory document.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   332
  In order to do this, we have to write another parser. Before we describe
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   333
  the process of writing parsers for theory syntax in more detail, we first
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   334
  show some examples of how we would like to use the inductive definition
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   335
  package.
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   336
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   337
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   338
  The definition of the transitive closure should look as follows:
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   339
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   340
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   341
text {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   342
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   343
  A proposition can be parsed using the function @{ML prop in OuterParse}.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   344
  Essentially, a proposition is just a string or an identifier, but using the
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   345
  specific parser function @{ML prop in OuterParse} leads to more instructive
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   346
  error messages, since the parser will complain that a proposition was expected
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   347
  when something else than a string or identifier is found.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   348
  An optional locale target specification of the form \isa{(\isacommand{in}\ $\ldots$)}
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   349
  can be parsed using @{ML opt_target in OuterParse}.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   350
  The lists of names of the predicates and parameters, together with optional
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   351
  types and syntax, are parsed using the functions @{ML "fixes" in OuterParse}
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   352
  and @{ML for_fixes in OuterParse}, respectively.
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   353
  In addition, the following function from @{ML_struct SpecParse} for parsing
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   354
  an optional theorem name and attribute, followed by a delimiter, will be useful:
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   355
  
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   356
  \begin{table}
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   357
  @{ML "opt_thm_name:
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   358
  string -> token list -> Attrib.binding * token list" in SpecParse}
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   359
  \end{table}
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   360
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   361
  We now have all the necessary tools to write the parser for our
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   362
  \isa{\isacommand{simple{\isacharunderscore}inductive}} command:
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   363
  
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   364
 
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   365
  Once all arguments of the command have been parsed, we apply the function
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   366
  @{ML add_inductive in SimpleInductivePackage}, which yields a local theory
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   367
  transformer of type @{ML_type "local_theory -> local_theory"}. Commands in
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   368
  Isabelle/Isar are realized by transition transformers of type
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   369
  @{ML_type [display] "Toplevel.transition -> Toplevel.transition"}
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   370
  We can turn a local theory transformer into a transition transformer by using
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   371
  the function
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   372
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   373
  @{ML [display] "Toplevel.local_theory : string option ->
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   374
  (local_theory -> local_theory) ->
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   375
  Toplevel.transition -> Toplevel.transition"}
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   376
 
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   377
  which, apart from the local theory transformer, takes an optional name of a locale
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   378
  to be used as a basis for the local theory. 
60
5b9c6010897b doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
Christian Urban <urbanc@in.tum.de>
parents: 55
diff changeset
   379
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   380
  (FIXME : needs to be adjusted to new parser type)
60
5b9c6010897b doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
Christian Urban <urbanc@in.tum.de>
parents: 55
diff changeset
   381
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   382
  {\it
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   383
  The whole parser for our command has type
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 88
diff changeset
   384
  @{text [display] "OuterLex.token list ->
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   385
  (Toplevel.transition -> Toplevel.transition) * OuterLex.token list"}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 88
diff changeset
   386
  which is abbreviated by @{text OuterSyntax.parser_fn}. The new command can be added
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   387
  to the system via the function
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 88
diff changeset
   388
  @{text [display] "OuterSyntax.command :
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   389
  string -> string -> OuterKeyword.T -> OuterSyntax.parser_fn -> unit"}
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   390
  which imperatively updates the parser table behind the scenes. }
60
5b9c6010897b doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
Christian Urban <urbanc@in.tum.de>
parents: 55
diff changeset
   391
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   392
  In addition to the parser, this
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   393
  function takes two strings representing the name of the command and a short description,
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   394
  as well as an element of type @{ML_type OuterKeyword.T} describing which \emph{kind} of
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   395
  command we intend to add. Since we want to add a command for declaring new concepts,
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   396
  we choose the kind @{ML "OuterKeyword.thy_decl"}. Other kinds include
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   397
  @{ML "OuterKeyword.thy_goal"}, which is similar to @{ML thy_decl in OuterKeyword},
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   398
  but requires the user to prove a goal before making the declaration, or
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   399
  @{ML "OuterKeyword.diag"}, which corresponds to a purely diagnostic command that does
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   400
  not change the context. For example, the @{ML thy_goal in OuterKeyword} kind is used
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   401
  by the \isa{\isacommand{function}} command \cite{Krauss-IJCAR06}, which requires the user
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   402
  to prove that a given set of equations is non-overlapping and covers all cases. The kind
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   403
  of the command should be chosen with care, since selecting the wrong one can cause strange
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 76
diff changeset
   404
  behaviour of the user interface, such as failure of the undo mechanism.
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   405
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   406
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   407
text {*
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   408
  Note that the @{text trcl} predicate has two different kinds of parameters: the
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   409
  first parameter @{text R} stays \emph{fixed} throughout the definition, whereas
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   410
  the second and third parameter changes in the ``recursive call''. This will
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   411
  become important later on when we deal with fixed parameters and locales. 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   412
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   413
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   414
 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   415
  The purpose of the package we show next is that the user just specifies the
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   416
  inductive predicate by stating some introduction rules and then the packages
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   417
  makes the equivalent definition and derives from it the needed properties.
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   418
*}
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   419
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   420
text {*
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   421
  From a high-level perspective the package consists of 6 subtasks:
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   422
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   423
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   424
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   425
*}
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   426
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   427
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   428
(*<*)
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   429
end
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   430
(*>*)