(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory Local_Method_Tests imports Eisbach_Methods begin text \ @{command supply_local_method} allows adding named aliases for methods to the local proof context, for when duplicating methods would be too difficult to maintain or debug but refactoring the proof would take too long. Usage: supply_local_method my_name = supply_local_method my_name_1 = my_name_2 = ... Where uses the same syntax as the input to @{command apply}. \ text \ @{method local_method} allows using a previously supplied local method to the current proof state. Usage: apply (local_method name) \ experiment begin section \Basic Examples\ lemma "A \ B \ A \ B" supply_local_method my_simp = simp apply (local_method my_simp) done text \Uses standard @{command apply} parser, allowing all the normal combinators\ lemma "A \ B \ A \ B" supply_local_method my_slightly_complicated_method = (rule conjI, assumption+) apply (local_method my_slightly_complicated_method) done text \Can supply multiple methods\ lemma "A \ B \ A \ B" supply_local_method my_rule = (rule conjI) my_asm = assumption+ apply (local_method my_rule) apply (local_method my_asm) done section \Failure Modes\ text \Fails on non-existent arguments\ lemma "A \ B \ A \ B" supply_local_method foo = simp apply (fails \local_method bar\) oops text \Doesn't persist between subgoals\ lemma "A \ B \ A \ B" apply (rule conjI) subgoal supply_local_method my_simp = simp apply (local_method my_simp) done apply (fails \local_method my_simp\) oops text \Does see top-level method definitions\ method foo = simp lemma "A \ B \ A \ B" supply_local_method magic = foo apply (local_method magic) done end end