| author | haftmann |
| Sat, 31 Dec 2016 08:12:31 +0100 | |
| changeset 64715 | 33d5fa0ce6e5 |
| parent 63527 | 59eff6e56d81 |
| child 67899 | 730fa992da38 |
| permissions | -rw-r--r-- |
(* Title: HOL/Eisbach/Eisbach.thy Author: Daniel Matichuk, NICTA/UNSW Main entry point for Eisbach proof method language. *) theory Eisbach imports Pure 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" method solves methods m = (m; fail) end