ProgTutorial/Helper/Command/Command.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 20 Mar 2012 09:39:44 +0000
changeset 514 7e25716c3744
parent 451 fc074e669f9f
child 520 615762b8d8cb
permissions -rw-r--r--
updated to outer syntax / parser changes
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
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 451
diff changeset
     3
keywords "foobar" "foobar_trace" :: thy_decl and
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 451
diff changeset
     4
         "foobar_goal" "foobar_prove" :: thy_goal
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
begin
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
     6
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
ML {*
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
let
449
f952f2679a11 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
     9
   val do_nothing = Scan.succeed (Local_Theory.background_theory I)
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
    10
   val kind = Keyword.thy_decl
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
in
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 451
diff changeset
    12
   Outer_Syntax.local_theory ("foobar", kind) "description of foobar" do_nothing
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
end
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
*}
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
ML {*
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
let
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    18
  fun trace_prop str =
449
f952f2679a11 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 426
diff changeset
    19
     Local_Theory.background_theory (fn lthy => (tracing str; lthy))
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
    20
  val trace_prop_parser = Parse.prop >> trace_prop
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
    21
  val kind = Keyword.thy_decl
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    22
in
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 451
diff changeset
    23
  Outer_Syntax.local_theory ("foobar_trace", kind) "traces a proposition"
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 451
diff changeset
    24
    trace_prop_parser
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    25
end
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    26
*}
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    27
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    28
ML {*
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
   fun prove_prop str lthy =
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    31
     let
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    32
       val prop = Syntax.read_prop lthy str
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    33
     in
422
e79563b14b0f updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
    34
       Proof.theorem NONE (K I) [[(prop,[])]] lthy
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    35
     end;
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
    36
   val prove_prop_parser = Parse.prop >> prove_prop
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
    37
   val kind = Keyword.thy_goal
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    38
in
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 451
diff changeset
    39
   Outer_Syntax.local_theory_to_proof ("foobar_goal", kind) "proving a proposition"
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 451
diff changeset
    40
     prove_prop_parser
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    41
end
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    42
*}
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    43
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    44
ML {*
451
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 449
diff changeset
    45
structure Result = Proof_Data(
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 449
diff changeset
    46
  type T = unit -> term
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 449
diff changeset
    47
  fun init thy () = error "Result")
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 449
diff changeset
    48
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 449
diff changeset
    49
val result_cookie = (Result.get, Result.put, "Result.put")
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    50
*}
451
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 449
diff changeset
    51
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    52
ML{*
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    53
let
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    54
   fun after_qed thm_name thms lthy =
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
    55
        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
    56
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    57
   fun setup_proof (thm_name, (txt, pos)) lthy =
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    58
   let
451
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 449
diff changeset
    59
     val trm = Code_Runtime.value lthy result_cookie ("", txt)
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    60
   in
422
e79563b14b0f updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
    61
     Proof.theorem NONE (after_qed thm_name) [[(trm,[])]] lthy
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    62
   end
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    63
426
d94755882e36 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
    64
   val parser = Parse_Spec.opt_thm_name ":" -- Parse.ML_source
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    65
 
321
e450fa467e3f polished the commands section
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    66
in
514
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 451
diff changeset
    67
   Outer_Syntax.local_theory_to_proof ("foobar_prove", Keyword.thy_goal) "proving a proposition" 
7e25716c3744 updated to outer syntax / parser changes
Christian Urban <urbanc@in.tum.de>
parents: 451
diff changeset
    68
      (parser >> setup_proof)
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    69
end*}
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 321
diff changeset
    70
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
451
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 449
diff changeset
    73
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
end