ProgTutorial/Advanced.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 02 Dec 2009 02:34:09 +0100
changeset 406 f399b6855546
parent 401 36d61044f9bf
child 408 ef048892d0f0
permissions -rw-r--r--
implemented suggestion from Sascha
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
395
2c392f61f400 spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
     1
theory Advanced
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports Base FirstSteps
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
     5
(*<*)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
     6
setup{*
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
     7
open_file_with_prelude 
395
2c392f61f400 spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
     8
  "Advanced_Code.thy"
401
36d61044f9bf updated to new Isabelle and clarified Skip_Proof
Christian Urban <urbanc@in.tum.de>
parents: 400
diff changeset
     9
  ["theory Advanced", "imports Base FirstSteps", "begin"]
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
    10
*}
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
    11
(*>*)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
    12
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
    13
395
2c392f61f400 spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
    14
chapter {* Advanced Isabelle *}
381
97518188ef0e added more to the unification section
Christian Urban <urbanc@in.tum.de>
parents: 380
diff changeset
    15
97518188ef0e added more to the unification section
Christian Urban <urbanc@in.tum.de>
parents: 380
diff changeset
    16
text {*
400
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    17
  While terms, types and theorems are the most basic data structures in
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    18
  Isabelle, there are a number of layers built on top of them. Most of these
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    19
  layers are concerned with storing and manipulating data. Handling
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    20
  them properly is an essential skill for Isabelle programming. The most basic
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    21
  layer are theories. They contain global data and can be seen as the
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    22
  ``long-term memory'' of Isabelle. There is usually only one theory
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    23
  active at each moment. Proof contexts and local theories, on the
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    24
  other hand, store local data for a task at hand. They act like the
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    25
  ``short-term memory'' and there can be many of them that are active
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    26
  in parallel. 
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    29
section {* Theories\label{sec:theories} (TBD) *}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    30
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    31
text {*
401
36d61044f9bf updated to new Isabelle and clarified Skip_Proof
Christian Urban <urbanc@in.tum.de>
parents: 400
diff changeset
    32
  Theories, as said above, are the most basic layer in Isabelle. They contain
400
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    33
  definitions, syntax declarations, axioms, proofs etc. If a definition is
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    34
  stated, it must be stored in a theory in order to be usable later
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    35
  on. Similar with proofs: once a proof is finished, the proved theorem needs
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    36
  to be stored in the theorem database of the theory in order to be
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    37
  usable. All relevant data of a theory can be queried as follows.
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    38
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    39
  \begin{isabelle}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    40
  \isacommand{print\_theory}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    41
  @{text "> names: Pure Code_Generator HOL \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    42
  @{text "> classes: Inf < type \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    43
  @{text "> default sort: type"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    44
  @{text "> syntactic types: #prop \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    45
  @{text "> logical types: 'a \<times> 'b \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    46
  @{text "> type arities: * :: (random, random) random \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    47
  @{text "> logical constants: == :: 'a \<Rightarrow> 'a \<Rightarrow> prop \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    48
  @{text "> abbreviations: \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    49
  @{text "> axioms: \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    50
  @{text "> oracles: \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    51
  @{text "> definitions: \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    52
  @{text "> theorems: \<dots>"}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    53
  \end{isabelle}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    54
400
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    55
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    56
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    57
  \begin{center}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    58
  \begin{tikzpicture}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    59
  \node[top color=white, bottom color=gray!30, draw=black!100, drop shadow] {A};
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    60
  \end{tikzpicture}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    61
  \end{center}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    62
392
47e5b71c7f6c modified the typ-pretty-printer and added parser exercise
Christian Urban <urbanc@in.tum.de>
parents: 389
diff changeset
    63
  \footnote{\bf FIXME: list append in merge operations can cause 
47e5b71c7f6c modified the typ-pretty-printer and added parser exercise
Christian Urban <urbanc@in.tum.de>
parents: 389
diff changeset
    64
  exponential blowups.}
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    65
*}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    66
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    67
section {* Setups (TBD) *}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    69
text {*
361
Christian Urban <urbanc@in.tum.de>
parents: 360
diff changeset
    70
  @{ML Sign.declare_const}
Christian Urban <urbanc@in.tum.de>
parents: 360
diff changeset
    71
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    72
  In the previous section we used \isacommand{setup} in order to make
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    73
  a theorem attribute known to Isabelle. What happens behind the scenes
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    74
  is that \isacommand{setup} expects a function of type 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    75
  @{ML_type "theory -> theory"}: the input theory is the current theory and the 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    76
  output the theory where the theory attribute has been stored.
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    77
  
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    78
  This is a fundamental principle in Isabelle. A similar situation occurs 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    79
  for example with declaring constants. The function that declares a 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    80
  constant on the ML-level is @{ML_ind  add_consts_i in Sign}. 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    81
  If you write\footnote{Recall that ML-code  needs to be 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    82
  enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    83
*}  
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    84
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    85
ML{*Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] @{theory} *}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    86
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    87
text {*
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    88
  for declaring the constant @{text "BAR"} with type @{typ nat} and 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    89
  run the code, then you indeed obtain a theory as result. But if you 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    90
  query the constant on the Isabelle level using the command \isacommand{term}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    91
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    92
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    93
  \isacommand{term}~@{text [quotes] "BAR"}\\
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    94
  @{text "> \"BAR\" :: \"'a\""}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    95
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    96
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    97
  you do not obtain a constant of type @{typ nat}, but a free variable (printed in 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    98
  blue) of polymorphic type. The problem is that the ML-expression above did 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    99
  not register the declaration with the current theory. This is what the command
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   100
  \isacommand{setup} is for. The constant is properly declared with
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   101
*}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   102
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   103
setup %gray {* Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] *}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   104
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   105
text {* 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   106
  Now 
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   107
  
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   108
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   109
  \isacommand{term}~@{text [quotes] "BAR"}\\
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   110
  @{text "> \"BAR\" :: \"nat\""}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   111
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   112
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   113
  returns a (black) constant with the type @{typ nat}.
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   114
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   115
  A similar command is \isacommand{local\_setup}, which expects a function
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   116
  of type @{ML_type "local_theory -> local_theory"}. Later on we will also
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   117
  use the commands \isacommand{method\_setup} for installing methods in the
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   118
  current theory and \isacommand{simproc\_setup} for adding new simprocs to
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   119
  the current simpset.
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   120
*}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
341
62dea749d5ed more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents: 340
diff changeset
   122
section {* Contexts (TBD) *}
62dea749d5ed more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents: 340
diff changeset
   123
62dea749d5ed more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents: 340
diff changeset
   124
section {* Local Theories (TBD) *}
62dea749d5ed more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents: 340
diff changeset
   125
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   126
text {*
400
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   127
  In contrast to an ordinary theory, which simply consists of a type
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   128
  signature, as well as tables for constants, axioms and theorems, a local
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   129
  theory contains additional context information, such as locally fixed
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   130
  variables and local assumptions that may be used by the package. The type
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   131
  @{ML_type local_theory} is identical to the type of \emph{proof contexts}
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   132
  @{ML_type "Proof.context"}, although not every proof context constitutes a
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   133
  valid local theory.
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   134
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   135
  @{ML "Context.>> o Context.map_theory"}
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   136
  @{ML_ind "Local_Theory.declaration"}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   137
*}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
(*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
setup {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
 Sign.add_consts_i [(Binding"bar", @{typ "nat"},NoSyn)] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
lemma "bar = (1::nat)"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  oops
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
setup {*   
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
  Sign.add_consts_i [("foo", @{typ "nat"},NoSyn)] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
 #> PureThy.add_defs false [((@{binding "foo_def"}, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
       Logic.mk_equals (Const ("FirstSteps.foo", @{typ "nat"}), @{term "1::nat"})), [])] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
 #> snd
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
*)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
(*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
lemma "foo = (1::nat)"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  apply(simp add: foo_def)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
  done
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
thm foo_def
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
*)
336
a12bb28fe2bd polished on the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 335
diff changeset
   160
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   161
section {* Morphisms (TBD) *}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   162
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   163
text {*
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   164
  Morphisms are arbitrary transformations over terms, types, theorems and bindings.
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   165
  They can be constructed using the function @{ML_ind morphism in Morphism},
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   166
  which expects a record with functions of type
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   167
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   168
  \begin{isabelle}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   169
  \begin{tabular}{rl}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   170
  @{text "binding:"} & @{text "binding -> binding"}\\
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   171
  @{text "typ:"}     & @{text "typ -> typ"}\\
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   172
  @{text "term:"}    & @{text "term -> term"}\\
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   173
  @{text "fact:"}    & @{text "thm list -> thm list"}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   174
  \end{tabular}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   175
  \end{isabelle}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   176
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   177
  The simplest morphism is the  @{ML_ind identity in Morphism}-morphism defined as
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   178
*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   179
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   180
ML{*val identity = Morphism.morphism {binding = I, typ = I, term = I, fact = I}*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   181
  
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   182
text {*
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   183
  Morphisms can be composed with the function @{ML_ind "$>" in Morphism}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   184
*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   185
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   186
ML{*fun trm_phi (Free (x, T)) = Var ((x, 0), T) 
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   187
  | trm_phi (Abs (x, T, t)) = Abs (x, T, trm_phi t)
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   188
  | trm_phi (t $ s) = (trm_phi t) $ (trm_phi s)
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   189
  | trm_phi t = t*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   190
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   191
ML{*val phi = Morphism.term_morphism trm_phi*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   192
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   193
ML{*Morphism.term phi @{term "P x y"}*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   194
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   195
text {*
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   196
  @{ML_ind term_morphism in Morphism}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   197
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   198
  @{ML_ind term in Morphism},
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   199
  @{ML_ind thm in Morphism}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   200
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   201
  \begin{readmore}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   202
  Morphisms are implemented in the file @{ML_file "Pure/morphism.ML"}.
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   203
  \end{readmore}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   204
*}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
section {* Misc (TBD) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
ML {*Datatype.get_info @{theory} "List.list"*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   210
text {* 
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   211
FIXME: association lists:
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   212
@{ML_file "Pure/General/alist.ML"}
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   213
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   214
FIXME: calling the ML-compiler
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   215
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   216
*}
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   217
335
163ac0662211 reorganised the certified terms section; tuned
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
   218
section {* Managing Name Spaces (TBD) *}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
360
Christian Urban <urbanc@in.tum.de>
parents: 359
diff changeset
   220
ML {* Sign.intern_type @{theory} "list" *}
Christian Urban <urbanc@in.tum.de>
parents: 359
diff changeset
   221
ML {* Sign.intern_const @{theory} "prod_fun" *}
Christian Urban <urbanc@in.tum.de>
parents: 359
diff changeset
   222
387
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   223
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   224
text {* 
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   225
  @{ML_ind "Binding.str_of"} returns the string with markup;
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   226
  @{ML_ind "Binding.name_of"} returns the string without markup
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   227
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   228
  @{ML_ind "Binding.conceal"} 
387
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   229
*}
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   230
388
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   231
section {* Concurrency (TBD) *}
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   232
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   233
text {*
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   234
  @{ML_ind prove_future in Goal}
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   235
  @{ML_ind future_result in Goal}
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   236
  @{ML_ind fork_pri in Future}
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   237
*}
387
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   238
396
Christian Urban <urbanc@in.tum.de>
parents: 395
diff changeset
   239
section {* Parse and Print Translations (TBD) *}
Christian Urban <urbanc@in.tum.de>
parents: 395
diff changeset
   240
349
9e374cd891e1 updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 348
diff changeset
   241
section {* Summary *}
9e374cd891e1 updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 348
diff changeset
   242
9e374cd891e1 updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 348
diff changeset
   243
text {*
395
2c392f61f400 spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   244
  TBD
349
9e374cd891e1 updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 348
diff changeset
   245
*}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
end