# HG changeset patch # User Cezary Kaliszyk # Date 1268336971 -3600 # Node ID 7a9217a7f68166d4468aa62c78f7881b462aa467 # Parent b41de1879daef54039f0b11484e7a94647bff3cb Do not fail if the finite support proof fails. diff -r b41de1879dae -r 7a9217a7f681 Nominal/Abs.thy --- a/Nominal/Abs.thy Thu Mar 11 19:43:50 2010 +0100 +++ b/Nominal/Abs.thy Thu Mar 11 20:49:31 2010 +0100 @@ -736,7 +736,7 @@ apply(simp add: inj_on_def) done -lemma +lemma supp_atom_image: fixes as::"'a::at_base set" shows "supp (atom ` as) = supp as" apply(simp add: supp_def) @@ -745,6 +745,12 @@ apply(simp add: atom_image_cong) done +lemma swap_atom_image_fresh: "\a \ atom ` (fn :: ('a :: at_base set)); b \ atom ` fn\ \ (a \ b) \ fn = fn" + apply (simp add: fresh_def) + apply (simp add: supp_atom_image) + apply (fold fresh_def) + apply (simp add: swap_fresh_fresh) + done end diff -r b41de1879dae -r 7a9217a7f681 Nominal/Fv.thy --- a/Nominal/Fv.thy Thu Mar 11 19:43:50 2010 +0100 +++ b/Nominal/Fv.thy Thu Mar 11 20:49:31 2010 +0100 @@ -714,7 +714,7 @@ simp_tac (HOL_ss addsimps @{thms supports_def not_in_union} @ perm) THEN_ALL_NEW ( REPEAT o rtac allI THEN' REPEAT o rtac impI THEN' split_conjs THEN' asm_full_simp_tac (HOL_ss addsimps @{thms fresh_def[symmetric] - swap_fresh_fresh fresh_atom swap_at_base_simps(3)})) + swap_fresh_fresh fresh_atom swap_at_base_simps(3) swap_atom_image_fresh})) *} ML {* @@ -766,7 +766,7 @@ ML {* fun fs_tac induct supports = ind_tac induct THEN_ALL_NEW ( rtac @{thm supports_finite} THEN' resolve_tac supports) THEN_ALL_NEW - asm_full_simp_tac (HOL_ss addsimps @{thms supp_atom finite_insert finite.emptyI finite_Un}) + asm_full_simp_tac (HOL_ss addsimps @{thms supp_atom supp_atom_image finite_insert finite.emptyI finite_Un}) *} ML {* diff -r b41de1879dae -r 7a9217a7f681 Nominal/Parser.thy --- a/Nominal/Parser.thy Thu Mar 11 19:43:50 2010 +0100 +++ b/Nominal/Parser.thy Thu Mar 11 20:49:31 2010 +0100 @@ -413,11 +413,11 @@ [Attrib.internal (fn _ => Nominal_ThmDecls.eqvt_add)]), q_eqvt) lthy19; val supports = map (prove_supports lthy20 q_perm) consts val _ = map tracing (map PolyML.makestring supports); - val fin_supp = HOLogic.conj_elims (prove_fs lthy20 q_induct supports q_tys); + val fin_supp = HOLogic.conj_elims (prove_fs lthy20 q_induct supports q_tys) handle _ => [] val thy3 = Local_Theory.exit_global lthy20; val lthy21 = Theory_Target.instantiation (qty_full_names, [], @{sort fs}) thy3; fun tac _ = Class.intro_classes_tac [] THEN (ALLGOALS (resolve_tac fin_supp)) - val lthy22 = Class.prove_instantiation_instance tac lthy21; + val lthy22 = Class.prove_instantiation_instance tac lthy21 handle _ => lthy20 in ((raw_dt_names, raw_bn_funs, raw_bn_eqs, raw_binds), lthy22) end