ProgTutorial/Base.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 02 Jun 2010 07:05:29 +0200
changeset 434 9aa8ef31c70a
parent 426 d94755882e36
child 438 d9a223c023f6
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23
1322990e4ee7 New theory for installing antiquotations.
berghofe
parents:
diff changeset
     1
theory Base
264
311830b43f8f minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 260
diff changeset
     2
imports Main LaTeXsugar
23
1322990e4ee7 New theory for installing antiquotations.
berghofe
parents:
diff changeset
     3
uses
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
     4
  ("output_tutorial.ML")
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
     5
  ("antiquote_setup.ML")
23
1322990e4ee7 New theory for installing antiquotations.
berghofe
parents:
diff changeset
     6
begin
1322990e4ee7 New theory for installing antiquotations.
berghofe
parents:
diff changeset
     7
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
     8
notation (latex output)
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
     9
  Cons ("_ # _" [66,65] 65)
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
    10
317
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    11
(* rebinding of writeln and tracing so that it can *)
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    12
(* be compared in intiquotations                   *)
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    13
ML {* 
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    14
fun writeln s = (Output.writeln s; s)
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    15
fun tracing s = (Output.tracing s; s)
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    16
*}
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    17
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    18
(* re-definition of various ML antiquotations     *)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    19
(* to have a special tag for text enclosed in ML; *)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    20
(* they also write the code into a separate file  *)
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 264
diff changeset
    21
64
9a6e5e0c4906 deleted old files and added code to give a special tag to the command ML
Christian Urban <urbanc@in.tum.de>
parents: 57
diff changeset
    22
ML {*
419
2e199c5faf76 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    23
val (filename, setup_filename) = Attrib.config_string "filename" (K "File_Code.ML")
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    24
*}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    25
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    26
setup {* setup_filename *}
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    27
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    28
ML {*
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    29
fun write_file txt thy =
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    30
let 
423
0a25b35178c3 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
    31
  val stream = Config.get_global thy filename
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    32
               |> TextIO.openAppend 
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    33
in
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    34
  TextIO.output (stream, txt); 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    35
  TextIO.flushOut stream;       (* needed ?*)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    36
  TextIO.closeOut stream
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    37
end
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    38
*}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    39
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    40
ML {*
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    41
fun write_file_ml_blk txt thy =
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    42
let
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    43
  val pre  = implode ["\n", "ML ", "{", "*", "\n"]
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    44
  val post = implode ["\n", "*", "}", "\n"]
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    45
  val _ = write_file (enclose pre post txt) thy
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    46
in
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    47
  thy
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    48
end
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    49
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    50
fun write_file_setup_blk txt thy =
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    51
let
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    52
  val pre  = implode ["\n", "setup ", "{", "*", "\n"]
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    53
  val post = implode ["\n", "*", "}", "\n"]
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    54
  val _ = write_file (enclose pre post txt) thy
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    55
in
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    56
  thy
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    57
end
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    58
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    59
fun write_file_lsetup_blk txt lthy =
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    60
let
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    61
  val pre  = implode ["\n", "local_setup ", "{", "*", "\n"]
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    62
  val post = implode ["\n", "*", "}", "\n"]
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    63
  val _ = write_file (enclose pre post txt) (ProofContext.theory_of lthy)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    64
in
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    65
  lthy
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    66
end
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    67
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    68
*}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    69
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    70
ML {*
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    71
fun open_file name thy =
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    72
let  
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    73
  val _ = tracing ("Open File: " ^ name)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    74
  val _ = TextIO.openOut name
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    75
in 
423
0a25b35178c3 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
    76
  Config.put_global filename name thy
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    77
end
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    78
*}
64
9a6e5e0c4906 deleted old files and added code to give a special tag to the command ML
Christian Urban <urbanc@in.tum.de>
parents: 57
diff changeset
    79
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    80
ML {*
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    81
fun open_file_with_prelude name txts thy =
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    82
let 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    83
  val thy' = open_file name thy 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    84
  val _ = write_file (cat_lines txts) thy'
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    85
in
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    86
  thy'
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    87
end
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    88
*}
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    89
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    90
ML {*
260
5accec94b6df updated to lates Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    91
fun propagate_env (context as Context.Proof lthy) =
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
    92
      Context.Proof (Local_Theory.map_contexts (ML_Env.inherit context) lthy)
264
311830b43f8f minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 260
diff changeset
    93
  | propagate_env context = context
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    94
260
5accec94b6df updated to lates Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    95
fun propagate_env_prf prf = Proof.map_contexts
264
311830b43f8f minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 260
diff changeset
    96
  (Context.proof_map (ML_Env.inherit (Context.Proof (Proof.context_of prf)))) prf
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    97
*}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    98
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    99
ML {*
64
9a6e5e0c4906 deleted old files and added code to give a special tag to the command ML
Christian Urban <urbanc@in.tum.de>
parents: 57
diff changeset
   100
val _ =
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   101
  Outer_Syntax.command "ML" "eval ML text within theory"
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   102
    (Keyword.tag "TutorialML" Keyword.thy_decl)
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   103
    (Parse.position Parse.text >> 
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   104
      (fn (txt, pos) =>
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   105
        Toplevel.generic_theory
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   106
         (ML_Context.exec 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   107
           (fn () => ML_Context.eval true pos txt) 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   108
              #> propagate_env #> Context.map_theory (write_file_ml_blk txt))))
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   109
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   110
val _ =
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   111
  Outer_Syntax.command "ML_prf" "ML text within proof" 
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   112
    (Keyword.tag "TutorialML" Keyword.prf_decl)
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   113
    (Parse.ML_source >> (fn (txt, pos) =>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   114
      Toplevel.proof (Proof.map_context (Context.proof_map
260
5accec94b6df updated to lates Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   115
        (ML_Context.exec (fn () => ML_Context.eval true pos txt))) #> propagate_env_prf)))
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   116
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   117
val _ =
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   118
  Outer_Syntax.command "ML_val" "diagnostic ML text" 
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   119
  (Keyword.tag "TutorialML" Keyword.diag)
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   120
    (Parse.ML_source >> IsarCmd.ml_diag true)
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   121
*}
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   122
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   123
ML {*
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   124
val _ =
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   125
  Outer_Syntax.command "setup" "ML theory setup" 
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   126
    (Keyword.tag_ml Keyword.thy_decl)
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   127
      (Parse.ML_source >> (fn (txt, pos) => 
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   128
        (Toplevel.theory (IsarCmd.global_setup (txt, pos) #> write_file_setup_blk txt))))
64
9a6e5e0c4906 deleted old files and added code to give a special tag to the command ML
Christian Urban <urbanc@in.tum.de>
parents: 57
diff changeset
   129
*}
239
b63c72776f03 replaced "warning" with "writeln"
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   130
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   131
ML {*
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   132
val _ =
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   133
  Outer_Syntax.local_theory "local_setup" "ML local theory setup" 
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   134
    (Keyword.tag_ml Keyword.thy_decl)
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   135
      (Parse.ML_source >> (fn (txt, pos) => 
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   136
         IsarCmd.local_setup (txt, pos) #> write_file_lsetup_blk txt));
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   137
*}
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   138
396
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   139
ML {*
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   140
fun P n = @{term "P::nat \<Rightarrow> bool"} $ (HOLogic.mk_number @{typ "nat"} n) 
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   141
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   142
fun rhs 1 = P 1
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   143
  | rhs n = HOLogic.mk_conj (P n, rhs (n - 1))
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   144
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   145
fun lhs 1 n = HOLogic.mk_imp (HOLogic.mk_eq (P 1, P n), rhs n)
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   146
  | lhs m n = HOLogic.mk_conj (HOLogic.mk_imp 
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   147
                 (HOLogic.mk_eq (P (m - 1), P m), rhs n), lhs (m - 1) n)
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   148
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   149
fun de_bruijn n =
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   150
  HOLogic.mk_Trueprop (HOLogic.mk_imp (lhs n n, rhs n))
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   151
*}
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   152
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   153
use "output_tutorial.ML"
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   154
use "antiquote_setup.ML"
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   155
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   156
239
b63c72776f03 replaced "warning" with "writeln"
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   157
end