ProgTutorial/Helper/Command/Command.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 22 Aug 2009 02:56:08 +0200
changeset 321 e450fa467e3f
parent 319 6bce4acf7f2a
child 324 4172c0743cf2
permissions -rw-r--r--
polished the commands section
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
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
   val do_nothing = Scan.succeed (LocalTheory.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 =
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    17
     LocalTheory.theory (fn lthy => (tracing str; lthy))
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 {*
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    43
val r = ref (NONE:(unit -> term) option)
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    44
*}
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    45
ML {*
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    46
let
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    47
   fun setup_proof (txt, pos) lthy =
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    48
   let
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    49
     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
    50
   in
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    51
       Proof.theorem_i NONE (K I) [[(trm,[])]] lthy
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    52
   end;
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    53
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    54
   val setup_proof_parser = OuterParse.ML_source >> setup_proof
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    55
        
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    56
   val kind = OuterKeyword.thy_goal
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    57
in
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    58
   OuterSyntax.local_theory_to_proof "foobar_prove" "proving a proposition" 
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    59
    kind setup_proof_parser
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    60
end
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    61
*}
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    62
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    63
(*
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    64
ML {*
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    65
let
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
   val do_nothing = Scan.succeed (LocalTheory.theory I)
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
   val kind = OuterKeyword.thy_decl
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
in
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    69
   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
    70
end
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    71
*}*)
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
end