FSet.thy
changeset 466 42082fc00903
parent 462 0911d3aabf47
child 467 5ca4a927d7f0
equal deleted inserted replaced
465:ce1f8a284920 466:42082fc00903
   306 
   306 
   307 lemma "IN x EMPTY = False"
   307 lemma "IN x EMPTY = False"
   308 apply(tactic {* procedure_tac @{context} @{thm m1} 1 *})
   308 apply(tactic {* procedure_tac @{context} @{thm m1} 1 *})
   309 apply(tactic {* regularize_tac @{context}  [rel_eqv] 1 *})
   309 apply(tactic {* regularize_tac @{context}  [rel_eqv] 1 *})
   310 apply(tactic {* all_inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
   310 apply(tactic {* all_inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
   311 ML_prf {* 
   311 apply(tactic {* clean_tac @{context} [quot] defs 1*})
   312   val rtrm = @{term "\<forall>x. IN x EMPTY = False"}
       
   313   val qtrm = @{term "\<forall>x. x memb [] = False"}
       
   314   val aps = find_aps rtrm qtrm 
       
   315 *}
       
   316 apply(tactic {* clean_tac @{context} [quot] defs aps 1*})
       
   317 done
   312 done
   318 
   313 
   319 lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
   314 lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
   320 by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *})
   315 by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *})
   321 
   316 
   356 apply (tactic {* (ObjectLogic.full_atomize_tac THEN' gen_frees_tac @{context}) 1 *})
   351 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 ())))) *}
   352 ML_prf {* val qtm = #concl (fst (Subgoal.focus @{context} 1 (#goal (Isar.goal ())))) *}
   358 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
   353 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
   359 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
   354 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
   360 prefer 2
   355 prefer 2
   361 apply(tactic {* clean_tac @{context} [quot] defs [] 1 *})
   356 apply(tactic {* clean_tac @{context} [quot] defs 1 *})
   362 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
   357 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
   363 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
   358 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
   364 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 2 *) (* lam-lam-elim for R = (===>) *) 
   359 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 2 *) (* lam-lam-elim for R = (===>) *) 
   365 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
   360 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 3 *) (* Ball-Ball *)
   366 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
   361 apply(tactic {* inj_repabs_tac_fset @{context} 1*}) (* 9 *) (* Rep-Abs-elim - can be complex Rep-Abs *)
   454 (* Construction site starts here *)
   449 (* Construction site starts here *)
   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"
   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"
   456 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *})
   451 apply (tactic {* procedure_tac @{context} @{thm list_induct_part} 1 *})
   457 apply (tactic {* regularize_tac @{context} [rel_eqv] 1 *})
   452 apply (tactic {* regularize_tac @{context} [rel_eqv] 1 *})
   458 prefer 2
   453 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 *})
   454 apply (tactic {* clean_tac @{context} [quot] defs 1 *})
   460 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
   455 apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *})
   461 apply (rule FUN_QUOTIENT)
   456 apply (rule FUN_QUOTIENT)
   462 apply (rule FUN_QUOTIENT)
   457 apply (rule FUN_QUOTIENT)
   463 apply (rule IDENTITY_QUOTIENT)
   458 apply (rule IDENTITY_QUOTIENT)
   464 apply (rule FUN_QUOTIENT)
   459 apply (rule FUN_QUOTIENT)