equal
deleted
inserted
replaced
716 apply (simp add: supp_atom_image) |
716 apply (simp add: supp_atom_image) |
717 apply (fold fresh_def) |
717 apply (fold fresh_def) |
718 apply (simp add: swap_fresh_fresh) |
718 apply (simp add: swap_fresh_fresh) |
719 done |
719 done |
720 |
720 |
|
721 (* TODO: The following lemmas can be moved somewhere... *) |
|
722 lemma split_rsp2[quot_respect]: "((R1 ===> R2 ===> prod_rel R1 R2 ===> op =) ===> |
|
723 prod_rel R1 R2 ===> prod_rel R1 R2 ===> op =) split split" |
|
724 by auto |
|
725 |
|
726 lemma split_prs2[quot_preserve]: |
|
727 assumes q1: "Quotient R1 Abs1 Rep1" |
|
728 and q2: "Quotient R2 Abs2 Rep2" |
|
729 shows "((Abs1 ---> Abs2 ---> prod_fun Abs1 Abs2 ---> id) ---> prod_fun Rep1 Rep2 ---> prod_fun Rep1 Rep2 ---> id) split = split" |
|
730 by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]) |
|
731 |
|
732 lemma alpha_gen2: |
|
733 "(bs, x1, x2) \<approx>gen (\<lambda>(x1, y1) (x2, y2). R1 x1 x2 \<and> R2 y1 y2) (\<lambda>(a, b). f1 a \<union> f2 b) pi (cs, y1, y2) = |
|
734 (f1 x1 \<union> f2 x2 - bs = f1 y1 \<union> f2 y2 - cs \<and> (f1 x1 \<union> f2 x2 - bs) \<sharp>* pi \<and> R1 (pi \<bullet> x1) y1 \<and> R2 (pi \<bullet> x2) y2)" |
|
735 by (simp add: alpha_gen) |
|
736 |
721 |
737 |
722 end |
738 end |
723 |
739 |