FSet.thy
changeset 462 0911d3aabf47
parent 459 020e27417b59
child 466 42082fc00903
equal deleted inserted replaced
459:020e27417b59 462:0911d3aabf47
   311 ML_prf {* 
   311 ML_prf {* 
   312   val rtrm = @{term "\<forall>x. IN x EMPTY = False"}
   312   val rtrm = @{term "\<forall>x. IN x EMPTY = False"}
   313   val qtrm = @{term "\<forall>x. x memb [] = False"}
   313   val qtrm = @{term "\<forall>x. x memb [] = False"}
   314   val aps = find_aps rtrm qtrm 
   314   val aps = find_aps rtrm qtrm 
   315 *}
   315 *}
   316 unfolding IN_def EMPTY_def
       
   317 apply(tactic {* clean_tac @{context} [quot] defs aps 1*})
   316 apply(tactic {* clean_tac @{context} [quot] defs aps 1*})
   318 done
   317 done
   319 
   318 
   320 lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
   319 lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
   321 by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *})
   320 by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *})
   352 done
   351 done
   353 
   352 
   354 lemma cheat: "P" sorry
   353 lemma cheat: "P" sorry
   355 
   354 
   356 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 {* (ObjectLogic.full_atomize_tac THEN' gen_frees_tac @{context}) 1 *})
       
   357 ML_prf {* val qtm = #concl (fst (Subgoal.focus @{context} 1 (#goal (Isar.goal ())))) *}
   357 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
   358 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
   358 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
   359 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
   359 prefer 2
   360 prefer 2
   360 apply(rule cheat)
   361 apply(tactic {* clean_tac @{context} [quot] defs [] 1 *})
   361 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
   362 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
   362 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
   363 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
   363 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 2 *) (* lam-lam-elim for R = (===>) *) 
   364 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 2 *) (* lam-lam-elim for R = (===>) *) 
   364 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
   365 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
   365 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
   366 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
   452 
   453 
   453 (* Construction site starts here *)
   454 (* Construction site starts here *)
   454 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"
   455 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"
   455 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *})
   456 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *})
   456 apply (tactic {* regularize_tac @{context} [rel_eqv] 1 *})
   457 apply (tactic {* regularize_tac @{context} [rel_eqv] 1 *})
       
   458 prefer 2
       
   459 apply (tactic {* clean_tac @{context} [quot] defs [(@{typ "('a list \<Rightarrow> 'c list \<Rightarrow> bool)"},@{typ "('a list \<Rightarrow> 'c fset \<Rightarrow> bool)"})] 1 *})
   457 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
   460 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
   458 apply (rule FUN_QUOTIENT)
   461 apply (rule FUN_QUOTIENT)
   459 apply (rule FUN_QUOTIENT)
   462 apply (rule FUN_QUOTIENT)
   460 apply (rule IDENTITY_QUOTIENT)
   463 apply (rule IDENTITY_QUOTIENT)
   461 apply (rule FUN_QUOTIENT)
   464 apply (rule FUN_QUOTIENT)
   521 apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP} @{context} THEN' (RANGE [quotient_tac [quot]])) 1 *})
   524 apply (tactic {* (instantiate_tac @{thm REP_ABS_RSP} @{context} THEN' (RANGE [quotient_tac [quot]])) 1 *})
   522 apply (tactic {* (inj_repabs_tac_fset @{context}) 1 *})
   525 apply (tactic {* (inj_repabs_tac_fset @{context}) 1 *})
   523 apply (tactic {* (inj_repabs_tac_fset @{context}) 1 *})
   526 apply (tactic {* (inj_repabs_tac_fset @{context}) 1 *})
   524 apply (tactic {* (inj_repabs_tac_fset @{context}) 1 *})
   527 apply (tactic {* (inj_repabs_tac_fset @{context}) 1 *})
   525 apply (tactic {* (inj_repabs_tac_fset @{context}) 1 *})
   528 apply (tactic {* (inj_repabs_tac_fset @{context}) 1 *})
   526 apply (tactic {* clean_tac @{context} [quot] defs [(@{typ "('a list \<Rightarrow> 'c list \<Rightarrow> bool)"},@{typ "('a list \<Rightarrow> 'c fset \<Rightarrow> bool)"})] 1 *})
       
   527 done
   529 done
   528 
   530 
   529 end
   531 end