equal
deleted
inserted
replaced
340 lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)" |
340 lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)" |
341 apply (tactic {* lift_tac_fset @{context} @{thm append_assoc} 1 *}) |
341 apply (tactic {* lift_tac_fset @{context} @{thm append_assoc} 1 *}) |
342 done |
342 done |
343 |
343 |
344 ML {* val aps = findaps rty (prop_of (atomize_thm @{thm list.induct})) *} |
344 ML {* val aps = findaps rty (prop_of (atomize_thm @{thm list.induct})) *} |
345 ML {* val app_prs_thms = map (applic_prs @{context} rty qty absrep) aps *} |
345 ML {* val app_prs_thms = map (applic_prs_old @{context} rty qty absrep) aps *} |
346 |
346 |
347 lemma cheat: "P" sorry |
347 lemma cheat: "P" sorry |
348 |
348 |
349 lemma imp_refl: "P \<longrightarrow> P" by simp |
349 lemma imp_refl: "P \<longrightarrow> P" by simp |
350 |
350 |
351 thm Set.conj_mono |
351 thm Set.conj_mono |
352 thm Set.imp_mono |
352 thm Set.imp_mono |
353 ML {* Inductive.get_monos @{context} *} |
353 ML {* Inductive.get_monos @{context} *} |
354 |
354 |
355 lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l" |
355 lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l" |
356 apply(tactic {* procedure_tac @{thm list.induct} @{context} 1 *}) |
356 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *}) |
357 defer |
357 defer |
358 apply(rule cheat) |
358 apply(rule cheat) |
359 apply(rule cheat) |
359 apply(rule cheat) |
360 apply(atomize (full)) |
360 apply(atomize (full)) |
361 apply(rule my_equiv_res_forallR) |
361 apply(rule my_equiv_res_forallR) |
376 apply(tactic {* resolve_tac (Inductive.get_monos @{context}) 1 *}) |
376 apply(tactic {* resolve_tac (Inductive.get_monos @{context}) 1 *}) |
377 apply(tactic {* resolve_tac (Inductive.get_monos @{context}) 1 *}) |
377 apply(tactic {* resolve_tac (Inductive.get_monos @{context}) 1 *}) |
378 done |
378 done |
379 |
379 |
380 lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l" |
380 lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l" |
381 apply(tactic {* procedure_tac @{thm list.induct} @{context} 1 *}) |
381 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *}) |
382 defer |
382 defer |
383 apply(rule cheat) |
383 apply(rule cheat) |
384 apply(rule cheat) |
384 apply(rule cheat) |
385 apply(tactic {* regularize_tac @{context} rel_eqv rel_refl 1 *}) |
385 apply(tactic {* regularize_tac @{context} rel_eqv rel_refl 1 *}) |
386 oops |
386 done |
387 |
387 |
388 quotient_def |
388 quotient_def |
389 fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a" |
389 fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a" |
390 where |
390 where |
391 "fset_rec \<equiv> list_rec" |
391 "fset_rec \<equiv> list_rec" |