# HG changeset patch # User Cezary Kaliszyk # Date 1260438885 -3600 # Node ID 2eba169533b560d3f5b976a4c1bc6edf42651d92 # Parent 2779da3cd02cfeb405f7e8f78e842d0de886af0e Found the problem with ttt3. diff -r 2779da3cd02c -r 2eba169533b5 Quot/Examples/FSet.thy --- a/Quot/Examples/FSet.thy Thu Dec 10 10:36:05 2009 +0100 +++ b/Quot/Examples/FSet.thy Thu Dec 10 10:54:45 2009 +0100 @@ -343,13 +343,14 @@ apply(rule list_eq_refl) done -lemma ttt3: "(\x. ((op @) x ((op #) e []))) = (\x. ((op #) e x))" +lemma ttt3: "(\x. ((op @) x (e # []))) = (op #) e" sorry - -lemma "(\x. (FUNION x (INSERT e EMPTY))) = (\x. (INSERT e x))" -(* apply (tactic {* procedure_tac @{context} @{thm ttt3} 1 *}) *) -sorry +lemma "(\x. (FUNION x (INSERT e EMPTY))) = INSERT e" +apply(lifting ttt3) +apply(regularize) +apply(auto simp add: cons_rsp) +done lemma hard: "(\P. \Q. P (Q (x::'a list))) = (\P. \Q. Q (P (x::'a list)))" sorry diff -r 2779da3cd02c -r 2eba169533b5 Quot/QuotMain.thy --- a/Quot/QuotMain.thy Thu Dec 10 10:36:05 2009 +0100 +++ b/Quot/QuotMain.thy Thu Dec 10 10:54:45 2009 +0100 @@ -345,7 +345,9 @@ else raise (LIFT_MATCH "regularize (bounds mismatch)") | (rt, qt) => - raise (LIFT_MATCH "regularize failed (default)") + let val (rts, qts) = (Syntax.string_of_term lthy rt, Syntax.string_of_term lthy qt) in + raise (LIFT_MATCH ("regularize failed (default: " ^ rts ^ "," ^ qts ^ ")")) + end *} section {* Regularize Tactic *}