equal
deleted
inserted
replaced
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] rty [quot] rsp_thms defs *} |
307 ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] rty [quot] 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 apply (tactic {* lift_tac_fset @{context} @{thm m1} 1 *}) |
|
311 done |
311 |
312 |
312 lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)" |
313 lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)" |
313 by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *}) |
314 by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *}) |
314 |
315 |
315 lemma "INSERT a (INSERT a x) = INSERT a x" |
316 lemma "INSERT a (INSERT a x) = INSERT a x" |
345 |
346 |
346 lemma cheat: "P" sorry |
347 lemma cheat: "P" sorry |
347 |
348 |
348 lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l" |
349 lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l" |
349 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *}) |
350 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *}) |
350 apply(tactic {* regularize_tac @{context} [rel_eqv] [rel_refl] 1 *}) |
351 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *}) |
351 prefer 2 |
352 prefer 2 |
352 apply(rule cheat) |
353 apply(rule cheat) |
353 apply(tactic {* r_mk_comb_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *) |
354 apply(tactic {* r_mk_comb_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *) |
354 apply(tactic {* r_mk_comb_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *) |
355 apply(tactic {* r_mk_comb_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *) |
355 apply(tactic {* r_mk_comb_tac_fset @{context} 1*}) (* 2 *) (* lam-lam-elim for R = (===>) *) |
356 apply(tactic {* r_mk_comb_tac_fset @{context} 1*}) (* 2 *) (* lam-lam-elim for R = (===>) *) |
447 ML {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty [quot] [rel_refl] [trans2] rsp_thms *} |
448 ML {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty [quot] [rel_refl] [trans2] rsp_thms *} |
448 |
449 |
449 (* Construction site starts here *) |
450 (* Construction site starts here *) |
450 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" |
451 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" |
451 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *}) |
452 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *}) |
452 apply (tactic {* regularize_tac @{context} [rel_eqv] [rel_refl] 1 *}) |
453 apply (tactic {* regularize_tac @{context} [rel_eqv] 1 *}) |
453 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *}) |
454 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *}) |
454 apply (rule FUN_QUOTIENT) |
455 apply (rule FUN_QUOTIENT) |
455 apply (rule FUN_QUOTIENT) |
456 apply (rule FUN_QUOTIENT) |
456 apply (rule IDENTITY_QUOTIENT) |
457 apply (rule IDENTITY_QUOTIENT) |
457 apply (rule FUN_QUOTIENT) |
458 apply (rule FUN_QUOTIENT) |