Solutions to IntEx tests.
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Sat, 05 Dec 2009 21:28:24 +0100
changeset 554 8395fc6a6945
parent 553 09cd71fac4ec
child 555 b460565dbb58
child 557 e67961288b12
Solutions to IntEx tests.
IntEx.thy
QuotMain.thy
QuotScript.thy
--- a/IntEx.thy	Sat Dec 05 16:26:18 2009 +0100
+++ b/IntEx.thy	Sat Dec 05 21:28:24 2009 +0100
@@ -173,7 +173,33 @@
 
 lemma "PLUS a = PLUS a"
 apply(tactic {* procedure_tac @{context} @{thm test2} 1 *})
-oops
+apply(rule ballI)
+apply(rule apply_rsp[OF Quotient_my_int ho_plus_rsp])
+apply(simp only: in_respects)
+
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(rule quot_rel_rsp)
+apply(tactic {* quotient_tac @{context} 1 *})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(fold PLUS_def)
+apply(tactic {* clean_tac @{context} 1 *})
+thm Quotient_rel_rep
+thm Quotient_rel_rep[OF fun_quotient[OF Quotient_my_int Quotient_my_int]]
+apply(simp only: Quotient_rel_rep[OF fun_quotient[OF Quotient_my_int Quotient_my_int]])
+done
 
 lemma test3: "my_plus = my_plus"
 apply(rule refl)
@@ -181,7 +207,18 @@
 
 lemma "PLUS = PLUS"
 apply(tactic {* procedure_tac @{context} @{thm test3} 1 *})
-oops
+apply(rule ho_plus_rsp)
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(rule quot_rel_rsp)
+apply(tactic {* quotient_tac @{context} 1 *})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* inj_repabs_tac_intex @{context} 1*})
+apply(tactic {* clean_tac @{context} 1 *})
+apply(simp only: Quotient_rel_rep[OF fun_quotient[OF Quotient_my_int fun_quotient[OF Quotient_my_int Quotient_my_int]]])
+done
 
 
 lemma "PLUS a b = PLUS b a"
--- a/QuotMain.thy	Sat Dec 05 16:26:18 2009 +0100
+++ b/QuotMain.thy	Sat Dec 05 21:28:24 2009 +0100
@@ -964,12 +964,6 @@
     | _ => no_tac)
 *}
 
-lemma fun_rel_id:
-  assumes a: "\<And>x y. R1 x y \<Longrightarrow> R2 (f x) (g y)"
-  shows "(R1 ===> R2) f g"
-using a by simp
-
-
 ML {*
 fun inj_repabs_tac_match ctxt trans2 = SUBGOAL (fn (goal, i) =>
 (case (bare_concl goal) of
--- a/QuotScript.thy	Sat Dec 05 16:26:18 2009 +0100
+++ b/QuotScript.thy	Sat Dec 05 21:28:24 2009 +0100
@@ -335,6 +335,19 @@
   using a unfolding Quotient_def
   by (metis COMBC_def Collect_def Collect_mem_eq in_respects fun_map.simps id_apply)
 
+lemma fun_rel_id:
+  assumes a: "\<And>x y. R1 x y \<Longrightarrow> R2 (f x) (g y)"
+  shows "(R1 ===> R2) f g"
+using a by simp
+
+lemma quot_rel_rsp:
+  assumes a: "Quotient R Abs Rep"
+  shows "(R ===> R ===> op =) R R"
+  apply(rule fun_rel_id)+
+  apply(rule equals_rsp[OF a])
+  apply(assumption)+
+  done
+