ProgTutorial/Base.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 20 Mar 2012 09:39:44 +0000
changeset 514 7e25716c3744
parent 488 780100cd4060
child 515 4b105b97069c
permissions -rw-r--r--
updated to outer syntax / parser changes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23
1322990e4ee7 New theory for installing antiquotations.
berghofe
parents:
diff changeset
     1
theory Base
459
4532577b61e0 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
     2
imports Main 
4532577b61e0 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
     3
        "~~/src/HOL/Library/LaTeXsugar"
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
     4
keywords "ML" "setup" "local_setup" :: thy_decl and
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
     5
         "ML_prf" :: prf_decl and
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
     6
         "ML_val" :: diag
23
1322990e4ee7 New theory for installing antiquotations.
berghofe
parents:
diff changeset
     7
uses
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
     8
  ("output_tutorial.ML")
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
     9
  ("antiquote_setup.ML")
23
1322990e4ee7 New theory for installing antiquotations.
berghofe
parents:
diff changeset
    10
begin
1322990e4ee7 New theory for installing antiquotations.
berghofe
parents:
diff changeset
    11
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
    12
notation (latex output)
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
    13
  Cons ("_ # _" [66,65] 65)
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
    14
317
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    15
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    16
(* 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
    17
(* 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
    18
(* 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
    19
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
    20
ML {*
462
1d1e795bc3ad updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 459
diff changeset
    21
val filename = Attrib.setup_config_string @{binding "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
    22
*}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    23
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    24
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    25
ML {*
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    26
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
    27
let 
423
0a25b35178c3 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
    28
  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
    29
               |> TextIO.openAppend 
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    30
in
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    31
  TextIO.output (stream, txt); 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    32
  TextIO.flushOut stream;       (* needed ?*)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    33
  TextIO.closeOut stream
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    34
end
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    35
*}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    36
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    37
ML {*
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    38
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
    39
let
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    40
  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
    41
  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
    42
  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
    43
in
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    44
  thy
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    45
end
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    46
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    47
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
    48
let
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    49
  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
    50
  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
    51
  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
    52
in
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    53
  thy
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    54
end
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    55
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    56
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
    57
let
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    58
  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
    59
  val post = implode ["\n", "*", "}", "\n"]
475
25371f74c768 updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
    60
  val _ = write_file (enclose pre post txt) (Proof_Context.theory_of lthy)
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    61
in
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    62
  lthy
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    63
end
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    64
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    65
*}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    66
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
    67
ML {*
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    68
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
    69
let  
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    70
  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
    71
  val _ = TextIO.openOut name
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    72
in 
423
0a25b35178c3 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 419
diff changeset
    73
  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
    74
end
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    75
*}
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
    76
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    77
ML {*
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    78
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
    79
let 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    80
  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
    81
  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
    82
in
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    83
  thy'
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    84
end
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    85
*}
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    86
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    87
ML {*
260
5accec94b6df updated to lates Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    88
fun propagate_env (context as Context.Proof lthy) =
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
    89
      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
    90
  | propagate_env context = context
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    91
260
5accec94b6df updated to lates Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    92
fun propagate_env_prf prf = Proof.map_contexts
264
311830b43f8f minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 260
diff changeset
    93
  (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
    94
*}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    95
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    96
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
    97
val _ =
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
    98
  Outer_Syntax.command ("ML", Keyword.tag "TutorialML" Keyword.thy_decl) "eval ML text within theory"
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
    99
    (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
   100
      (fn (txt, pos) =>
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   101
        Toplevel.generic_theory
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   102
         (ML_Context.exec 
438
d9a223c023f6 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   103
           (fn () => ML_Context.eval_text true pos txt) 
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   104
              #> propagate_env #> Context.map_theory (write_file_ml_blk txt))))
438
d9a223c023f6 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   105
*}
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 475
diff changeset
   106
438
d9a223c023f6 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   107
ML {*
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   108
val _ =
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   109
  Outer_Syntax.command ("ML_prf", Keyword.tag "TutorialMLprf" Keyword.prf_decl) "ML text within proof" 
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   110
    (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
   111
      Toplevel.proof (Proof.map_context (Context.proof_map
438
d9a223c023f6 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   112
        (ML_Context.exec (fn () => ML_Context.eval_text true pos txt))) #> propagate_env_prf)))
d9a223c023f6 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   113
*}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   114
438
d9a223c023f6 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   115
ML {*
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   116
val _ =
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   117
  Outer_Syntax.command ("ML_val", Keyword.tag "TutorialML" Keyword.diag) "diagnostic ML text" 
438
d9a223c023f6 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   118
    (Parse.ML_source >> Isar_Cmd.ml_diag true)
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   119
*}
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   120
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   121
ML {*
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   122
val _ =
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   123
  Outer_Syntax.command ("setup", Keyword.tag_ml Keyword.thy_decl) "ML theory setup" 
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   124
      (Parse.ML_source >> (fn (txt, pos) => 
438
d9a223c023f6 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   125
        (Toplevel.theory (Isar_Cmd.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
   126
*}
239
b63c72776f03 replaced "warning" with "writeln"
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   127
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   128
ML {*
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   129
val _ =
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   130
  Outer_Syntax.local_theory ("local_setup", Keyword.tag_ml Keyword.thy_decl) "ML local theory setup" 
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   131
      (Parse.ML_source >> (fn (txt, pos) => 
438
d9a223c023f6 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
   132
         Isar_Cmd.local_setup (txt, pos) #> write_file_lsetup_blk txt));
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   133
*}
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   134
396
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   135
ML {*
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   136
fun P n = @{term "P::nat \<Rightarrow> bool"} $ (HOLogic.mk_number @{typ "nat"} n) 
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   137
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   138
fun rhs 1 = P 1
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   139
  | rhs n = HOLogic.mk_conj (P n, rhs (n - 1))
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   140
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   141
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
   142
  | lhs m n = HOLogic.mk_conj (HOLogic.mk_imp 
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   143
                 (HOLogic.mk_eq (P (m - 1), P m), rhs n), lhs (m - 1) n)
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   144
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   145
fun de_bruijn n =
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   146
  HOLogic.mk_Trueprop (HOLogic.mk_imp (lhs n n, rhs n))
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   147
*}
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   148
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   149
use "output_tutorial.ML"
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   150
use "antiquote_setup.ML"
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   151
471
f65b9f14d5de updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   152
setup {* OutputTutorial.setup *}
f65b9f14d5de updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   153
setup {* AntiquoteSetup.setup *}
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   154
239
b63c72776f03 replaced "warning" with "writeln"
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   155
end