(* * Copyright 2018, Data61, CSIRO * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(DATA61_BSD) *) theory Trace_Schematic_Insts_Test imports Lib.Trace_Schematic_Insts begin text \ Trace the schematic variables and types that a method instantiates. This only works for the variables already in the proof goal; new variables introduced by the traced method are not tracked. \ experiment begin section \Examples\ text \Schematic variables\ lemma "\ \x. P x \ \ P x" apply (drule spec) \ \introduces schematic var "?x"\ apply (trace_schematic_insts \assumption\) done definition foo :: "'a \ bool" where "foo x = True" lemma fooI1: "foo 0 \ foo x" by (simp add: foo_def) lemma fooI2: "foo x \ foo 0" by (simp add: foo_def) lemma fooI2': "foo x \ foo (0 :: nat)" by (erule fooI2) text \Schematic type variables\ lemma "foo x \ foo y" apply (rule fooI1) \ \introduces schematic type "0 :: ?'a"\ apply (trace_schematic_insts \erule fooI2'\) done text \When backtracking, every recursive invocation is traced\ lemma "\ \x. Q x \ R x; \x. P x \ Q x; P x; P y \ R x \ \ R x" apply (drule spec) apply (drule spec) text \For more clarity, methods can be named\ apply (trace_schematic_insts impE1 \erule impE\, trace_schematic_insts impE2 \erule impE\, (trace_schematic_insts "try assumption" \assumption\)+; fail) done text \Interactive example\ ML \ fun trace_resolve_tac ctxt = Trace_Schematic_Insts.trace_schematic_insts_tac ctxt (Trace_Schematic_Insts.default_rule_report ctxt "demo title") (fn t => resolve_tac ctxt [t]) \ lemma assumes Pf: "\f (x :: nat). Q f \ P x \ P (f x)" assumes Q: "Q ((*) 2)" assumes P: "P a" shows "\x. P (x + a + a)" apply (tactic \trace_resolve_tac @{context} @{thm exI} 1\) apply (trace_schematic_insts \subst add_0\) \\ This picks *some* instantiation of `f` in @{thm Pf}. The first one is @{term "\a. a"}, which isn't what we want. \ apply (tactic \trace_resolve_tac @{context} @{thm Pf} 1\) \\ This picks the *next* instantiation of `f`, in this case @{term "\a. a + a"} Notice that the reporting callback gets called with the new instantiations! \ back apply (tactic \ Trace_Schematic_Insts.trace_schematic_insts_tac @{context} (Trace_Schematic_Insts.default_rule_report @{context} "demo title") (fn t => EqSubst.eqsubst_tac @{context} [0] [t]) @{thm mult_2[symmetric]} 1 \) apply (tactic \trace_resolve_tac @{context} @{thm Q} 1\) apply (tactic \trace_resolve_tac @{context} @{thm P} 1\) done section \Tests\ ML \ fun trace_schematic_assert ctxt test_name tac expected_vars expected_tvars = let fun skip_dummy_state tac = fn st => case Thm.prop_of st of Const (@{const_name Pure.prop}, _) $ (Const (@{const_name Pure.term}, _) $ Const (@{const_name Pure.dummy_pattern}, _)) => Seq.succeed st | _ => tac st fun check insts = if expected_vars = #terms insts andalso expected_tvars = #typs insts then () else error ("Trace_Schematic_Insts failed test: " ^ test_name) in Method.NO_CONTEXT_TACTIC ctxt (Trace_Schematic_Insts.trace_schematic_insts (SIMPLE_METHOD tac) check []) |> skip_dummy_state end \ text \Schematic variables\ lemma "\ \x. P x \ \ P x" apply (drule spec) apply (tactic \let val alpha = TFree ("'a", @{sort type}) val expected_vars = [(Var (("x", 0), alpha), Free ("x", alpha))] val expected_tvars = [] in trace_schematic_assert @{context} "basic Var test" (assume_tac @{context} 1) expected_vars expected_tvars end\) done text \Schematic type variables\ lemma "foo x \ foo y" apply (rule fooI1) apply (tactic \let val expected_vars = [] val expected_tvars = [(TVar (("'a", 0), @{sort zero}), @{typ nat})] in trace_schematic_assert @{context} "basic TVar test" (eresolve_tac @{context} @{thms fooI2'} 1) expected_vars expected_tvars end\) done ML \ fun trace_schematic_resolve_tac_assert ctxt test_name thm expected_rule_insts expected_proof_insts = let fun check rule_insts proof_insts = if expected_rule_insts = rule_insts andalso expected_proof_insts = proof_insts then () else let val _ = tracing (@{make_string} (rule_insts, proof_insts)) in error ("Trace_Schematic_Insts failed test: " ^ test_name) end fun tactic thm = resolve_tac ctxt [thm] in HEADGOAL (Trace_Schematic_Insts.trace_schematic_insts_tac ctxt check tactic thm) end \ text \Simultaneous rule and goal instantiations\ lemma "\a. foo (a :: nat)" apply (rule exI) apply (tactic \ let val a' = TVar (("'a", 0), @{sort type}) val b' = TVar (("'b", 0), @{sort zero}) val a'' = TVar (("'a", 2), @{sort type}) val expected_rule_vars = [ (Var (("x", 0), a'), Var(("x", 2), a'')) ] val expected_rule_tvars = [ (a', a''), (b', @{typ nat}) ] val expected_goal_vars = [ (Var (("a", 0), @{typ nat}), @{term "0 :: nat"}) ] in trace_schematic_resolve_tac_assert @{context} "basic rule tracing" @{thm fooI2} {bounds = [], terms = expected_rule_vars, typs = expected_rule_tvars} {bounds = [], terms = expected_goal_vars, typs = []} end \) by (simp add: foo_def) text \Rule instantiations with bound variables\ lemma "\X. X \ Y \ Y \ X" apply (tactic \ let val expected_rule_bounds = [("X", @{typ bool})] val expected_rule_vars = [ (Var (("P", 0), @{typ bool}), @{term "\X :: bool. Y :: bool"}), (Var (("Q", 0), @{typ bool}), @{term "\X :: bool. X :: bool"}) ] in trace_schematic_resolve_tac_assert @{context} "rule tracing with bound variables" @{thm conjI} {bounds = expected_rule_bounds, terms = expected_rule_vars, typs = []} {bounds = [], terms = [], typs = []} end \) by simp+ text \Rule instantiations with function terms\ lemma "\f. \x. f x = x" apply (intro exI allI) apply (rule fun_cong) apply (tactic \ let val a' = TVar (("'a", 0), @{sort type}) \\ New lambda abstraction gets an anonymous variable name. Usually rendered as @{term "\x a. a"}. \ val lambda = Abs ("x", @{typ 'a}, Abs ("", @{typ 'a}, Bound 0)) val expected_rule_bounds = [("x", @{typ 'a})] val expected_rule_vars = [ (Var (("t", 0), a'), lambda) ] val expected_rule_typs = [ (a', @{typ "'a \ 'a"}) ] val expected_goal_vars = [ (Var (("f", 2), @{typ "'a \ 'a \ 'a"}), lambda) ] in trace_schematic_resolve_tac_assert @{context} "rule tracing with function term instantiations" @{thm refl} {bounds = expected_rule_bounds, terms = expected_rule_vars, typs = expected_rule_typs} {bounds = [], terms = expected_goal_vars, typs = []} end \) done end end