ProgTutorial/Package/Ind_Interface.thy
author Norbert Schirmer <norbert.schirmer@web.de>
Tue, 14 May 2019 17:10:47 +0200
changeset 565 cecd7a941885
parent 562 daf404920ab9
child 566 6103b0eadbf2
permissions -rw-r--r--
isabelle update_cartouches -t
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
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
     2
imports Ind_Intro Simple_Inductive_Package
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
     3
keywords "simple_inductive2" :: thy_decl
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     4
begin
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     5
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
     6
section \<open>Parsing and Typing the Specification\label{sec:interface}\<close>
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
     7
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
     8
text_raw \<open>
296
fa2228a1d159 tuned figure placement
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
     9
\begin{figure}[t]
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    10
\begin{boxedminipage}{\textwidth}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    11
\begin{isabelle}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    12
\<close>
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    13
simple_inductive
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    14
  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
    15
where
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    16
  base: "trcl R x x"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    17
| 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
    18
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    19
simple_inductive
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    20
  even and odd
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    21
where
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    22
  even0: "even 0"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    23
| 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
    24
| 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
    25
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    26
simple_inductive
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    27
  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
    28
where
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    29
  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
    30
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    31
(*<*)
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    32
datatype trm =
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    33
  Var "string"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    34
| App "trm" "trm"
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    35
| Lam "string" "trm"
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
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    38
simple_inductive 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    39
  fresh :: "string \<Rightarrow> trm \<Rightarrow> bool" 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    40
where
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    41
  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
    42
| 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
    43
| 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
    44
| fresh_lam2: "\<lbrakk>a\<noteq>b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    45
text_raw \<open>
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    46
\end{isabelle}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    47
\end{boxedminipage}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    48
\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
    49
@{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
    50
@{term "fresh"}.\label{fig:specs}}
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    51
\end{figure}  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    52
\<close>
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    53
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    54
text \<open>
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    55
  To be able to write down the specifications of inductive predicates, we have
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    56
  to introduce a new command (see Section~\ref{sec:newcommand}).  As the
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    57
  keyword for the new command we chose \simpleinductive{}. Examples of
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    58
  specifications from the previous section are shown in
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    59
  Figure~\ref{fig:specs}. The syntax used in these examples more or
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    60
  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
    61
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    62
\<close>
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    63
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    64
ML %grayML\<open>val spec_parser = 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 558
diff changeset
    65
   Parse.vars -- 
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    66
   Scan.optional 
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 365
diff changeset
    67
     (Parse.$$$ "where" |--
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 365
diff changeset
    68
        Parse.!!! 
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 365
diff changeset
    69
          (Parse.enum1 "|" 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    70
             (Parse_Spec.opt_thm_name ":" -- Parse.prop))) []\<close>
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    71
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    72
text \<open>
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    73
  which we explained in Section~\ref{sec:parsingspecs}.  There is no code included 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    74
  for parsing the keyword and what is called a \emph{target}. The latter  can be given 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    75
  optionally after the keyword. The target is an ``advanced'' feature which we will 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    76
  inherit for ``free'' from the infrastructure on which we shall build the package. 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    77
  The target stands for a locale and allows us to specify
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    78
\<close>
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    79
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    80
locale rel =
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    81
  fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    82
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    83
text \<open>
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
    84
  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
    85
  locale as follows:
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    86
\<close>
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    87
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
    88
simple_inductive (in rel) 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
    89
  trcl' 
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    90
where
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    91
  base: "trcl' x x"
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    92
| 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
    93
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
    94
simple_inductive (in rel) 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
    95
  accpart'
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    96
where
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
    97
  accpartI: "(\<And>y. R y x \<Longrightarrow> accpart' y) \<Longrightarrow> accpart' x"
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    98
(*<*)ML %no\<open>fun filtered_input str = 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 558
diff changeset
    99
  filter Token.is_proper (Token.explode (Thy_Header.get_keywords' @{context}) Position.none str) 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   100
fun parse p input = Scan.finite Token.stopper (Scan.error p) input\<close>(*>*)
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   101
text \<open>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   102
  Note that in these definitions the parameter \<open>R\<close>, standing for the
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   103
  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
   104
  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
   105
  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
   106
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   107
  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
   108
  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
   109
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   110
  @{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
   111
"let
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   112
  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
   113
     (\"even and odd \" ^  
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   114
      \"where \" ^
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   115
      \"  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
   116
      \"| evenS[intro]: \\\"odd n \<Longrightarrow> even (Suc n)\\\" \" ^ 
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   117
      \"| oddS[intro]:  \\\"even n \<Longrightarrow> odd (Suc n)\\\" \")
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   118
in
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   119
  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
   120
end"
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   121
"(([(even, NONE, NoSyn), (odd, NONE, NoSyn)],
553
c53d74b34123 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 542
diff changeset
   122
     [((even0,\<dots>), \<dots>),
c53d74b34123 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 542
diff changeset
   123
      ((evenS,\<dots>), \<dots>),
c53d74b34123 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 542
diff changeset
   124
      ((oddS,\<dots>), \<dots>)]), [])"}
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   125
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   126
  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
   127
  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
   128
  need for calling the package and setting up the keyword. The latter is
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   129
  done in Lines 5 to 7 in the code below.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   130
\<close>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   131
(*<*)ML %no\<open>fun add_inductive_cmd pred_specs rule_specs lthy = lthy 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   132
   fun add_inductive pred_specs rule_specs lthy = lthy\<close>(*>*) (* FIXME: Is the dummy simple_inductive2 installed with ML, before there was ML_val*)
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   133
ML %linenosgray\<open>val specification : (local_theory -> local_theory) parser =
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   134
  spec_parser >>
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   135
    (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
   136
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 558
diff changeset
   137
val _ = Outer_Syntax.local_theory @{command_keyword "simple_inductive2"}
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   138
          "definition of simple inductive predicates"
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   139
             specification\<close>
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   140
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   141
text \<open>
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 365
diff changeset
   142
  We call @{ML_ind  local_theory in Outer_Syntax} with the kind-indicator 
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 365
diff changeset
   143
  @{ML_ind  thy_decl in Keyword} since the package does not need to open 
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   144
  up any proof (see Section~\ref{sec:newcommand}). 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   145
  The auxiliary function \<open>specification\<close> in Lines 1 to 3
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   146
  gathers the information from the parser to be processed further
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   147
  by the function \<open>add_inductive_cmd\<close>, which we describe below. 
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   148
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   149
  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
   150
  ``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
   151
  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
   152
  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
   153
  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
   154
  datastructures that can be processed further. For this we can use the
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 558
diff changeset
   155
  function @{ML_ind  read_multi_specs in Specification}. This function takes some strings
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   156
  (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
   157
  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
   158
  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
   159
  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
   160
  introduction rules. So at the heart of the function 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   161
  \<open>add_inductive_cmd\<close> is a call to @{ML read_multi_specs in Specification}.
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   162
\<close>
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   163
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   164
ML_val%grayML\<open>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
   165
let
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   166
  val ((pred_specs', rule_specs'), _) = 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 558
diff changeset
   167
         Specification.read_multi_specs pred_specs rule_specs lthy
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   168
in
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   169
  add_inductive pred_specs' rule_specs' lthy
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   170
end\<close>
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   171
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   172
text \<open>
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   173
  Once we have the input data as some internal datastructure, we call
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   174
  the function \<open>add_inductive\<close>. 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
   175
  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
   176
  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
   177
  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
   178
  the next two sections.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   179
\<close>
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 514
diff changeset
   180
(*<*)end(*>*)