| author | wenzelm |
| Mon, 27 Feb 2012 19:54:50 +0100 | |
| changeset 46716 | c45a4427db39 |
| parent 42795 | 66fcc9882784 |
| child 48891 | c0eafbd55de3 |
| permissions | -rw-r--r-- |
(* Title: HOL/UNITY/UNITY_Main.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 2003 University of Cambridge *) header{*Comprehensive UNITY Theory*} theory UNITY_Main imports Detects PPROD Follows ProgressSets uses "UNITY_tactics.ML" begin method_setup safety = {* Scan.succeed (SIMPLE_METHOD' o constrains_tac) *} "for proving safety properties" method_setup ensures_tac = {* Args.goal_spec -- Scan.lift Args.name_source >> (fn (quant, s) => fn ctxt => SIMPLE_METHOD'' quant (ensures_tac ctxt s)) *} "for proving progress properties" setup {* Simplifier.map_simpset_global (fn ss => ss addsimps (make_o_equivs @{thm fst_o_funPair} @ make_o_equivs @{thm snd_o_funPair}) addsimps (make_o_equivs @{thm fst_o_lift_map} @ make_o_equivs @{thm snd_o_lift_map})) *} end