FSet.thy
changeset 489 2b7b349e470f
parent 483 74348dc2f8bb
child 495 76fa93b1fe8b
equal deleted inserted replaced
488:508f3106b89c 489:2b7b349e470f
   298   @{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp ho_map_rsp ho_append_rsp ho_fold_rsp} *}
   298   @{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp ho_map_rsp ho_append_rsp ho_fold_rsp} *}
   299 
   299 
   300 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
   300 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
   301 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset"; *}
   301 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset"; *}
   302 ML {* val consts = lookup_quot_consts defs *}
   302 ML {* val consts = lookup_quot_consts defs *}
   303 ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] defs *}
   303 ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] [quot] defs *}
   304 
   304 
   305 lemma "IN x EMPTY = False"
   305 lemma "IN x EMPTY = False"
   306 apply(tactic {* procedure_tac @{context} @{thm m1} 1 *})
   306 apply(tactic {* procedure_tac @{context} @{thm m1} 1 *})
   307 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
   307 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
   308 apply(tactic {* all_inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
   308 apply(tactic {* all_inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
   309 apply(tactic {* clean_tac @{context} [quot] defs 1*})
   309 apply(tactic {* clean_tac @{context} [quot] defs 1*})
   310 done
   310 done
   311 
   311 
   312 lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
   312 lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
   313 by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *})
   313 by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *})
   331 lemma "FOLD f g (z::'b) (INSERT a x) =
   331 lemma "FOLD f g (z::'b) (INSERT a x) =
   332   (if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)"
   332   (if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)"
   333 apply(tactic {* lift_tac_fset @{context} @{thm fold1.simps(2)} 1 *})
   333 apply(tactic {* lift_tac_fset @{context} @{thm fold1.simps(2)} 1 *})
   334 done
   334 done
   335 
   335 
   336 ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy rty [quot] [rel_refl] [trans2] *}
   336 ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy [quot] [rel_refl] [trans2] *}
   337 
   337 
   338 lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"
   338 lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"
   339 apply (tactic {* lift_tac_fset @{context} @{thm map_append} 1 *})
   339 apply (tactic {* lift_tac_fset @{context} @{thm map_append} 1 *})
   340 done
   340 done
   341 
   341 
   433 where
   433 where
   434   "INSERT2 \<equiv> op #"
   434   "INSERT2 \<equiv> op #"
   435 
   435 
   436 ML {* val defs = @{thms EMPTY2_def INSERT2_def} @ defs *}
   436 ML {* val defs = @{thms EMPTY2_def INSERT2_def} @ defs *}
   437 ML {* val quot = @{thms QUOTIENT_fset QUOTIENT_fset2} *}
   437 ML {* val quot = @{thms QUOTIENT_fset QUOTIENT_fset2} *}
   438 ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy rty quot [rel_refl] [trans2] *}
   438 ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy quot [rel_refl] [trans2] *}
   439 ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty quot defs *}
   439 ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] quot defs *}
   440 
   440 
   441 lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
   441 lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
   442 apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
   442 apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
   443 done
   443 done
   444 
   444 
   469   apply (auto simp add: FUN_REL_EQ)
   469   apply (auto simp add: FUN_REL_EQ)
   470   sorry
   470   sorry
   471 
   471 
   472 ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *}
   472 ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *}
   473 ML {* val defs = @{thms fset_rec_def fset_case_def} @ defs *}
   473 ML {* val defs = @{thms fset_rec_def fset_case_def} @ defs *}
   474 ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] defs *}
   474 ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] quot defs *}
   475 
   475 
   476 lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
   476 lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
   477 apply (tactic {* lift_tac_fset @{context} @{thm list.recs(2)} 1 *})
   477 apply (tactic {* lift_tac_fset @{context} @{thm list.recs(2)} 1 *})
   478 done
   478 done
   479 
   479