FSet.thy
changeset 416 3f3927f793d4
parent 414 4dad34ca50db
child 419 b1cd040ff5f7
equal deleted inserted replaced
415:5a9bdf81672d 416:3f3927f793d4
   302   @ @{thms ho_all_prs ho_ex_prs} *}
   302   @ @{thms ho_all_prs ho_ex_prs} *}
   303 
   303 
   304 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
   304 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
   305 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset"; *}
   305 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset"; *}
   306 ML {* val consts = lookup_quot_consts defs *}
   306 ML {* val consts = lookup_quot_consts defs *}
   307 ML {* fun lift_tac_fset lthy t = lift_tac lthy t rel_eqv rel_refl rty quot trans2 rsp_thms reps_same absrep defs *}
   307 ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] trans2 rsp_thms defs *}
   308 
   308 
   309 lemma "IN x EMPTY = False"
   309 lemma "IN x EMPTY = False"
   310 by (tactic {* lift_tac_fset @{context} @{thm m1} 1 *})
   310 by (tactic {* lift_tac_fset @{context} @{thm m1} 1 *})
   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)"
   343 
   343 
   344 lemma cheat: "P" sorry
   344 lemma cheat: "P" sorry
   345 
   345 
   346 lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"
   346 lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"
   347 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
   347 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
   348 apply(tactic {* regularize_tac @{context} [rel_eqv] rel_refl 1 *})
   348 apply(tactic {* regularize_tac @{context} [rel_eqv] [rel_refl] 1 *})
   349 prefer 2
   349 prefer 2
   350 apply (tactic {* clean_tac @{context} quot defs reps_same absrep [(@{typ "('a list \<Rightarrow> bool)"},@{typ "('a fset \<Rightarrow> bool)"})] 1 *})
   350 apply (tactic {* clean_tac @{context} [quot] defs [(@{typ "('a list \<Rightarrow> bool)"},@{typ "('a fset \<Rightarrow> bool)"})] 1 *})
   351 apply(tactic {* r_mk_comb_tac' @{context} rty quot rel_refl trans2 rsp_thms 1*})
   351 apply(tactic {* r_mk_comb_tac' @{context} rty [quot] rel_refl trans2 rsp_thms 1*})
   352 done
   352 done
   353 
   353 
   354 quotient_def
   354 quotient_def
   355   fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
   355   fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
   356 where
   356 where
   374   apply (auto simp add: FUN_REL_EQ)
   374   apply (auto simp add: FUN_REL_EQ)
   375   sorry
   375   sorry
   376 
   376 
   377 ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *}
   377 ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *}
   378 ML {* val defs = @{thms fset_rec_def fset_case_def} @ defs *}
   378 ML {* val defs = @{thms fset_rec_def fset_case_def} @ defs *}
   379 ML {* fun lift_tac_fset lthy t = lift_tac lthy t rel_eqv rel_refl rty quot trans2 rsp_thms reps_same absrep defs *}
   379 ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] trans2 rsp_thms defs *}
   380 
   380 
   381 lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
   381 lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
   382 apply (tactic {* lift_tac_fset @{context} @{thm list.recs(2)} 1 *})
   382 apply (tactic {* lift_tac_fset @{context} @{thm list.recs(2)} 1 *})
   383 done
   383 done
   384 
   384 
   395   apply (rule b)
   395   apply (rule b)
   396   apply (assumption)
   396   apply (assumption)
   397   done
   397   done
   398 
   398 
   399 
   399 
   400 ML {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms *}
   400 ML {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty [quot] [rel_refl] trans2 rsp_thms *}
   401 
   401 
   402 
   402 
   403 
   403 
   404 
   404 
   405 (* Construction site starts here *)
   405 (* Construction site starts here *)
   406 lemma "P (x :: 'a list) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
   406 lemma "P (x :: 'a list) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
   407 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *})
   407 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *})
   408 apply (tactic {* regularize_tac @{context} [rel_eqv] rel_refl 1 *})
   408 apply (tactic {* regularize_tac @{context} [rel_eqv] [rel_refl] 1 *})
   409 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
   409 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
   410 apply (rule FUN_QUOTIENT)
   410 apply (rule FUN_QUOTIENT)
   411 apply (rule FUN_QUOTIENT)
   411 apply (rule FUN_QUOTIENT)
   412 apply (rule IDENTITY_QUOTIENT)
   412 apply (rule IDENTITY_QUOTIENT)
   413 apply (rule FUN_QUOTIENT)
   413 apply (rule FUN_QUOTIENT)
   452 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   452 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   453 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   453 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   454 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   454 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   455 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   455 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   456 apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
   456 apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
   457 apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
   457 apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac [quot]])) 1 *})
   458 apply assumption
   458 apply assumption
   459 apply (rule refl)
   459 apply (rule refl)
   460 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   460 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   461 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   461 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   462 apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
   462 apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
   463 apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
   463 apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
   464 apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
   464 apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac [quot]])) 1 *})
   465 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   465 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   466 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   466 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   467 apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
   467 apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
   468 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   468 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   469 apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
   469 apply (tactic {* instantiate_tac @{thm APPLY_RSP2} @{context} 1 *})
   470 apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac quot])) 1 *})
   470 apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP(1)} @{context} THEN' (RANGE [quotient_tac [quot]])) 1 *})
   471 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   471 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   472 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   472 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   473 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   473 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   474 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   474 apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})
   475 apply (tactic {* clean_tac @{context} quot defs reps_same absrep [(@{typ "('a list \<Rightarrow> 'c list \<Rightarrow> bool)"},@{typ "('a list \<Rightarrow> 'c fset \<Rightarrow> bool)"})] 1 *})
   475 apply (tactic {* clean_tac @{context} [quot] defs [(@{typ "('a list \<Rightarrow> 'c list \<Rightarrow> bool)"},@{typ "('a list \<Rightarrow> 'c fset \<Rightarrow> bool)"})] 1 *})
   476 done
   476 done
   477 
   477 
   478 end
   478 end