Nominal/Abs.thy
changeset 1469 0c5dfd2866bb
parent 1465 4de35639fef0
parent 1467 77b86f1fc936
child 1470 3127c75275a6
equal deleted inserted replaced
1466:d18defacda25 1469:0c5dfd2866bb
   746   apply (simp add: supp_atom_image)
   746   apply (simp add: supp_atom_image)
   747   apply (fold fresh_def)
   747   apply (fold fresh_def)
   748   apply (simp add: swap_fresh_fresh)
   748   apply (simp add: swap_fresh_fresh)
   749   done
   749   done
   750 
   750 
       
   751 (* TODO: The following lemmas can be moved somewhere... *)
       
   752 lemma split_rsp2[quot_respect]: "((R1 ===> R2 ===> prod_rel R1 R2 ===> op =) ===>
       
   753   prod_rel R1 R2 ===> prod_rel R1 R2 ===> op =) split split"
       
   754   by auto
       
   755 
       
   756 lemma split_prs2[quot_preserve]:
       
   757   assumes q1: "Quotient R1 Abs1 Rep1"
       
   758   and q2: "Quotient R2 Abs2 Rep2"
       
   759   shows "((Abs1 ---> Abs2 ---> prod_fun Abs1 Abs2 ---> id) ---> prod_fun Rep1 Rep2 ---> prod_fun Rep1 Rep2 ---> id) split = split"
       
   760   by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
       
   761 
       
   762 lemma alpha_gen2:
       
   763   "(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) =
       
   764   (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)"
       
   765 by (simp add: alpha_gen)
       
   766 
   751 
   767 
   752 end
   768 end
   753 
   769