# HG changeset patch # User bulwahn # Date 1244714773 -7200 # Node ID 995d6b90e9d6a326d91ba0db2400735cf966d36d # Parent b63d253ed9e21386bad258d8905e2a5f8175d3ef making isatest happy; but misunderstanding remains diff -r b63d253ed9e2 -r 995d6b90e9d6 src/HOL/ex/Predicate_Compile_ex.thy --- a/src/HOL/ex/Predicate_Compile_ex.thy Wed Jun 10 21:04:36 2009 +0200 +++ b/src/HOL/ex/Predicate_Compile_ex.thy Thu Jun 11 12:06:13 2009 +0200 @@ -44,13 +44,15 @@ | "f x \ partition f xs ys zs \ partition f (x # xs) (x # ys) zs" | "\ f x \ partition f xs ys zs \ partition f (x # xs) ys (x # zs)" +(* FIXME: correct handling of parameters *) +(* ML {* reset Predicate_Compile.do_proofs *} - code_pred partition . thm partition.equation +ML {* set Predicate_Compile.do_proofs *} +*) -ML {* set Predicate_Compile.do_proofs *} (* TODO: requires to handle abstractions in parameter positions correctly *) (*FIXME values 10 "{(ys, zs). partition (\n. n mod 2 = 0) [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}" *) @@ -63,11 +65,11 @@ "r a b ==> tranclp r b c ==> tranclp r a c" by auto *) - +(* code_pred tranclp . thm tranclp.equation - +*) inductive succ :: "nat \ nat \ bool" where "succ 0 1" | "succ m n \ succ (Suc m) (Suc n)" diff -r b63d253ed9e2 -r 995d6b90e9d6 src/HOL/ex/predicate_compile.ML --- a/src/HOL/ex/predicate_compile.ML Wed Jun 10 21:04:36 2009 +0200 +++ b/src/HOL/ex/predicate_compile.ML Thu Jun 11 12:06:13 2009 +0200 @@ -1443,12 +1443,14 @@ (* TODO: must create state to prove multiple cases *) fun generic_code_pred prep_const raw_const lthy = let + val thy = ProofContext.theory_of lthy val const = prep_const thy raw_const + val lthy' = LocalTheory.theory (PredData.map (Graph.extend (dependencies_of thy) const)) lthy val thy' = ProofContext.theory_of lthy' val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy') - val _ = Output.tracing ("preds: " ^ commas preds) + fun mk_cases const = let val nparams = nparams_of thy' const @@ -1458,7 +1460,8 @@ fun after_qed thms = LocalTheory.theory (fold set_elim (map the_single thms) #> add_equations const) in - Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy' + Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy + (* FIXME: expected the actual local_theory to be lthy'; but works with lthy ??*) end; structure P = OuterParse