src/HOL/Eisbach/Eisbach.thy
author paulson <lp15@cam.ac.uk>
Mon, 07 Dec 2015 16:44:26 +0000
changeset 61806 d2e62ae01cd8
parent 60623 be39fe6c5620
child 61853 fb7756087101
permissions -rw-r--r--
Cauchy's integral formula for circles. Starting to fix eventually_mono.

(*  Title:      HOL/Eisbach/Eisbach.thy
    Author:     Daniel Matichuk, NICTA/UNSW

Main entry point for Eisbach proof method language.
*)

theory Eisbach
imports Main
keywords
  "method" :: thy_decl and
  "conclusion"
  "declares"
  "methods"
  "\<bar>" "\<Rightarrow>"
  "uses"
begin

ML_file "parse_tools.ML"
ML_file "method_closure.ML"
ML_file "eisbach_rule_insts.ML"
ML_file "match_method.ML"
ML_file "eisbach_antiquotations.ML"

(* FIXME reform Isabelle/Pure attributes to make this work by default *)
setup \<open>
  fold (Method_Closure.wrap_attribute {handle_all_errs = true, declaration = true})
    [@{binding intro}, @{binding elim}, @{binding dest}, @{binding simp}] #>
  fold (Method_Closure.wrap_attribute {handle_all_errs = false, declaration = false})
    [@{binding THEN}, @{binding OF}, @{binding rotated}, @{binding simplified}]
\<close>

method solves methods m = (m; fail)

end