diff -r 2d9de77d5687 -r d616a0912245 Quot/Examples/FSet.thy --- a/Quot/Examples/FSet.thy Wed Dec 09 06:21:09 2009 +0100 +++ b/Quot/Examples/FSet.thy Wed Dec 09 15:11:49 2009 +0100 @@ -373,10 +373,10 @@ apply (lifting ttt) done + lemma ttt2: "(\e. ((op @) x ((op #) e []))) = (\e. ((op #) e x))" sorry -(* PROBLEM *) lemma "(\e. (FUNION x (INSERT e EMPTY))) = (\e. (INSERT e x))" apply(lifting ttt2) apply(regularize) @@ -389,6 +389,7 @@ lemma ttt3: "(\x. ((op @) x ((op #) e []))) = (\x. ((op #) e x))" sorry + lemma "(\x. (FUNION x (INSERT e EMPTY))) = (\x. (INSERT e x))" (* apply (tactic {* procedure_tac @{context} @{thm ttt3} 1 *}) *) sorry @@ -396,9 +397,10 @@ lemma hard: "(\P. \Q. P (Q (x::'a list))) = (\P. \Q. Q (P (x::'a list)))" sorry -(* PROBLEM *) lemma hard_lift: "(\P. \Q. P (Q (x::'a fset))) = (\P. \Q. Q (P (x::'a fset)))" apply(lifting hard) +apply(regularize) +apply(auto) sorry end