Prove.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 25 Sep 2009 19:26:18 +0200
changeset 40 c8187c293587
parent 6 6a1b4c22a386
child 46 e801b929216b
permissions -rw-r--r--
slightly tuned the comment for unlam
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Prove
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports Main
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
ML {*
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
val r = ref (NONE:(unit -> term) option)
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
*}
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
ML {*
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
let
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
   fun setup_proof (txt, pos) lthy =
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
   let
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
     val trm = ML_Context.evaluate lthy true ("r", r) txt
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
   in
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
       Proof.theorem_i NONE (K I) [[(trm,[])]] lthy
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
   end;
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
   val setup_proof_parser = OuterParse.ML_source >> setup_proof
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
        
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
   val kind = OuterKeyword.thy_goal
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
in
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
   OuterSyntax.local_theory_to_proof "prove" "proving a proposition" 
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
    kind setup_proof_parser
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
end
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
*}
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
6a1b4c22a386 added the prove command
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
end