Prove.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Tue, 27 Oct 2009 18:05:45 +0100
changeset 211 80859cc80332
parent 183 6acf9e001038
child 319 0ae9d9e66cb7
permissions -rw-r--r--
Manual conversion of equality to equivalence allows lifting append_assoc.

theory Prove
imports Main 
begin

ML {*
val r = Unsynchronized.ref (NONE:(unit -> term) option)
*}

ML {*
let
  fun after_qed thm_name thms lthy =
       LocalTheory.note Thm.theoremK (thm_name, (flat thms)) lthy |> snd
    
  fun setup_proof (name_spec, (txt, pos)) lthy =
  let
    val trm = ML_Context.evaluate lthy true ("r", r) txt
  in
    Proof.theorem_i NONE (after_qed name_spec) [[(trm,[])]] lthy
  end

  val parser = SpecParse.opt_thm_name ":" -- OuterParse.ML_source
in
  OuterSyntax.local_theory_to_proof "prove" "proving a proposition" 
    OuterKeyword.thy_goal (parser >> setup_proof)
end
*}

end