equal
deleted
inserted
replaced
320 apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(5)} 1 *}) |
320 apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(5)} 1 *}) |
321 done |
321 done |
322 |
322 |
323 lemma "CARD x = Suc n \<Longrightarrow> (\<exists>a b. \<not> IN a b & x = INSERT a b)" |
323 lemma "CARD x = Suc n \<Longrightarrow> (\<exists>a b. \<not> IN a b & x = INSERT a b)" |
324 apply (tactic {* lift_tac_fset @{context} @{thm card1_suc} 1 *}) |
324 apply (tactic {* lift_tac_fset @{context} @{thm card1_suc} 1 *}) |
325 done |
325 oops |
326 |
326 |
327 lemma "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))" |
327 lemma "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))" |
328 apply (tactic {* lift_tac_fset @{context} @{thm not_mem_card1} 1 *}) |
328 apply (tactic {* lift_tac_fset @{context} @{thm not_mem_card1} 1 *}) |
329 done |
329 done |
330 |
330 |
345 |
345 |
346 lemma cheat: "P" sorry |
346 lemma cheat: "P" sorry |
347 |
347 |
348 lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l" |
348 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 *}) |
349 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *}) |
350 apply(tactic {* regularize_tac @{context} [rel_eqv] [rel_refl] 1 *}) |
350 apply(tactic {* regularize_tac @{context} rel_eqv [rel_refl] 1 *}) |
351 prefer 2 |
351 prefer 2 |
352 apply(rule cheat) |
352 apply(rule cheat) |
353 apply(tactic {* r_mk_comb_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *) |
353 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 *) |
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*}) (* 2 *) (* lam-lam-elim for R = (===>) *) |
355 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 *} |
447 ML {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty [quot] [rel_refl] [trans2] rsp_thms *} |
448 |
448 |
449 (* Construction site starts here *) |
449 (* 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" |
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 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *}) |
451 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *}) |
452 apply (tactic {* regularize_tac @{context} [rel_eqv] [rel_refl] 1 *}) |
452 apply (tactic {* regularize_tac @{context} rel_eqv [rel_refl] 1 *}) |
453 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *}) |
453 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *}) |
454 apply (rule FUN_QUOTIENT) |
454 apply (rule FUN_QUOTIENT) |
455 apply (rule FUN_QUOTIENT) |
455 apply (rule FUN_QUOTIENT) |
456 apply (rule IDENTITY_QUOTIENT) |
456 apply (rule IDENTITY_QUOTIENT) |
457 apply (rule FUN_QUOTIENT) |
457 apply (rule FUN_QUOTIENT) |