FSet.thy
changeset 400 7ef153ded7e2
parent 398 fafcc54e531d
child 401 211229d6c66f
equal deleted inserted replaced
398:fafcc54e531d 400:7ef153ded7e2
   339 
   339 
   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})) *}
       
   345 ML {* val app_prs_thms = map (applic_prs_old @{context} rty qty absrep) aps *}
       
   346 
       
   347 lemma cheat: "P" sorry
   344 lemma cheat: "P" sorry
   348 
   345 
   349 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"
   350 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
   347 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
   351 apply(rule cheat)
   348 apply(rule cheat)
   352 prefer 2
   349 prefer 2
   353 apply(rule cheat)
   350 apply(rule cheat)
   354 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 sorry
   355 
   353 
   356 quotient_def
   354 quotient_def
   357   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"
   358 where
   356 where
   359   "fset_rec \<equiv> list_rec"
   357   "fset_rec \<equiv> list_rec"
   405 
   403 
   406 
   404 
   407 (* Construction site starts here *)
   405 (* Construction site starts here *)
   408 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"
   409 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *})
   407 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *})
   410 apply (tactic {* regularize_tac @{context} rel_eqv rel_refl 1 *})
   408 apply (tactic {* regularize_tac @{context} [rel_eqv] rel_refl 1 *})
   411 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
   409 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
   412 apply (rule FUN_QUOTIENT)
   410 apply (rule FUN_QUOTIENT)
   413 apply (rule FUN_QUOTIENT)
   411 apply (rule FUN_QUOTIENT)
   414 apply (rule IDENTITY_QUOTIENT)
   412 apply (rule IDENTITY_QUOTIENT)
   415 apply (rule FUN_QUOTIENT)
   413 apply (rule FUN_QUOTIENT)