ProgTutorial/Package/Ind_Interface.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 05 May 2009 03:21:49 +0200
changeset 244 dc95a56b1953
parent 224 647cab4a72c2
child 245 53112deda119
permissions -rw-r--r--
fixed the problem with double definition of even and odd
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
244
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
     2
imports "../Base" 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
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
     5
section {* Parsing and Typing the Specification\label{sec:interface} *}
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
     6
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
     7
text_raw {*
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
     8
\begin{figure}[p]
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
     9
\begin{boxedminipage}{\textwidth}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    10
\begin{isabelle}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    11
*}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    12
simple_inductive
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    13
  trcl :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    14
where
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    15
  base: "trcl R x x"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    16
| step: "trcl R x y \<Longrightarrow> R y z \<Longrightarrow> trcl R x z"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    17
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    18
simple_inductive
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    19
  even and odd
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    20
where
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    21
  even0: "even 0"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    22
| evenS: "odd n \<Longrightarrow> even (Suc n)"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    23
| oddS: "even n \<Longrightarrow> odd (Suc n)"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    24
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    25
simple_inductive
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    26
  accpart :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    27
where
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    28
  accpartI: "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    29
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    30
(*<*)
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    31
datatype trm =
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    32
  Var "string"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    33
| App "trm" "trm"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    34
| Lam "string" "trm"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    35
(*>*)
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    36
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    37
simple_inductive 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    38
  fresh :: "string \<Rightarrow> trm \<Rightarrow> bool" 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    39
where
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    40
  fresh_var: "a\<noteq>b \<Longrightarrow> fresh a (Var b)"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    41
| fresh_app: "\<lbrakk>fresh a t; fresh a s\<rbrakk> \<Longrightarrow> fresh a (App t s)"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    42
| fresh_lam1: "fresh a (Lam a t)"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    43
| fresh_lam2: "\<lbrakk>a\<noteq>b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    44
text_raw {*
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    45
\end{isabelle}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    46
\end{boxedminipage}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    47
\caption{Specification given by the user for the inductive predicates
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    48
@{term "trcl"}, @{term "even"} and @{term "odd"}, @{term "accpart"} and 
244
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    49
@{term "fresh"}.\label{fig:specs}}
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    50
\end{figure}  
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    51
*}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    52
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    53
text {*
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    54
  \begin{figure}[p]
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    55
  \begin{boxedminipage}{\textwidth}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    56
  \begin{isabelle}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    57
  \railnontermfont{\rmfamily\itshape}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    58
  \railterm{simpleinductive,where,for}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    59
  \railalias{simpleinductive}{\simpleinductive{}}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    60
  \railalias{where}{\isacommand{where}}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    61
  \railalias{for}{\isacommand{for}}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    62
  \begin{rail}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    63
  simpleinductive target?\\ fixes
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    64
  (where (thmdecl? prop + '|'))?
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    65
  ;
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    66
  \end{rail}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    67
  \end{isabelle}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    68
  \end{boxedminipage}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    69
  \caption{A railroad diagram describing the syntax of \simpleinductive{}. 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    70
  The \emph{target} indicates an optional locale; the \emph{fixes} are an 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    71
  \isacommand{and}-separated list of names for the inductive predicates (they
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
    72
  can also contain typing- and syntax annotations); \emph{prop} stands for a 
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    73
  introduction rule with an optional theorem declaration (\emph{thmdecl}).
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    74
  \label{fig:railroad}}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    75
  \end{figure}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    76
*}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    77
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
    78
text {* 
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
    79
  To be able to write down the specifications or inductive predicates, we have 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
    80
  to introduce
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
    81
  a new command (see Section~\ref{sec:newcommand}).  As the keyword for the
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    82
  new command we chose \simpleinductive{}. Examples of specifications we expect
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    83
  the user gives for the inductive predicates from the previous section are
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
    84
  shown in Figure~\ref{fig:specs}. The general syntax we will parse is specified
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
    85
  in the railroad diagram shown in Figure~\ref{fig:railroad}. This diagram more 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
    86
  or less translates directly into the parser:
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    87
*}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    88
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    89
ML{*val spec_parser = 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    90
   OuterParse.fixes -- 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    91
   Scan.optional 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    92
     (OuterParse.$$$ "where" |--
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    93
        OuterParse.!!! 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    94
          (OuterParse.enum1 "|" 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    95
             (SpecParse.opt_thm_name ":" -- OuterParse.prop))) []*}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    96
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    97
text {*
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
    98
  which we explained in Section~\ref{sec:parsingspecs}.  However, if you look
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
    99
  closely, there is no code for parsing the target given optionally after the
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   100
  keyword. This is an ``advanced'' feature which we will inherit for ``free''
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   101
  from the infrastructure on which we shall build the package. The target
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   102
  stands for a locale and allows us to specify
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   103
*}
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   104
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   105
locale rel =
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   106
  fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   107
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   108
text {*
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 215
diff changeset
   109
  and then define the transitive closure and the accessible part of this 
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 215
diff changeset
   110
  locale as follows:
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   111
*}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   112
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   113
simple_inductive (in rel) 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   114
  trcl' 
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   115
where
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   116
  base: "trcl' x x"
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   117
| 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
   118
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   119
simple_inductive (in rel) 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   120
  accpart'
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   121
where
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   122
  accpartI: "(\<And>y. R y x \<Longrightarrow> accpart' y) \<Longrightarrow> accpart' x"
244
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   123
(*<*)ML %no{*fun filtered_input str = 
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   124
  filter OuterLex.is_proper (OuterSyntax.scan Position.none str) 
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   125
fun parse p input = Scan.finite OuterLex.stopper (Scan.error p) input*}(*>*)
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
   126
text {*
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   127
  Note that in these definitions the parameter @{text R}, standing for the
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   128
  relation, is left implicit. For the moment we will ignore this kind
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   129
  of implicit parameters and rely on the fact that the infrastructure will
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   130
  deal with them. Later, however, we will come back to them.
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   131
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   132
  If we feed into the parser the string that corresponds to our definition 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   133
  of @{term even} and @{term odd}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   134
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   135
  @{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
   136
"let
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   137
  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
   138
     (\"even and odd \" ^  
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   139
      \"where \" ^
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   140
      \"  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
   141
      \"| 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
   142
      \"| 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
   143
in
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   144
  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
   145
end"
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   146
"(([(even, NONE, NoSyn), (odd, NONE, NoSyn)],
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   147
     [((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
   148
      ((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
   149
      ((oddS,\<dots>), \"\\^E\\^Ftoken\\^Eeven n \<Longrightarrow> odd (Suc n)\\^E\\^F\\^E\")]), [])"}
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   150
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   151
  then we get back the specifications of the predicates (with type and syntax annotations), 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   152
  and specifications of the introduction rules. This is all the information we
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   153
  need for calling the package and setting up the keyword. The latter is
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   154
  done in Lines 5 to 7 in the code below.
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   155
*}
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   156
244
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   157
(*<*)
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   158
ML %linenos 
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   159
{* fun add_inductive_cmd pred_specs rule_specs lthy = lthy 
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   160
   fun add_inductive pred_specs rule_specs lthy = lthy *}
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   161
(*>*)
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   162
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   163
ML_val %linenosgray{*val specification =
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   164
  spec_parser >>
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   165
    (fn (pred_specs, rule_specs) => add_inductive_cmd pred_specs rule_specs)
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   166
244
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   167
val _ = OuterSyntax.local_theory "simple_inductive2" 
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   168
          "definition of simple inductive predicates"
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   169
             OuterKeyword.thy_decl specification*}
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   170
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   171
text {*
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   172
  We call @{ML local_theory in OuterSyntax} with the kind-indicator 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   173
  @{ML thy_decl in OuterKeyword} since the package does not need to open 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   174
  up any proof (see Section~\ref{sec:newcommand}). 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   175
  The auxiliary function @{text specification} in Lines 1 to 3
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   176
  gathers the information from the parser to be processed further
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   177
  by the function @{text "add_inductive_cmd"}, which we describe below. 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   178
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   179
  Note that the predicates when they come out of the parser are just some
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   180
  ``naked'' strings: they have no type yet (even if we annotate them with
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   181
  types) and they are also not defined constants yet (which the predicates 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   182
  eventually will be).  Also the introduction rules are just strings. What we have
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   183
  to do first is to transform the parser's output into some internal
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   184
  datastructures that can be processed further. For this we can use the
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   185
  function @{ML read_spec in Specification}. This function takes some strings
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   186
  (with possible typing annotations) and some rule specifications, and attempts
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   187
  to find a typing according to the given type constraints given by the 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   188
  user and the type constraints by the ``ambient'' theory. It returns 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   189
  the type for the predicates and also returns typed terms for the
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   190
  introduction rules. So at the heart of the function 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   191
  @{text "add_inductive_cmd"} is a call to @{ML read_spec in Specification}.
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   192
*}
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   193
244
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   194
ML_val{*fun add_inductive_cmd pred_specs rule_specs lthy =
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   195
let
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   196
  val ((pred_specs', rule_specs'), _) = 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   197
         Specification.read_spec pred_specs rule_specs lthy
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   198
in
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   199
  add_inductive pred_specs' rule_specs' lthy
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   200
end*}
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   201
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   202
text {*
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   203
  Once we have the input data as some internal datastructure, we call
244
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   204
  the function @{text add_inductive}. This function does the  heavy duty 
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   205
  lifting in the package: it generates definitions for the 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   206
  predicates and derives from them corresponding induction principles and 
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   207
  introduction rules. The description of this function will span over
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   208
  the next two sections.
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   209
*}
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   210
(*<*)end(*>*)