diff -r 9c6991411e1f -r a68c51dd85b3 IntEx.thy --- a/IntEx.thy Sun Dec 06 02:41:35 2009 +0100 +++ b/IntEx.thy Sun Dec 06 04:03:08 2009 +0100 @@ -147,15 +147,37 @@ apply(rule refl) done -abbreviation - "Abs \ ABS_my_int" -abbreviation - "Rep \ REP_my_int" -abbreviation - "Resp \ Respects" -thm apply_rsp rep_abs_rsp lambda_prs -ML {* map (Pattern.pattern o bare_concl o prop_of) @{thms apply_rsp rep_abs_rsp lambda_prs} *} +(* +lemma yy: + "(REP_my_int ---> id) + (\x. Ball (Respects op \) + ((ABS_my_int ---> id) + ((REP_my_int ---> id) + (\b. (ABS_my_int ---> ABS_my_int ---> REP_my_int) + ((REP_my_int ---> REP_my_int ---> ABS_my_int) my_plus) + (REP_my_int (ABS_my_int x)) (REP_my_int (ABS_my_int b)) \ + (ABS_my_int ---> ABS_my_int ---> REP_my_int) + ((REP_my_int ---> REP_my_int ---> ABS_my_int) my_plus) + (REP_my_int (ABS_my_int x)) (REP_my_int (ABS_my_int b)))))) = +(\x. Ball (Respects op \) + ((ABS_my_int ---> id) + ((REP_my_int ---> id) + (\b. (ABS_my_int ---> ABS_my_int ---> REP_my_int) + ((REP_my_int ---> REP_my_int ---> ABS_my_int) my_plus) (REP_my_int x) + (REP_my_int (ABS_my_int b)) \ + (ABS_my_int ---> ABS_my_int ---> REP_my_int) + ((REP_my_int ---> REP_my_int ---> ABS_my_int) my_plus) (REP_my_int x) + (REP_my_int (ABS_my_int b))))))" +apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [lambda_prs2 @{theory}]) 1*}) + +apply(rule lambda_prs) +apply(tactic {* quotient_tac @{context} 1 *}) +apply(simp add: id_simps) +apply(tactic {* quotient_tac @{context} 1 *}) +done +*) + lemma "PLUS a b = PLUS a b"