ProgTutorial/Helper/Command/Command.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 19 Nov 2009 14:11:50 +0100
changeset 394 0019ebf76e10
parent 328 c0cae24b9d46
child 422 e79563b14b0f
permissions -rw-r--r--
updated to new Isabelle
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Command
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports Main
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
     4
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
ML {*
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
let
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
     7
   val do_nothing = Scan.succeed (Local_Theory.theory I)
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
     8
   val kind = OuterKeyword.thy_decl
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
in
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
   OuterSyntax.local_theory "foobar" "description of foobar" kind do_nothing
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
end
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
*}
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
ML {*
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
let
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    16
  fun trace_prop str =
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    17
     Local_Theory.theory (fn lthy => (tracing str; lthy))
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    18
  val trace_prop_parser = OuterParse.prop >> trace_prop
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    19
  val kind = OuterKeyword.thy_decl
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    20
in
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    21
  OuterSyntax.local_theory "foobar_trace" "traces a proposition"
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    22
    kind trace_prop_parser
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    23
end
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    24
*}
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    25
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    26
ML {*
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    27
let
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    28
   fun prove_prop str lthy =
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    29
     let
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    30
       val prop = Syntax.read_prop lthy str
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    31
     in
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    32
       Proof.theorem_i NONE (K I) [[(prop,[])]] lthy
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    33
     end;
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    34
   val prove_prop_parser = OuterParse.prop >> prove_prop
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    35
   val kind = OuterKeyword.thy_goal
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    36
in
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    37
   OuterSyntax.local_theory_to_proof "foobar_goal" "proving a proposition"
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    38
     kind prove_prop_parser
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    39
end
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    40
*}
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    41
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    42
ML {*
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
    43
val r = Unsynchronized.ref (NONE:(unit -> term) option)
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    44
*}
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    45
ML{*
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    46
let
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    47
   fun after_qed thm_name thms lthy =
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    48
        Local_Theory.note (thm_name, (flat thms)) lthy |> snd
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    49
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    50
   fun setup_proof (thm_name, (txt, pos)) lthy =
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    51
   let
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    52
     val trm = ML_Context.evaluate lthy true ("r", r) txt
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    53
   in
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    54
     Proof.theorem_i NONE (after_qed thm_name) [[(trm,[])]] lthy
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    55
   end
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    56
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    57
   val parser = SpecParse.opt_thm_name ":" -- OuterParse.ML_source
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    58
 
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    59
in
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    60
   OuterSyntax.local_theory_to_proof "foobar_prove" "proving a proposition" 
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    61
     OuterKeyword.thy_goal (parser >> setup_proof)
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    62
end*}
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    63
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    64
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    65
(*
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    66
ML {*
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    67
let
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    68
   val do_nothing = Scan.succeed (Local_Theory.theory I)
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
   val kind = OuterKeyword.thy_decl
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
in
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    71
   OuterSyntax.local_theory "simple_inductive" "description of foobar" kind do_nothing
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
end
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    73
*}*)
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
end