diff -r a2f2214dc881 -r 6088fea1c8b1 IntEx.thy --- a/IntEx.thy Mon Dec 07 12:14:25 2009 +0100 +++ b/IntEx.thy Mon Dec 07 14:00:36 2009 +0100 @@ -246,27 +246,31 @@ lemma "(\x. (x, x)) (y::my_int) = (y, y)" apply(tactic {* procedure_tac @{context} @{thm lam_tst} 1 *}) -(*apply(tactic {* regularize_tac @{context} 1 *}) *) -defer +apply(tactic {* regularize_tac @{context} 1 *}) apply(tactic {* all_inj_repabs_tac_intex @{context} 1*}) apply(tactic {* clean_tac @{context} 1 *}) apply(simp add: pair_prs) -sorry +done lemma lam_tst2: "(\(y :: nat \ nat). y) = (\(x :: nat \ nat). x)" by simp + + + lemma "(\(y :: my_int). y) = (\(x :: my_int). x)" -apply(tactic {* ObjectLogic.full_atomize_tac 1 *}) -apply(tactic {* gen_frees_tac @{context} 1 *}) -ML_prf {* regularize_trm @{context} (prop_of @{thm lam_tst2}) @{prop "(\(y :: my_int). y) = (\(x :: my_int). x)"} *} -ML_prf {* print_depth 50 *} -ML_prf {* Syntax.check_term @{context} it *} -ML_prf {* procedure_inst @{context} (prop_of @{thm lam_tst2}) @{prop "(\(y :: my_int). y) = (\(x :: my_int). x)"} *} +apply(tactic {* procedure_tac @{context} @{thm lam_tst2} 1 *}) +defer +apply(tactic {* all_inj_repabs_tac_intex @{context} 1*}) +(*apply(tactic {* lambda_prs_tac @{context} 1 *})*) +sorry -apply(tactic {* procedure_tac @{context} @{thm lam_tst2} 1 *}) -apply(tactic {* regularize_tac @{context} 1 *}) +lemma lam_tst3: "(\(y :: nat \ nat \ nat \ nat). y) = (\(x :: nat \ nat \ nat \ nat). x)" +by auto + +lemma "(\(y :: my_int \ my_int). y) = (\(x :: my_int \ my_int). x)" +apply(tactic {* procedure_tac @{context} @{thm lam_tst3} 1 *}) +defer apply(tactic {* all_inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* clean_tac @{context} 1 *}) -done - +apply(tactic {* lambda_prs_tac @{context} 1 *}) +sorry