# HG changeset patch # User traytel # Date 1437042202 -7200 # Node ID 26ffdb9667598195681e346db8ce3370c656cc01 # Parent 53697011b03af96595ce90c13d2b42433ba03d70 {r,e,d,f}tac with proper context in BNF diff -r 53697011b03a -r 26ffdb966759 src/HOL/Library/bnf_lfp_countable.ML --- a/src/HOL/Library/bnf_lfp_countable.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Library/bnf_lfp_countable.ML Thu Jul 16 12:23:22 2015 +0200 @@ -23,17 +23,17 @@ val countableS = @{sort countable}; -fun nchotomy_tac nchotomy = - HEADGOAL (rtac (nchotomy RS @{thm all_reg[rotated]}) THEN' +fun nchotomy_tac ctxt nchotomy = + HEADGOAL (rtac ctxt (nchotomy RS @{thm all_reg[rotated]}) THEN' REPEAT_ALL_NEW (resolve0_tac [allI, impI] ORELSE' eresolve0_tac [exE, disjE])); -fun meta_spec_mp_tac 0 = K all_tac - | meta_spec_mp_tac depth = - dtac meta_spec THEN' meta_spec_mp_tac (depth - 1) THEN' dtac meta_mp THEN' atac; +fun meta_spec_mp_tac ctxt 0 = K all_tac + | meta_spec_mp_tac ctxt depth = + dtac ctxt meta_spec THEN' meta_spec_mp_tac ctxt (depth - 1) THEN' dtac ctxt meta_mp THEN' atac; -val use_induction_hypothesis_tac = +fun use_induction_hypothesis_tac ctxt = DEEPEN (1, 64 (* large number *)) - (fn depth => meta_spec_mp_tac depth THEN' etac allE THEN' etac impE THEN' atac THEN' atac) 0; + (fn depth => meta_spec_mp_tac ctxt depth THEN' etac ctxt allE THEN' etac ctxt impE THEN' atac THEN' atac) 0; val same_ctr_simps = @{thms sum_encode_eq prod_encode_eq sum.inject prod.inject to_nat_split id_apply snd_conv simp_thms}; @@ -42,29 +42,29 @@ fun same_ctr_tac ctxt injects recs map_congs' inj_map_strongs' = HEADGOAL (asm_full_simp_tac (ss_only (injects @ recs @ map_congs' @ same_ctr_simps) ctxt) THEN_MAYBE' - TRY o REPEAT_ALL_NEW (rtac conjI) THEN_ALL_NEW + TRY o REPEAT_ALL_NEW (rtac ctxt conjI) THEN_ALL_NEW REPEAT_ALL_NEW (eresolve_tac ctxt (conjE :: inj_map_strongs')) THEN_ALL_NEW - (atac ORELSE' use_induction_hypothesis_tac)); + (atac ORELSE' use_induction_hypothesis_tac ctxt)); fun distinct_ctrs_tac ctxt recs = HEADGOAL (asm_full_simp_tac (ss_only (recs @ distinct_ctrs_simps) ctxt)); fun mk_encode_injective_tac ctxt n nchotomy injects recs map_comps' inj_map_strongs' = let val ks = 1 upto n in - EVERY (maps (fn k => nchotomy_tac nchotomy :: map (fn k' => + EVERY (maps (fn k => nchotomy_tac ctxt nchotomy :: map (fn k' => if k = k' then same_ctr_tac ctxt injects recs map_comps' inj_map_strongs' else distinct_ctrs_tac ctxt recs) ks) ks) end; fun mk_encode_injectives_tac ctxt ns induct nchotomys injectss recss map_comps' inj_map_strongs' = - HEADGOAL (rtac induct) THEN + HEADGOAL (rtac ctxt induct) THEN EVERY (@{map 4} (fn n => fn nchotomy => fn injects => fn recs => mk_encode_injective_tac ctxt n nchotomy injects recs map_comps' inj_map_strongs') ns nchotomys injectss recss); fun endgame_tac ctxt encode_injectives = unfold_thms_tac ctxt @{thms inj_on_def ball_UNIV} THEN - ALLGOALS (rtac exI THEN' rtac allI THEN' resolve_tac ctxt encode_injectives); + ALLGOALS (rtac ctxt exI THEN' rtac ctxt allI THEN' resolve_tac ctxt encode_injectives); fun encode_sumN n k t = Balanced_Tree.access {init = t, diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_comp.ML --- a/src/HOL/Tools/BNF/bnf_comp.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_comp.ML Thu Jul 16 12:23:22 2015 +0200 @@ -253,12 +253,12 @@ else (bd, NONE); - fun map_id0_tac _ = - mk_comp_map_id0_tac (map_id0_of_bnf outer) (map_cong0_of_bnf outer) + fun map_id0_tac ctxt = + mk_comp_map_id0_tac ctxt (map_id0_of_bnf outer) (map_cong0_of_bnf outer) (map map_id0_of_bnf inners); - fun map_comp0_tac _ = - mk_comp_map_comp0_tac (map_comp0_of_bnf outer) (map_cong0_of_bnf outer) + fun map_comp0_tac ctxt = + mk_comp_map_comp0_tac ctxt (map_comp0_of_bnf outer) (map_cong0_of_bnf outer) (map map_comp0_of_bnf inners); fun mk_single_set_map0_tac i ctxt = @@ -268,11 +268,11 @@ val set_map0_tacs = map mk_single_set_map0_tac (0 upto ilive - 1); - fun bd_card_order_tac _ = - mk_comp_bd_card_order_tac (map bd_card_order_of_bnf inners) (bd_card_order_of_bnf outer); + fun bd_card_order_tac ctxt = + mk_comp_bd_card_order_tac ctxt (map bd_card_order_of_bnf inners) (bd_card_order_of_bnf outer); - fun bd_cinfinite_tac _ = - mk_comp_bd_cinfinite_tac (bd_cinfinite_of_bnf inner) (bd_cinfinite_of_bnf outer); + fun bd_cinfinite_tac ctxt = + mk_comp_bd_cinfinite_tac ctxt (bd_cinfinite_of_bnf inner) (bd_cinfinite_of_bnf outer); val set_alt_thms = if Config.get lthy quick_and_dirty then @@ -319,7 +319,7 @@ |> Thm.close_derivation end; - fun le_rel_OO_tac _ = mk_le_rel_OO_tac (le_rel_OO_of_bnf outer) (rel_mono_of_bnf outer) + fun le_rel_OO_tac ctxt = mk_le_rel_OO_tac ctxt (le_rel_OO_of_bnf outer) (rel_mono_of_bnf outer) (map le_rel_OO_of_bnf inners); fun rel_OO_Grp_tac ctxt = @@ -334,7 +334,7 @@ trans OF [rel_OO_of_bnf outer RS sym, outer_rel_cong OF (map (fn bnf => rel_OO_Grp_of_bnf bnf RS sym) inners)]]] RS sym); in - unfold_thms_tac ctxt set'_eq_sets THEN rtac thm 1 + unfold_thms_tac ctxt set'_eq_sets THEN rtac ctxt thm 1 end; val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac @@ -408,16 +408,16 @@ val bd = mk_bd_of_bnf Ds As bnf; - fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1; + fun map_id0_tac ctxt = rtac ctxt (map_id0_of_bnf bnf) 1; fun map_comp0_tac ctxt = unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) :: - @{thms comp_assoc id_comp comp_id}) THEN rtac refl 1; + @{thms comp_assoc id_comp comp_id}) THEN rtac ctxt refl 1; fun map_cong0_tac ctxt = mk_kill_map_cong0_tac ctxt n (live - n) (map_cong0_of_bnf bnf); - val set_map0_tacs = map (fn thm => fn _ => rtac thm 1) (drop n (set_map0_of_bnf bnf)); - fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1; - fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1; - val set_bd_tacs = map (fn thm => fn _ => rtac thm 1) (drop n (set_bd_of_bnf bnf)); + val set_map0_tacs = map (fn thm => fn ctxt => rtac ctxt thm 1) (drop n (set_map0_of_bnf bnf)); + fun bd_card_order_tac ctxt = rtac ctxt (bd_card_order_of_bnf bnf) 1; + fun bd_cinfinite_tac ctxt = rtac ctxt (bd_cinfinite_of_bnf bnf) 1; + val set_bd_tacs = map (fn thm => fn ctxt => rtac ctxt thm 1) (drop n (set_bd_of_bnf bnf)); val in_alt_thm = let @@ -425,14 +425,15 @@ val in_alt = mk_in (map HOLogic.mk_UNIV killedAs @ Asets) bnf_sets T; val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt)); in - Goal.prove_sorry lthy [] [] goal (K kill_in_alt_tac) |> Thm.close_derivation + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + kill_in_alt_tac ctxt) |> Thm.close_derivation end; fun le_rel_OO_tac ctxt = - EVERY' [rtac @{thm ord_le_eq_trans}, rtac (le_rel_OO_of_bnf bnf)] 1 THEN - unfold_thms_tac ctxt @{thms eq_OO} THEN rtac refl 1; + EVERY' [rtac ctxt @{thm ord_le_eq_trans}, rtac ctxt (le_rel_OO_of_bnf bnf)] 1 THEN + unfold_thms_tac ctxt @{thms eq_OO} THEN rtac ctxt refl 1; - fun rel_OO_Grp_tac _ = + fun rel_OO_Grp_tac ctxt = let val rel_Grp = rel_Grp_of_bnf bnf RS sym val thm = @@ -444,7 +445,7 @@ (replicate n @{thm trans[OF Grp_UNIV_id[OF refl] eq_alt[symmetric]]} @ replicate (live - n) @{thm Grp_fst_snd})]]] RS sym); in - rtac thm 1 + rtac ctxt thm 1 end; val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac @@ -506,26 +507,26 @@ val bd = mk_bd_of_bnf Ds As bnf; - fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1; + fun map_id0_tac ctxt = rtac ctxt (map_id0_of_bnf bnf) 1; fun map_comp0_tac ctxt = unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) :: - @{thms comp_assoc id_comp comp_id}) THEN rtac refl 1; + @{thms comp_assoc id_comp comp_id}) THEN rtac ctxt refl 1; fun map_cong0_tac ctxt = - rtac (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1); + rtac ctxt (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1); val set_map0_tacs = if Config.get lthy quick_and_dirty then replicate (n + live) (K all_tac) else - replicate n (K empty_natural_tac) @ - map (fn thm => fn _ => rtac thm 1) (set_map0_of_bnf bnf); - fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1; - fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1; + replicate n empty_natural_tac @ + map (fn thm => fn ctxt => rtac ctxt thm 1) (set_map0_of_bnf bnf); + fun bd_card_order_tac ctxt = rtac ctxt (bd_card_order_of_bnf bnf) 1; + fun bd_cinfinite_tac ctxt = rtac ctxt (bd_cinfinite_of_bnf bnf) 1; val set_bd_tacs = if Config.get lthy quick_and_dirty then replicate (n + live) (K all_tac) else - replicate n (K (mk_lift_set_bd_tac (bd_Card_order_of_bnf bnf))) @ - (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf)); + replicate n (fn ctxt => mk_lift_set_bd_tac ctxt (bd_Card_order_of_bnf bnf)) @ + (map (fn thm => fn ctxt => rtac ctxt thm 1) (set_bd_of_bnf bnf)); val in_alt_thm = let @@ -533,12 +534,13 @@ val in_alt = mk_in (drop n Asets) bnf_sets T; val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt)); in - Goal.prove_sorry lthy [] [] goal (K lift_in_alt_tac) |> Thm.close_derivation + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => lift_in_alt_tac ctxt) + |> Thm.close_derivation end; - fun le_rel_OO_tac _ = rtac (le_rel_OO_of_bnf bnf) 1; + fun le_rel_OO_tac ctxt = rtac ctxt (le_rel_OO_of_bnf bnf) 1; - fun rel_OO_Grp_tac _ = mk_simple_rel_OO_Grp_tac (rel_OO_Grp_of_bnf bnf) in_alt_thm; + fun rel_OO_Grp_tac ctxt = mk_simple_rel_OO_Grp_tac ctxt (rel_OO_Grp_of_bnf bnf) in_alt_thm; val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac; @@ -598,14 +600,14 @@ val bd = mk_bd_of_bnf Ds As bnf; - fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1; - fun map_comp0_tac _ = rtac (map_comp0_of_bnf bnf) 1; + fun map_id0_tac ctxt = rtac ctxt (map_id0_of_bnf bnf) 1; + fun map_comp0_tac ctxt = rtac ctxt (map_comp0_of_bnf bnf) 1; fun map_cong0_tac ctxt = - rtac (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1); - val set_map0_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_map0_of_bnf bnf)); - fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1; - fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1; - val set_bd_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf)); + rtac ctxt (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1); + val set_map0_tacs = permute (map (fn thm => fn ctxt => rtac ctxt thm 1) (set_map0_of_bnf bnf)); + fun bd_card_order_tac ctxt = rtac ctxt (bd_card_order_of_bnf bnf) 1; + fun bd_cinfinite_tac ctxt = rtac ctxt (bd_cinfinite_of_bnf bnf) 1; + val set_bd_tacs = permute (map (fn thm => fn ctxt => rtac ctxt thm 1) (set_bd_of_bnf bnf)); val in_alt_thm = let @@ -613,13 +615,14 @@ val in_alt = mk_in (unpermute Asets) bnf_sets T; val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt)); in - Goal.prove_sorry lthy [] [] goal (K (mk_permute_in_alt_tac src dest)) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_permute_in_alt_tac ctxt src dest) |> Thm.close_derivation end; - fun le_rel_OO_tac _ = rtac (le_rel_OO_of_bnf bnf) 1; + fun le_rel_OO_tac ctxt = rtac ctxt (le_rel_OO_of_bnf bnf) 1; - fun rel_OO_Grp_tac _ = mk_simple_rel_OO_Grp_tac (rel_OO_Grp_of_bnf bnf) in_alt_thm; + fun rel_OO_Grp_tac ctxt = mk_simple_rel_OO_Grp_tac ctxt (rel_OO_Grp_of_bnf bnf) in_alt_thm; val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac; @@ -796,7 +799,7 @@ else inter (op =) (Term.add_tfreesT repTA []) all_TA_params_in_order); val ((TA, (Rep_name, Abs_name, type_definition, Abs_inverse, Abs_inject, _)), lthy) = maybe_typedef lthy has_live_phantoms (T_bind, TA_params, NoSyn) (HOLogic.mk_UNIV repTA) NONE - (fn _ => EVERY' [rtac exI, rtac UNIV_I] 1) lthy; + (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val repTB = mk_T_of_bnf Ds Bs bnf; val TB = Term.typ_subst_atomic (As ~~ Bs) TA; @@ -826,7 +829,7 @@ val ((bdT, (_, Abs_bd_name, _, _, Abs_bdT_inject, Abs_bdT_cases)), lthy) = maybe_typedef lthy false (bdT_bind, params, NoSyn) - (HOLogic.mk_UNIV bd_repT) NONE (fn _ => EVERY' [rtac exI, rtac UNIV_I] 1) lthy; + (HOLogic.mk_UNIV bd_repT) NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val (bnf_bd', bd_ordIso, bd_card_order, bd_cinfinite) = if bdT = bd_repT then (bnf_bd, bd_Card_order_of_bnf bnf RS @{thm ordIso_refl}, @@ -847,31 +850,32 @@ (bnf_bd', bd_ordIso, bd_card_order, bd_cinfinite) end; - fun map_id0_tac _ = - rtac (@{thm type_copy_map_id0} OF [type_definition, map_id0_of_bnf bnf]) 1; - fun map_comp0_tac _ = - rtac (@{thm type_copy_map_comp0} OF [type_definition, map_comp0_of_bnf bnf]) 1; - fun map_cong0_tac _ = - EVERY' (rtac @{thm type_copy_map_cong0} :: rtac (map_cong0_of_bnf bnf) :: - map (fn i => EVERY' [select_prem_tac live (dtac meta_spec) i, etac meta_mp, - etac (o_apply RS equalityD2 RS set_mp)]) (1 upto live)) 1; - fun set_map0_tac thm _ = - rtac (@{thm type_copy_set_map0} OF [type_definition, thm]) 1; - val set_bd_tacs = map (fn thm => fn _ => rtac (@{thm ordLeq_ordIso_trans} OF + fun map_id0_tac ctxt = + rtac ctxt (@{thm type_copy_map_id0} OF [type_definition, map_id0_of_bnf bnf]) 1; + fun map_comp0_tac ctxt = + rtac ctxt (@{thm type_copy_map_comp0} OF [type_definition, map_comp0_of_bnf bnf]) 1; + fun map_cong0_tac ctxt = + EVERY' (rtac ctxt @{thm type_copy_map_cong0} :: rtac ctxt (map_cong0_of_bnf bnf) :: + map (fn i => EVERY' [select_prem_tac ctxt live (dtac ctxt meta_spec) i, etac ctxt meta_mp, + etac ctxt (o_apply RS equalityD2 RS set_mp)]) (1 upto live)) 1; + fun set_map0_tac thm ctxt = + rtac ctxt (@{thm type_copy_set_map0} OF [type_definition, thm]) 1; + val set_bd_tacs = map (fn thm => fn ctxt => rtac ctxt (@{thm ordLeq_ordIso_trans} OF [thm, bd_ordIso] RS @{thm type_copy_set_bd}) 1) (set_bd_of_bnf bnf); - fun le_rel_OO_tac _ = - rtac (le_rel_OO_of_bnf bnf RS @{thm vimage2p_relcompp_mono}) 1; + fun le_rel_OO_tac ctxt = + rtac ctxt (le_rel_OO_of_bnf bnf RS @{thm vimage2p_relcompp_mono}) 1; fun rel_OO_Grp_tac ctxt = - (rtac (rel_OO_Grp_of_bnf bnf RS @{thm vimage2p_cong} RS trans) THEN' + (rtac ctxt (rel_OO_Grp_of_bnf bnf RS @{thm vimage2p_cong} RS trans) THEN' (if has_live_phantoms then subst_tac ctxt NONE else SELECT_GOAL o unfold_thms_tac ctxt) [type_definition RS @{thm vimage2p_relcompp_converse}] THEN' SELECT_GOAL (unfold_thms_tac ctxt [o_apply, type_definition RS @{thm type_copy_vimage2p_Grp_Rep}, type_definition RS @{thm vimage2p_relcompp_converse}]) THEN' - rtac refl) 1; + rtac ctxt refl) 1; val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac - (map set_map0_tac (set_map0_of_bnf bnf)) (K (rtac bd_card_order 1)) (K (rtac bd_cinfinite 1)) + (map set_map0_tac (set_map0_of_bnf bnf)) + (fn ctxt => rtac ctxt bd_card_order 1) (fn ctxt => rtac ctxt bd_cinfinite 1) set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac; val bnf_wits = map (fn (I, t) => @@ -879,7 +883,7 @@ (AbsA $ Term.list_comb (t, map Bound (0 upto length I - 1)))) (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf); - fun wit_tac _ = ALLGOALS (dtac (type_definition RS @{thm type_copy_wit})) THEN + fun wit_tac ctxt = ALLGOALS (dtac ctxt (type_definition RS @{thm type_copy_wit})) THEN mk_simple_wit_tac (wit_thms_of_bnf bnf); val (bnf', lthy') = diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_comp_tactics.ML --- a/src/HOL/Tools/BNF/bnf_comp_tactics.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_comp_tactics.ML Thu Jul 16 12:23:22 2015 +0200 @@ -8,28 +8,28 @@ signature BNF_COMP_TACTICS = sig - val mk_comp_bd_card_order_tac: thm list -> thm -> tactic - val mk_comp_bd_cinfinite_tac: thm -> thm -> tactic + val mk_comp_bd_card_order_tac: Proof.context -> thm list -> thm -> tactic + val mk_comp_bd_cinfinite_tac: Proof.context -> thm -> thm -> tactic val mk_comp_in_alt_tac: Proof.context -> thm list -> tactic - val mk_comp_map_comp0_tac: thm -> thm -> thm list -> tactic + val mk_comp_map_comp0_tac: Proof.context -> thm -> thm -> thm list -> tactic val mk_comp_map_cong0_tac: Proof.context -> thm list -> thm list -> thm -> thm list -> tactic - val mk_comp_map_id0_tac: thm -> thm -> thm list -> tactic + val mk_comp_map_id0_tac: Proof.context -> thm -> thm -> thm list -> tactic val mk_comp_set_alt_tac: Proof.context -> thm -> tactic val mk_comp_set_bd_tac: Proof.context -> thm -> thm option -> thm -> thm list -> tactic val mk_comp_set_map0_tac: Proof.context -> thm -> thm -> thm -> thm -> thm list -> tactic val mk_comp_wit_tac: Proof.context -> thm list -> thm list -> thm -> thm list -> tactic - val kill_in_alt_tac: tactic + val kill_in_alt_tac: Proof.context -> tactic val mk_kill_map_cong0_tac: Proof.context -> int -> int -> thm -> tactic - val empty_natural_tac: tactic - val lift_in_alt_tac: tactic - val mk_lift_set_bd_tac: thm -> tactic + val empty_natural_tac: Proof.context -> tactic + val lift_in_alt_tac: Proof.context -> tactic + val mk_lift_set_bd_tac: Proof.context -> thm -> tactic - val mk_permute_in_alt_tac: ''a list -> ''a list -> tactic + val mk_permute_in_alt_tac: Proof.context -> ''a list -> ''a list -> tactic - val mk_le_rel_OO_tac: thm -> thm -> thm list -> tactic - val mk_simple_rel_OO_Grp_tac: thm -> thm -> tactic + val mk_le_rel_OO_tac: Proof.context -> thm -> thm -> thm list -> tactic + val mk_simple_rel_OO_Grp_tac: Proof.context -> thm -> thm -> tactic val mk_simple_wit_tac: thm list -> tactic val mk_simplified_set_tac: Proof.context -> thm -> tactic val bd_ordIso_natLeq_tac: tactic @@ -52,96 +52,96 @@ fun mk_comp_set_alt_tac ctxt collect_set_map = unfold_thms_tac ctxt @{thms comp_assoc} THEN unfold_thms_tac ctxt [collect_set_map RS sym] THEN - rtac refl 1; + rtac ctxt refl 1; -fun mk_comp_map_id0_tac Gmap_id0 Gmap_cong0 map_id0s = - EVERY' ([rtac @{thm ext}, rtac (Gmap_cong0 RS trans)] @ - map (fn thm => rtac (thm RS fun_cong)) map_id0s @ [rtac (Gmap_id0 RS fun_cong)]) 1; +fun mk_comp_map_id0_tac ctxt Gmap_id0 Gmap_cong0 map_id0s = + EVERY' ([rtac ctxt @{thm ext}, rtac ctxt (Gmap_cong0 RS trans)] @ + map (fn thm => rtac ctxt (thm RS fun_cong)) map_id0s @ [rtac ctxt (Gmap_id0 RS fun_cong)]) 1; -fun mk_comp_map_comp0_tac Gmap_comp0 Gmap_cong0 map_comp0s = - EVERY' ([rtac @{thm ext}, rtac sym, rtac trans_o_apply, - rtac (Gmap_comp0 RS sym RS comp_eq_dest_lhs RS trans), rtac Gmap_cong0] @ - map (fn thm => rtac (thm RS sym RS fun_cong)) map_comp0s) 1; +fun mk_comp_map_comp0_tac ctxt Gmap_comp0 Gmap_cong0 map_comp0s = + EVERY' ([rtac ctxt @{thm ext}, rtac ctxt sym, rtac ctxt trans_o_apply, + rtac ctxt (Gmap_comp0 RS sym RS comp_eq_dest_lhs RS trans), rtac ctxt Gmap_cong0] @ + map (fn thm => rtac ctxt (thm RS sym RS fun_cong)) map_comp0s) 1; fun mk_comp_set_map0_tac ctxt set'_eq_set Gmap_comp0 Gmap_cong0 Gset_map0 set_map0s = unfold_thms_tac ctxt [set'_eq_set] THEN - EVERY' ([rtac @{thm ext}] @ - replicate 3 (rtac trans_o_apply) @ - [rtac (arg_cong_Union RS trans), - rtac (@{thm arg_cong2[of _ _ _ _ collect, OF refl]} RS trans), - rtac (Gmap_comp0 RS sym RS comp_eq_dest_lhs RS trans), - rtac Gmap_cong0] @ - map (fn thm => rtac (thm RS fun_cong)) set_map0s @ - [rtac (Gset_map0 RS comp_eq_dest_lhs), rtac sym, rtac trans_o_apply, - rtac trans_image_cong_o_apply, rtac trans_image_cong_o_apply, - rtac (@{thm image_cong} OF [Gset_map0 RS comp_eq_dest_lhs RS arg_cong_Union, refl] RS trans), - rtac @{thm trans[OF comp_eq_dest[OF Union_natural[symmetric]]]}, rtac arg_cong_Union, - rtac @{thm trans[OF comp_eq_dest_lhs[OF image_o_collect[symmetric]]]}, - rtac @{thm fun_cong[OF arg_cong[of _ _ collect]]}] @ - [REPEAT_DETERM_N (length set_map0s) o EVERY' [rtac @{thm trans[OF image_insert]}, - rtac @{thm arg_cong2[of _ _ _ _ insert]}, rtac @{thm ext}, rtac trans_o_apply, - rtac trans_image_cong_o_apply, rtac @{thm trans[OF image_image]}, - rtac @{thm sym[OF trans[OF o_apply]]}, rtac @{thm image_cong[OF refl o_apply]}], - rtac @{thm image_empty}]) 1; + EVERY' ([rtac ctxt @{thm ext}] @ + replicate 3 (rtac ctxt trans_o_apply) @ + [rtac ctxt (arg_cong_Union RS trans), + rtac ctxt (@{thm arg_cong2[of _ _ _ _ collect, OF refl]} RS trans), + rtac ctxt (Gmap_comp0 RS sym RS comp_eq_dest_lhs RS trans), + rtac ctxt Gmap_cong0] @ + map (fn thm => rtac ctxt (thm RS fun_cong)) set_map0s @ + [rtac ctxt (Gset_map0 RS comp_eq_dest_lhs), rtac ctxt sym, rtac ctxt trans_o_apply, + rtac ctxt trans_image_cong_o_apply, rtac ctxt trans_image_cong_o_apply, + rtac ctxt (@{thm image_cong} OF [Gset_map0 RS comp_eq_dest_lhs RS arg_cong_Union, refl] RS trans), + rtac ctxt @{thm trans[OF comp_eq_dest[OF Union_natural[symmetric]]]}, rtac ctxt arg_cong_Union, + rtac ctxt @{thm trans[OF comp_eq_dest_lhs[OF image_o_collect[symmetric]]]}, + rtac ctxt @{thm fun_cong[OF arg_cong[of _ _ collect]]}] @ + [REPEAT_DETERM_N (length set_map0s) o EVERY' [rtac ctxt @{thm trans[OF image_insert]}, + rtac ctxt @{thm arg_cong2[of _ _ _ _ insert]}, rtac ctxt @{thm ext}, rtac ctxt trans_o_apply, + rtac ctxt trans_image_cong_o_apply, rtac ctxt @{thm trans[OF image_image]}, + rtac ctxt @{thm sym[OF trans[OF o_apply]]}, rtac ctxt @{thm image_cong[OF refl o_apply]}], + rtac ctxt @{thm image_empty}]) 1; fun mk_comp_map_cong0_tac ctxt set'_eq_sets comp_set_alts map_cong0 map_cong0s = let val n = length comp_set_alts; in unfold_thms_tac ctxt set'_eq_sets THEN - (if n = 0 then rtac refl 1 - else rtac map_cong0 1 THEN + (if n = 0 then rtac ctxt refl 1 + else rtac ctxt map_cong0 1 THEN EVERY' (map_index (fn (i, map_cong0) => - rtac map_cong0 THEN' EVERY' (map_index (fn (k, set_alt) => - EVERY' [select_prem_tac n (dtac @{thm meta_spec}) (k + 1), etac meta_mp, - rtac (equalityD2 RS set_mp), rtac (set_alt RS fun_cong RS trans), - rtac trans_o_apply, rtac (@{thm collect_def} RS arg_cong_Union), - rtac @{thm UnionI}, rtac @{thm UN_I}, REPEAT_DETERM_N i o rtac @{thm insertI2}, - rtac @{thm insertI1}, rtac (o_apply RS equalityD2 RS set_mp), - etac @{thm imageI}, atac]) + rtac ctxt map_cong0 THEN' EVERY' (map_index (fn (k, set_alt) => + EVERY' [select_prem_tac ctxt n (dtac ctxt @{thm meta_spec}) (k + 1), etac ctxt meta_mp, + rtac ctxt (equalityD2 RS set_mp), rtac ctxt (set_alt RS fun_cong RS trans), + rtac ctxt trans_o_apply, rtac ctxt (@{thm collect_def} RS arg_cong_Union), + rtac ctxt @{thm UnionI}, rtac ctxt @{thm UN_I}, REPEAT_DETERM_N i o rtac ctxt @{thm insertI2}, + rtac ctxt @{thm insertI1}, rtac ctxt (o_apply RS equalityD2 RS set_mp), + etac ctxt @{thm imageI}, atac]) comp_set_alts)) map_cong0s) 1) end; -fun mk_comp_bd_card_order_tac Fbd_card_orders Gbd_card_order = - rtac @{thm natLeq_card_order} 1 ORELSE +fun mk_comp_bd_card_order_tac ctxt Fbd_card_orders Gbd_card_order = + rtac ctxt @{thm natLeq_card_order} 1 ORELSE let val (card_orders, last_card_order) = split_last Fbd_card_orders; - fun gen_before thm = rtac @{thm card_order_csum} THEN' rtac thm; + fun gen_before thm = rtac ctxt @{thm card_order_csum} THEN' rtac ctxt thm; in - (rtac @{thm card_order_cprod} THEN' - WRAP' gen_before (K (K all_tac)) card_orders (rtac last_card_order) THEN' - rtac Gbd_card_order) 1 + (rtac ctxt @{thm card_order_cprod} THEN' + WRAP' gen_before (K (K all_tac)) card_orders (rtac ctxt last_card_order) THEN' + rtac ctxt Gbd_card_order) 1 end; -fun mk_comp_bd_cinfinite_tac Fbd_cinfinite Gbd_cinfinite = - (rtac @{thm natLeq_cinfinite} ORELSE' - rtac @{thm cinfinite_cprod} THEN' - ((K (TRY ((rtac @{thm cinfinite_csum} THEN' rtac disjI1) 1)) THEN' - ((rtac @{thm cinfinite_csum} THEN' rtac disjI1 THEN' rtac Fbd_cinfinite) ORELSE' - rtac Fbd_cinfinite)) ORELSE' - rtac Fbd_cinfinite) THEN' - rtac Gbd_cinfinite) 1; +fun mk_comp_bd_cinfinite_tac ctxt Fbd_cinfinite Gbd_cinfinite = + (rtac ctxt @{thm natLeq_cinfinite} ORELSE' + rtac ctxt @{thm cinfinite_cprod} THEN' + ((K (TRY ((rtac ctxt @{thm cinfinite_csum} THEN' rtac ctxt disjI1) 1)) THEN' + ((rtac ctxt @{thm cinfinite_csum} THEN' rtac ctxt disjI1 THEN' rtac ctxt Fbd_cinfinite) ORELSE' + rtac ctxt Fbd_cinfinite)) ORELSE' + rtac ctxt Fbd_cinfinite) THEN' + rtac ctxt Gbd_cinfinite) 1; fun mk_comp_set_bd_tac ctxt set'_eq_set bd_ordIso_natLeq_opt comp_set_alt Gset_Fset_bds = let val (bds, last_bd) = split_last Gset_Fset_bds; fun gen_before bd = - rtac ctrans THEN' rtac @{thm Un_csum} THEN' - rtac ctrans THEN' rtac @{thm csum_mono} THEN' - rtac bd; - fun gen_after _ = rtac @{thm ordIso_imp_ordLeq} THEN' rtac @{thm cprod_csum_distrib1}; + rtac ctxt ctrans THEN' rtac ctxt @{thm Un_csum} THEN' + rtac ctxt ctrans THEN' rtac ctxt @{thm csum_mono} THEN' + rtac ctxt bd; + fun gen_after _ = rtac ctxt @{thm ordIso_imp_ordLeq} THEN' rtac ctxt @{thm cprod_csum_distrib1}; in (case bd_ordIso_natLeq_opt of - SOME thm => rtac (thm RSN (2, @{thm ordLeq_ordIso_trans})) 1 + SOME thm => rtac ctxt (thm RSN (2, @{thm ordLeq_ordIso_trans})) 1 | NONE => all_tac) THEN unfold_thms_tac ctxt [set'_eq_set, comp_set_alt] THEN - rtac @{thm comp_set_bd_Union_o_collect} 1 THEN + rtac ctxt @{thm comp_set_bd_Union_o_collect} 1 THEN unfold_thms_tac ctxt @{thms Union_image_insert Union_image_empty Union_Un_distrib o_apply} THEN - (rtac ctrans THEN' - WRAP' gen_before gen_after bds (rtac last_bd) THEN' - rtac @{thm ordIso_imp_ordLeq} THEN' - rtac @{thm cprod_com}) 1 + (rtac ctxt ctrans THEN' + WRAP' gen_before gen_after bds (rtac ctxt last_bd) THEN' + rtac ctxt @{thm ordIso_imp_ordLeq} THEN' + rtac ctxt @{thm cprod_com}) 1 end; val comp_in_alt_thms = @{thms o_apply collect_def image_insert image_empty Union_insert UN_insert @@ -152,84 +152,85 @@ unfold_thms_tac ctxt comp_set_alts THEN unfold_thms_tac ctxt comp_in_alt_thms THEN unfold_thms_tac ctxt @{thms set_eq_subset} THEN - rtac conjI 1 THEN + rtac ctxt conjI 1 THEN REPEAT_DETERM ( - rtac @{thm subsetI} 1 THEN + rtac ctxt @{thm subsetI} 1 THEN unfold_thms_tac ctxt @{thms mem_Collect_eq Ball_def} THEN - (REPEAT_DETERM (CHANGED (etac conjE 1)) THEN + (REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN REPEAT_DETERM (CHANGED (( - (rtac conjI THEN' (atac ORELSE' rtac subset_UNIV)) ORELSE' + (rtac ctxt conjI THEN' (atac ORELSE' rtac ctxt subset_UNIV)) ORELSE' atac ORELSE' - (rtac subset_UNIV)) 1)) ORELSE rtac subset_UNIV 1)); + (rtac ctxt subset_UNIV)) 1)) ORELSE rtac ctxt subset_UNIV 1)); val comp_wit_thms = @{thms Union_empty_conv o_apply collect_def UN_insert UN_empty Un_empty_right Union_image_insert Union_image_empty}; fun mk_comp_wit_tac ctxt set'_eq_sets Gwit_thms collect_set_map Fwit_thms = unfold_thms_tac ctxt set'_eq_sets THEN - ALLGOALS (dtac @{thm in_Union_o_assoc}) THEN + ALLGOALS (dtac ctxt @{thm in_Union_o_assoc}) THEN unfold_thms_tac ctxt [collect_set_map] THEN unfold_thms_tac ctxt comp_wit_thms THEN REPEAT_DETERM ((atac ORELSE' REPEAT_DETERM o eresolve_tac ctxt @{thms UnionE UnE} THEN' - etac imageE THEN' TRY o dresolve_tac ctxt Gwit_thms THEN' - (etac FalseE ORELSE' + etac ctxt imageE THEN' TRY o dresolve_tac ctxt Gwit_thms THEN' + (etac ctxt FalseE ORELSE' hyp_subst_tac ctxt THEN' dresolve_tac ctxt Fwit_thms THEN' - (etac FalseE ORELSE' atac))) 1); + (etac ctxt FalseE ORELSE' atac))) 1); (* Kill operation *) fun mk_kill_map_cong0_tac ctxt n m map_cong0 = - (rtac map_cong0 THEN' EVERY' (replicate n (rtac refl)) THEN' + (rtac ctxt map_cong0 THEN' EVERY' (replicate n (rtac ctxt refl)) THEN' EVERY' (replicate m (Goal.assume_rule_tac ctxt))) 1; -val kill_in_alt_tac = - ((rtac @{thm Collect_cong} THEN' rtac iffI) 1 THEN - REPEAT_DETERM (CHANGED (etac conjE 1)) THEN - REPEAT_DETERM (CHANGED ((etac conjI ORELSE' - rtac conjI THEN' rtac subset_UNIV) 1)) THEN - (rtac subset_UNIV ORELSE' atac) 1 THEN - REPEAT_DETERM (CHANGED (etac conjE 1)) THEN - REPEAT_DETERM (CHANGED ((etac conjI ORELSE' atac) 1))) ORELSE - ((rtac @{thm UNIV_eq_I} THEN' rtac CollectI) 1 THEN - REPEAT_DETERM (TRY (rtac conjI 1) THEN rtac subset_UNIV 1)); +fun kill_in_alt_tac ctxt = + ((rtac ctxt @{thm Collect_cong} THEN' rtac ctxt iffI) 1 THEN + REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN + REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE' + rtac ctxt conjI THEN' rtac ctxt subset_UNIV) 1)) THEN + (rtac ctxt subset_UNIV ORELSE' atac) 1 THEN + REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN + REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE' atac) 1))) ORELSE + ((rtac ctxt @{thm UNIV_eq_I} THEN' rtac ctxt CollectI) 1 THEN + REPEAT_DETERM (TRY (rtac ctxt conjI 1) THEN rtac ctxt subset_UNIV 1)); (* Lift operation *) -val empty_natural_tac = rtac @{thm empty_natural} 1; +fun empty_natural_tac ctxt = rtac ctxt @{thm empty_natural} 1; -fun mk_lift_set_bd_tac bd_Card_order = (rtac @{thm Card_order_empty} THEN' rtac bd_Card_order) 1; +fun mk_lift_set_bd_tac ctxt bd_Card_order = + (rtac ctxt @{thm Card_order_empty} THEN' rtac ctxt bd_Card_order) 1; -val lift_in_alt_tac = - ((rtac @{thm Collect_cong} THEN' rtac iffI) 1 THEN - REPEAT_DETERM (CHANGED (etac conjE 1)) THEN - REPEAT_DETERM (CHANGED ((etac conjI ORELSE' atac) 1)) THEN - REPEAT_DETERM (CHANGED (etac conjE 1)) THEN - REPEAT_DETERM (CHANGED ((etac conjI ORELSE' - rtac conjI THEN' rtac @{thm empty_subsetI}) 1)) THEN - (rtac @{thm empty_subsetI} ORELSE' atac) 1) ORELSE - ((rtac sym THEN' rtac @{thm UNIV_eq_I} THEN' rtac CollectI) 1 THEN - REPEAT_DETERM (TRY (rtac conjI 1) THEN rtac @{thm empty_subsetI} 1)); +fun lift_in_alt_tac ctxt = + ((rtac ctxt @{thm Collect_cong} THEN' rtac ctxt iffI) 1 THEN + REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN + REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE' atac) 1)) THEN + REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN + REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE' + rtac ctxt conjI THEN' rtac ctxt @{thm empty_subsetI}) 1)) THEN + (rtac ctxt @{thm empty_subsetI} ORELSE' atac) 1) ORELSE + ((rtac ctxt sym THEN' rtac ctxt @{thm UNIV_eq_I} THEN' rtac ctxt CollectI) 1 THEN + REPEAT_DETERM (TRY (rtac ctxt conjI 1) THEN rtac ctxt @{thm empty_subsetI} 1)); (* Permute operation *) -fun mk_permute_in_alt_tac src dest = - (rtac @{thm Collect_cong} THEN' - mk_rotate_eq_tac (rtac refl) trans @{thm conj_assoc} @{thm conj_commute} @{thm conj_cong} +fun mk_permute_in_alt_tac ctxt src dest = + (rtac ctxt @{thm Collect_cong} THEN' + mk_rotate_eq_tac ctxt (rtac ctxt refl) trans @{thm conj_assoc} @{thm conj_commute} @{thm conj_cong} dest src) 1; (* Miscellaneous *) -fun mk_le_rel_OO_tac outer_le_rel_OO outer_rel_mono inner_le_rel_OOs = - EVERY' (map rtac (@{thm order_trans} :: outer_le_rel_OO :: outer_rel_mono :: inner_le_rel_OOs)) 1; +fun mk_le_rel_OO_tac ctxt outer_le_rel_OO outer_rel_mono inner_le_rel_OOs = + EVERY' (map (rtac ctxt) (@{thm order_trans} :: outer_le_rel_OO :: outer_rel_mono :: inner_le_rel_OOs)) 1; -fun mk_simple_rel_OO_Grp_tac rel_OO_Grp in_alt_thm = - rtac (trans OF [rel_OO_Grp, in_alt_thm RS @{thm OO_Grp_cong} RS sym]) 1; +fun mk_simple_rel_OO_Grp_tac ctxt rel_OO_Grp in_alt_thm = + rtac ctxt (trans OF [rel_OO_Grp, in_alt_thm RS @{thm OO_Grp_cong} RS sym]) 1; fun mk_simple_wit_tac wit_thms = ALLGOALS (atac ORELSE' eresolve0_tac (@{thm emptyE} :: wit_thms)); @@ -244,7 +245,7 @@ fun mk_simplified_set_tac ctxt collect_set_map = unfold_thms_tac ctxt (collect_set_map :: @{thms comp_assoc}) THEN - unfold_thms_tac ctxt simplified_set_simps THEN rtac refl 1; + unfold_thms_tac ctxt simplified_set_simps THEN rtac ctxt refl 1; val bd_ordIso_natLeq_tac = HEADGOAL (REPEAT_DETERM o resolve0_tac diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_def.ML --- a/src/HOL/Tools/BNF/bnf_def.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_def.ML Thu Jul 16 12:23:22 2015 +0200 @@ -1154,7 +1154,8 @@ (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*) val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect)); in - Goal.prove_sorry lthy [] [] goal (K (mk_collect_set_map_tac (#set_map0 axioms))) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_collect_set_map_tac ctxt (#set_map0 axioms)) |> Thm.close_derivation end; @@ -1168,7 +1169,8 @@ (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop (mk_leq (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA')))); in - Goal.prove_sorry lthy [] [] in_mono_goal (K (mk_in_mono_tac live)) + Goal.prove_sorry lthy [] [] in_mono_goal (fn {context = ctxt, prems = _} => + mk_in_mono_tac ctxt live) |> Thm.close_derivation end; @@ -1183,7 +1185,7 @@ mk_Trueprop_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA'))); in Goal.prove_sorry lthy [] [] in_cong_goal - (K ((TRY o hyp_subst_tac lthy THEN' rtac refl) 1)) + (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation end; @@ -1217,8 +1219,9 @@ val concl = HOLogic.mk_Trueprop (mk_inj (Term.list_comb (bnf_map_AsBs, fs))); val goal = fold_rev Logic.all fs (Logic.list_implies (prems, concl)); in - Goal.prove_sorry lthy [] [] goal (fn _ => mk_inj_map_tac live (Lazy.force map_id) - (Lazy.force map_comp) (#map_cong0 axioms) (Lazy.force map_cong)) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_inj_map_tac ctxt live (Lazy.force map_id) (Lazy.force map_comp) (#map_cong0 axioms) + (Lazy.force map_cong)) |> Thm.close_derivation end; @@ -1295,7 +1298,8 @@ in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl))) - (K (mk_rel_mono_tac rel_OO_Grps (Lazy.force in_mono))) + (fn {context = ctxt, prems = _} => + mk_rel_mono_tac ctxt rel_OO_Grps (Lazy.force in_mono)) |> Thm.close_derivation end; @@ -1306,7 +1310,7 @@ in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl))) - (fn _ => (TRY o hyp_subst_tac lthy THEN' rtac refl) 1) + (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation end; @@ -1317,7 +1321,8 @@ Goal.prove_sorry lthy [] [] (mk_Trueprop_eq (Term.list_comb (relAsAs, map HOLogic.eq_const As'), HOLogic.eq_const CA')) - (K (mk_rel_eq_tac live (Lazy.force rel_Grp) (Lazy.force rel_cong) (#map_id0 axioms))) + (fn {context = ctxt, prems = _} => + mk_rel_eq_tac ctxt live (Lazy.force rel_Grp) (Lazy.force rel_cong) (#map_id0 axioms)) |> Thm.close_derivation; val rel_eq = Lazy.lazy mk_rel_eq; @@ -1336,7 +1341,8 @@ val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs)); in Goal.prove_sorry lthy [] [] goal - (K (mk_rel_conversep_tac le_thm (Lazy.force rel_mono))) + (fn {context = ctxt, prems = _} => + mk_rel_conversep_tac ctxt le_thm (Lazy.force rel_mono)) |> Thm.close_derivation end; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_def_tactics.ML --- a/src/HOL/Tools/BNF/bnf_def_tactics.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_def_tactics.ML Thu Jul 16 12:23:22 2015 +0200 @@ -9,10 +9,10 @@ signature BNF_DEF_TACTICS = sig - val mk_collect_set_map_tac: thm list -> tactic - val mk_in_mono_tac: int -> tactic + val mk_collect_set_map_tac: Proof.context -> thm list -> tactic + val mk_in_mono_tac: Proof.context -> int -> tactic val mk_inj_map_strong_tac: Proof.context -> thm -> thm list -> thm -> tactic - val mk_inj_map_tac: int -> thm -> thm -> thm -> thm -> tactic + val mk_inj_map_tac: Proof.context -> int -> thm -> thm -> thm -> thm -> tactic val mk_map_id: thm -> thm val mk_map_ident: Proof.context -> thm -> thm val mk_map_comp: thm -> thm @@ -21,12 +21,12 @@ val mk_set_transfer_tac: Proof.context -> thm -> thm list -> tactic val mk_rel_Grp_tac: Proof.context -> thm list -> thm -> thm -> thm -> thm -> thm list -> tactic - val mk_rel_eq_tac: int -> thm -> thm -> thm -> tactic + val mk_rel_eq_tac: Proof.context -> int -> thm -> thm -> thm -> tactic val mk_rel_OO_le_tac: Proof.context -> thm list -> thm -> thm -> thm -> thm list -> tactic - val mk_rel_conversep_tac: thm -> thm -> tactic + val mk_rel_conversep_tac: Proof.context -> thm -> thm -> tactic val mk_rel_conversep_le_tac: Proof.context -> thm list -> thm -> thm -> thm -> thm list -> tactic val mk_rel_map0_tac: Proof.context -> int -> thm -> thm -> thm -> thm -> tactic - val mk_rel_mono_tac: thm list -> thm -> tactic + val mk_rel_mono_tac: Proof.context -> thm list -> thm -> tactic val mk_rel_mono_strong0_tac: Proof.context -> thm -> thm list -> tactic val mk_rel_transfer_tac: Proof.context -> thm -> thm list -> thm -> tactic @@ -52,24 +52,24 @@ fun mk_map_ident ctxt = unfold_thms ctxt @{thms id_def}; fun mk_map_comp comp = @{thm comp_eq_dest_lhs} OF [mk_sym comp]; fun mk_map_cong_tac ctxt cong0 = - (hyp_subst_tac ctxt THEN' rtac cong0 THEN' - REPEAT_DETERM o (dtac meta_spec THEN' etac meta_mp THEN' atac)) 1; + (hyp_subst_tac ctxt THEN' rtac ctxt cong0 THEN' + REPEAT_DETERM o (dtac ctxt meta_spec THEN' etac ctxt meta_mp THEN' atac)) 1; fun mk_set_map set_map0 = set_map0 RS @{thm comp_eq_dest}; -fun mk_in_mono_tac n = if n = 0 then rtac subset_UNIV 1 - else (rtac subsetI THEN' - rtac CollectI) 1 THEN +fun mk_in_mono_tac ctxt n = if n = 0 then rtac ctxt subset_UNIV 1 + else (rtac ctxt subsetI THEN' + rtac ctxt CollectI) 1 THEN REPEAT_DETERM (eresolve0_tac [CollectE, conjE] 1) THEN REPEAT_DETERM_N (n - 1) - ((rtac conjI THEN' etac subset_trans THEN' atac) 1) THEN - (etac subset_trans THEN' atac) 1; + ((rtac ctxt conjI THEN' etac ctxt subset_trans THEN' atac) 1) THEN + (etac ctxt subset_trans THEN' atac) 1; -fun mk_inj_map_tac n map_id map_comp map_cong0 map_cong = +fun mk_inj_map_tac ctxt n map_id map_comp map_cong0 map_cong = let val map_cong' = map_cong OF (asm_rl :: replicate n refl); val map_cong0' = map_cong0 OF (replicate n @{thm the_inv_f_o_f_id}); in - HEADGOAL (rtac @{thm injI} THEN' etac (map_cong' RS box_equals) THEN' - REPEAT_DETERM_N 2 o (rtac (box_equals OF [map_cong0', map_comp RS sym, map_id]) THEN' + HEADGOAL (rtac ctxt @{thm injI} THEN' etac ctxt (map_cong' RS box_equals) THEN' + REPEAT_DETERM_N 2 o (rtac ctxt (box_equals OF [map_cong0', map_comp RS sym, map_id]) THEN' REPEAT_DETERM_N n o atac)) end; @@ -78,150 +78,150 @@ val rel_eq' = rel_eq RS @{thm predicate2_eqD}; val rel_maps' = map (fn thm => thm RS iffD1) rel_maps; in - HEADGOAL (dtac (rel_eq' RS iffD2) THEN' rtac (rel_eq' RS iffD1)) THEN - EVERY (map (HEADGOAL o dtac) rel_maps') THEN - HEADGOAL (etac rel_mono_strong) THEN + HEADGOAL (dtac ctxt (rel_eq' RS iffD2) THEN' rtac ctxt (rel_eq' RS iffD1)) THEN + EVERY (map (HEADGOAL o dtac ctxt) rel_maps') THEN + HEADGOAL (etac ctxt rel_mono_strong) THEN TRYALL (Goal.assume_rule_tac ctxt) end; -fun mk_collect_set_map_tac set_map0s = - (rtac (@{thm collect_comp} RS trans) THEN' rtac @{thm arg_cong[of _ _ collect]} THEN' +fun mk_collect_set_map_tac ctxt set_map0s = + (rtac ctxt (@{thm collect_comp} RS trans) THEN' rtac ctxt @{thm arg_cong[of _ _ collect]} THEN' EVERY' (map (fn set_map0 => - rtac (mk_trans @{thm image_insert} @{thm arg_cong2[of _ _ _ _ insert]}) THEN' - rtac set_map0) set_map0s) THEN' - rtac @{thm image_empty}) 1; + rtac ctxt (mk_trans @{thm image_insert} @{thm arg_cong2[of _ _ _ _ insert]}) THEN' + rtac ctxt set_map0) set_map0s) THEN' + rtac ctxt @{thm image_empty}) 1; fun mk_rel_Grp_tac ctxt rel_OO_Grps map_id0 map_cong0 map_id map_comp set_maps = let val n = length set_maps; - val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac else rtac (hd rel_OO_Grps RS trans); + val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac else rtac ctxt (hd rel_OO_Grps RS trans); in if null set_maps then unfold_thms_tac ctxt ((map_id0 RS @{thm Grp_UNIV_id}) :: rel_OO_Grps) THEN - rtac @{thm Grp_UNIV_idI[OF refl]} 1 + rtac ctxt @{thm Grp_UNIV_idI[OF refl]} 1 else - EVERY' [rel_OO_Grps_tac, rtac @{thm antisym}, rtac @{thm predicate2I}, + EVERY' [rel_OO_Grps_tac, rtac ctxt @{thm antisym}, rtac ctxt @{thm predicate2I}, REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}], - hyp_subst_tac ctxt, rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac map_cong0, - REPEAT_DETERM_N n o EVERY' [rtac @{thm Collect_split_Grp_eqD}, etac @{thm set_mp}, atac], - rtac CollectI, - CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS ord_eq_le_trans), - rtac @{thm image_subsetI}, rtac @{thm Collect_split_Grp_inD}, etac @{thm set_mp}, atac]) + hyp_subst_tac ctxt, rtac ctxt @{thm GrpI}, rtac ctxt trans, rtac ctxt map_comp, rtac ctxt map_cong0, + REPEAT_DETERM_N n o EVERY' [rtac ctxt @{thm Collect_split_Grp_eqD}, etac ctxt @{thm set_mp}, atac], + rtac ctxt CollectI, + CONJ_WRAP' (fn thm => EVERY' [rtac ctxt (thm RS ord_eq_le_trans), + rtac ctxt @{thm image_subsetI}, rtac ctxt @{thm Collect_split_Grp_inD}, etac ctxt @{thm set_mp}, atac]) set_maps, - rtac @{thm predicate2I}, REPEAT_DETERM o eresolve_tac ctxt [@{thm GrpE}, exE, conjE], + rtac ctxt @{thm predicate2I}, REPEAT_DETERM o eresolve_tac ctxt [@{thm GrpE}, exE, conjE], hyp_subst_tac ctxt, - rtac @{thm relcomppI}, rtac @{thm conversepI}, + rtac ctxt @{thm relcomppI}, rtac ctxt @{thm conversepI}, EVERY' (map2 (fn convol => fn map_id0 => - EVERY' [rtac @{thm GrpI}, - rtac (@{thm box_equals} OF [map_cong0, map_comp RS sym, map_id0]), - REPEAT_DETERM_N n o rtac (convol RS fun_cong), + EVERY' [rtac ctxt @{thm GrpI}, + rtac ctxt (@{thm box_equals} OF [map_cong0, map_comp RS sym, map_id0]), + REPEAT_DETERM_N n o rtac ctxt (convol RS fun_cong), REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], - rtac CollectI, + rtac ctxt CollectI, CONJ_WRAP' (fn thm => - EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI}, - rtac @{thm convol_mem_GrpI}, etac set_mp, atac]) + EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt thm, rtac ctxt @{thm image_subsetI}, + rtac ctxt @{thm convol_mem_GrpI}, etac ctxt set_mp, atac]) set_maps]) @{thms fst_convol snd_convol} [map_id, refl])] 1 end; -fun mk_rel_eq_tac n rel_Grp rel_cong map_id0 = - (EVERY' (rtac (rel_cong RS trans) :: replicate n (rtac @{thm eq_alt})) THEN' - rtac (rel_Grp RSN (2, @{thm box_equals[OF _ sym sym[OF eq_alt]]})) THEN' - (if n = 0 then rtac refl - else EVERY' [rtac @{thm arg_cong2[of _ _ _ _ "Grp"]}, - rtac @{thm equalityI}, rtac subset_UNIV, rtac subsetI, rtac CollectI, - CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto n), rtac map_id0])) 1; +fun mk_rel_eq_tac ctxt n rel_Grp rel_cong map_id0 = + (EVERY' (rtac ctxt (rel_cong RS trans) :: replicate n (rtac ctxt @{thm eq_alt})) THEN' + rtac ctxt (rel_Grp RSN (2, @{thm box_equals[OF _ sym sym[OF eq_alt]]})) THEN' + (if n = 0 then rtac ctxt refl + else EVERY' [rtac ctxt @{thm arg_cong2[of _ _ _ _ "Grp"]}, + rtac ctxt @{thm equalityI}, rtac ctxt subset_UNIV, rtac ctxt subsetI, rtac ctxt CollectI, + CONJ_WRAP' (K (rtac ctxt subset_UNIV)) (1 upto n), rtac ctxt map_id0])) 1; fun mk_rel_map0_tac ctxt live rel_compp rel_conversep rel_Grp map_id = if live = 0 then HEADGOAL (Goal.conjunction_tac) THEN unfold_thms_tac ctxt @{thms id_apply} THEN - ALLGOALS (rtac refl) + ALLGOALS (rtac ctxt refl) else let val ks = 1 upto live; in Goal.conjunction_tac 1 THEN unfold_thms_tac ctxt [rel_compp, rel_conversep, rel_Grp, @{thm vimage2p_Grp}] THEN - TRYALL (EVERY' [rtac iffI, rtac @{thm relcomppI}, rtac @{thm GrpI}, - resolve_tac ctxt [map_id, refl], rtac CollectI, - CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks, rtac @{thm relcomppI}, atac, - rtac @{thm conversepI}, rtac @{thm GrpI}, resolve_tac ctxt [map_id, refl], rtac CollectI, - CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks, + TRYALL (EVERY' [rtac ctxt iffI, rtac ctxt @{thm relcomppI}, rtac ctxt @{thm GrpI}, + resolve_tac ctxt [map_id, refl], rtac ctxt CollectI, + CONJ_WRAP' (K (rtac ctxt @{thm subset_UNIV})) ks, rtac ctxt @{thm relcomppI}, atac, + rtac ctxt @{thm conversepI}, rtac ctxt @{thm GrpI}, resolve_tac ctxt [map_id, refl], rtac ctxt CollectI, + CONJ_WRAP' (K (rtac ctxt @{thm subset_UNIV})) ks, REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE conversepE GrpE}, - dtac (trans OF [sym, map_id]), hyp_subst_tac ctxt, atac]) + dtac ctxt (trans OF [sym, map_id]), hyp_subst_tac ctxt, atac]) end; -fun mk_rel_mono_tac rel_OO_Grps in_mono = +fun mk_rel_mono_tac ctxt rel_OO_Grps in_mono = let val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac - else rtac (hd rel_OO_Grps RS ord_eq_le_trans) THEN' - rtac (hd rel_OO_Grps RS sym RSN (2, ord_le_eq_trans)); + else rtac ctxt (hd rel_OO_Grps RS ord_eq_le_trans) THEN' + rtac ctxt (hd rel_OO_Grps RS sym RSN (2, ord_le_eq_trans)); in - EVERY' [rel_OO_Grps_tac, rtac @{thm relcompp_mono}, rtac @{thm iffD2[OF conversep_mono]}, - rtac @{thm Grp_mono}, rtac in_mono, REPEAT_DETERM o etac @{thm Collect_split_mono}, - rtac @{thm Grp_mono}, rtac in_mono, REPEAT_DETERM o etac @{thm Collect_split_mono}] 1 + EVERY' [rel_OO_Grps_tac, rtac ctxt @{thm relcompp_mono}, rtac ctxt @{thm iffD2[OF conversep_mono]}, + rtac ctxt @{thm Grp_mono}, rtac ctxt in_mono, REPEAT_DETERM o etac ctxt @{thm Collect_split_mono}, + rtac ctxt @{thm Grp_mono}, rtac ctxt in_mono, REPEAT_DETERM o etac ctxt @{thm Collect_split_mono}] 1 end; fun mk_rel_conversep_le_tac ctxt rel_OO_Grps rel_eq map_cong0 map_comp set_maps = let val n = length set_maps; val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac - else rtac (hd rel_OO_Grps RS ord_eq_le_trans) THEN' - rtac (hd rel_OO_Grps RS sym RS @{thm arg_cong[of _ _ conversep]} RSN (2, ord_le_eq_trans)); + else rtac ctxt (hd rel_OO_Grps RS ord_eq_le_trans) THEN' + rtac ctxt (hd rel_OO_Grps RS sym RS @{thm arg_cong[of _ _ conversep]} RSN (2, ord_le_eq_trans)); in - if null set_maps then rtac (rel_eq RS @{thm leq_conversepI}) 1 + if null set_maps then rtac ctxt (rel_eq RS @{thm leq_conversepI}) 1 else - EVERY' [rel_OO_Grps_tac, rtac @{thm predicate2I}, + EVERY' [rel_OO_Grps_tac, rtac ctxt @{thm predicate2I}, REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}], - hyp_subst_tac ctxt, rtac @{thm conversepI}, rtac @{thm relcomppI}, rtac @{thm conversepI}, - EVERY' (map (fn thm => EVERY' [rtac @{thm GrpI}, rtac sym, rtac trans, - rtac map_cong0, REPEAT_DETERM_N n o rtac thm, - rtac (map_comp RS sym), rtac CollectI, - CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS ord_eq_le_trans), - etac @{thm flip_pred}]) set_maps]) [@{thm snd_fst_flip}, @{thm fst_snd_flip}])] 1 + hyp_subst_tac ctxt, rtac ctxt @{thm conversepI}, rtac ctxt @{thm relcomppI}, rtac ctxt @{thm conversepI}, + EVERY' (map (fn thm => EVERY' [rtac ctxt @{thm GrpI}, rtac ctxt sym, rtac ctxt trans, + rtac ctxt map_cong0, REPEAT_DETERM_N n o rtac ctxt thm, + rtac ctxt (map_comp RS sym), rtac ctxt CollectI, + CONJ_WRAP' (fn thm => EVERY' [rtac ctxt (thm RS ord_eq_le_trans), + etac ctxt @{thm flip_pred}]) set_maps]) [@{thm snd_fst_flip}, @{thm fst_snd_flip}])] 1 end; -fun mk_rel_conversep_tac le_conversep rel_mono = - EVERY' [rtac @{thm antisym}, rtac le_conversep, rtac @{thm xt1(6)}, rtac conversep_shift, - rtac le_conversep, rtac @{thm iffD2[OF conversep_mono]}, rtac rel_mono, - REPEAT_DETERM o rtac @{thm eq_refl[OF sym[OF conversep_conversep]]}] 1; +fun mk_rel_conversep_tac ctxt le_conversep rel_mono = + EVERY' [rtac ctxt @{thm antisym}, rtac ctxt le_conversep, rtac ctxt @{thm xt1(6)}, rtac ctxt conversep_shift, + rtac ctxt le_conversep, rtac ctxt @{thm iffD2[OF conversep_mono]}, rtac ctxt rel_mono, + REPEAT_DETERM o rtac ctxt @{thm eq_refl[OF sym[OF conversep_conversep]]}] 1; fun mk_rel_OO_le_tac ctxt rel_OO_Grps rel_eq map_cong0 map_comp set_maps = let val n = length set_maps; - fun in_tac nthO_in = rtac CollectI THEN' - CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS ord_eq_le_trans), - rtac @{thm image_subsetI}, rtac nthO_in, etac set_mp, atac]) set_maps; + fun in_tac nthO_in = rtac ctxt CollectI THEN' + CONJ_WRAP' (fn thm => EVERY' [rtac ctxt (thm RS ord_eq_le_trans), + rtac ctxt @{thm image_subsetI}, rtac ctxt nthO_in, etac ctxt set_mp, atac]) set_maps; val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac - else rtac (hd rel_OO_Grps RS ord_eq_le_trans) THEN' - rtac (@{thm arg_cong2[of _ _ _ _ "op OO"]} OF (replicate 2 (hd rel_OO_Grps RS sym)) RSN + else rtac ctxt (hd rel_OO_Grps RS ord_eq_le_trans) THEN' + rtac ctxt (@{thm arg_cong2[of _ _ _ _ "op OO"]} OF (replicate 2 (hd rel_OO_Grps RS sym)) RSN (2, ord_le_eq_trans)); in - if null set_maps then rtac (rel_eq RS @{thm leq_OOI}) 1 + if null set_maps then rtac ctxt (rel_eq RS @{thm leq_OOI}) 1 else - EVERY' [rel_OO_Grps_tac, rtac @{thm predicate2I}, + EVERY' [rel_OO_Grps_tac, rtac ctxt @{thm predicate2I}, REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}], hyp_subst_tac ctxt, - rtac @{thm relcomppI}, rtac @{thm relcomppI}, rtac @{thm conversepI}, rtac @{thm GrpI}, - rtac trans, rtac map_comp, rtac sym, rtac map_cong0, - REPEAT_DETERM_N n o rtac @{thm fst_fstOp}, + rtac ctxt @{thm relcomppI}, rtac ctxt @{thm relcomppI}, rtac ctxt @{thm conversepI}, rtac ctxt @{thm GrpI}, + rtac ctxt trans, rtac ctxt map_comp, rtac ctxt sym, rtac ctxt map_cong0, + REPEAT_DETERM_N n o rtac ctxt @{thm fst_fstOp}, in_tac @{thm fstOp_in}, - rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac map_cong0, - REPEAT_DETERM_N n o EVERY' [rtac trans, rtac o_apply, - rtac ballE, rtac subst, - rtac @{thm csquare_def}, rtac @{thm csquare_fstOp_sndOp}, atac, etac notE, - etac set_mp, atac], + rtac ctxt @{thm GrpI}, rtac ctxt trans, rtac ctxt map_comp, rtac ctxt map_cong0, + REPEAT_DETERM_N n o EVERY' [rtac ctxt trans, rtac ctxt o_apply, + rtac ctxt ballE, rtac ctxt subst, + rtac ctxt @{thm csquare_def}, rtac ctxt @{thm csquare_fstOp_sndOp}, atac, etac ctxt notE, + etac ctxt set_mp, atac], in_tac @{thm fstOp_in}, - rtac @{thm relcomppI}, rtac @{thm conversepI}, rtac @{thm GrpI}, - rtac trans, rtac map_comp, rtac map_cong0, - REPEAT_DETERM_N n o rtac o_apply, + rtac ctxt @{thm relcomppI}, rtac ctxt @{thm conversepI}, rtac ctxt @{thm GrpI}, + rtac ctxt trans, rtac ctxt map_comp, rtac ctxt map_cong0, + REPEAT_DETERM_N n o rtac ctxt o_apply, in_tac @{thm sndOp_in}, - rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac sym, rtac map_cong0, - REPEAT_DETERM_N n o rtac @{thm snd_sndOp}, + rtac ctxt @{thm GrpI}, rtac ctxt trans, rtac ctxt map_comp, rtac ctxt sym, rtac ctxt map_cong0, + REPEAT_DETERM_N n o rtac ctxt @{thm snd_sndOp}, in_tac @{thm sndOp_in}] 1 end; @@ -231,51 +231,51 @@ unfold_tac ctxt [in_rel] THEN REPEAT_DETERM (eresolve_tac ctxt [exE, CollectE, conjE] 1) THEN hyp_subst_tac ctxt 1 THEN - EVERY' [rtac exI, rtac @{thm conjI[OF CollectI conjI[OF refl refl]]}, + EVERY' [rtac ctxt exI, rtac ctxt @{thm conjI[OF CollectI conjI[OF refl refl]]}, CONJ_WRAP' (fn thm => - (etac (@{thm Collect_split_mono_strong} OF [thm, thm]) THEN' atac)) + (etac ctxt (@{thm Collect_split_mono_strong} OF [thm, thm]) THEN' atac)) set_maps] 1; fun mk_rel_transfer_tac ctxt in_rel rel_map rel_mono_strong = let fun last_tac iffD = - HEADGOAL (etac rel_mono_strong) THEN - REPEAT_DETERM (HEADGOAL (etac (@{thm predicate2_transferD} RS iffD) THEN' + HEADGOAL (etac ctxt rel_mono_strong) THEN + REPEAT_DETERM (HEADGOAL (etac ctxt (@{thm predicate2_transferD} RS iffD) THEN' REPEAT_DETERM o atac)); in - REPEAT_DETERM (HEADGOAL (rtac rel_funI)) THEN - (HEADGOAL (hyp_subst_tac ctxt THEN' rtac refl) ORELSE + REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN + (HEADGOAL (hyp_subst_tac ctxt THEN' rtac ctxt refl) ORELSE REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt (Tactic.make_elim (in_rel RS iffD1) :: @{thms exE conjE CollectE}))) THEN - HEADGOAL (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt rel_map THEN HEADGOAL (rtac iffI) THEN + HEADGOAL (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt rel_map THEN HEADGOAL (rtac ctxt iffI) THEN last_tac iffD1 THEN last_tac iffD2) end; fun mk_map_transfer_tac ctxt rel_mono in_rel set_maps map_cong0 map_comp = let val n = length set_maps; - val in_tac = if n = 0 then rtac UNIV_I else - rtac CollectI THEN' CONJ_WRAP' (fn thm => - etac (thm RS + val in_tac = if n = 0 then rtac ctxt UNIV_I else + rtac ctxt CollectI THEN' CONJ_WRAP' (fn thm => + etac ctxt (thm RS @{thm ord_eq_le_trans[OF _ subset_trans[OF image_mono convol_image_vimage2p]]})) set_maps; in - REPEAT_DETERM_N n (HEADGOAL (rtac rel_funI)) THEN + REPEAT_DETERM_N n (HEADGOAL (rtac ctxt rel_funI)) THEN unfold_thms_tac ctxt @{thms rel_fun_iff_leq_vimage2p} THEN - HEADGOAL (EVERY' [rtac @{thm order_trans}, rtac rel_mono, REPEAT_DETERM_N n o atac, - rtac @{thm predicate2I}, dtac (in_rel RS iffD1), + HEADGOAL (EVERY' [rtac ctxt @{thm order_trans}, rtac ctxt rel_mono, REPEAT_DETERM_N n o atac, + rtac ctxt @{thm predicate2I}, dtac ctxt (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, CollectE, conjE], hyp_subst_tac ctxt, - rtac @{thm vimage2pI}, rtac (in_rel RS iffD2), rtac exI, rtac conjI, in_tac, - rtac conjI, + rtac ctxt @{thm vimage2pI}, rtac ctxt (in_rel RS iffD2), rtac ctxt exI, rtac ctxt conjI, in_tac, + rtac ctxt conjI, EVERY' (map (fn convol => - rtac (@{thm box_equals} OF [map_cong0, map_comp RS sym, map_comp RS sym]) THEN' - REPEAT_DETERM_N n o rtac (convol RS fun_cong)) @{thms fst_convol snd_convol})]) + rtac ctxt (@{thm box_equals} OF [map_cong0, map_comp RS sym, map_comp RS sym]) THEN' + REPEAT_DETERM_N n o rtac ctxt (convol RS fun_cong)) @{thms fst_convol snd_convol})]) end; fun mk_in_bd_tac ctxt live surj_imp_ordLeq_inst map_comp map_id map_cong0 set_maps set_bds bd_card_order bd_Card_order bd_Cinfinite bd_Cnotzero = if live = 0 then - rtac @{thm ordLeq_transitive[OF ordLeq_csum2[OF card_of_Card_order] + rtac ctxt @{thm ordLeq_transitive[OF ordLeq_csum2[OF card_of_Card_order] ordLeq_cexp2[OF ordLeq_refl[OF Card_order_ctwo] Card_order_csum]]} 1 else let @@ -286,67 +286,67 @@ [set_bd, bd_Card_order RS @{thm card_of_Field_ordIso} RS @{thm ordIso_symmetric}]]) set_bds; in - EVERY' [rtac (Drule.rotate_prems 1 ctrans), rtac @{thm cprod_cinfinite_bound}, - rtac (ctrans OF @{thms ordLeq_csum2 ordLeq_cexp2}), rtac @{thm card_of_Card_order}, - rtac @{thm ordLeq_csum2}, rtac @{thm Card_order_ctwo}, rtac @{thm Card_order_csum}, - rtac @{thm ordIso_ordLeq_trans}, rtac @{thm cexp_cong1}, - if live = 1 then rtac @{thm ordIso_refl[OF Card_order_csum]} + EVERY' [rtac ctxt (Drule.rotate_prems 1 ctrans), rtac ctxt @{thm cprod_cinfinite_bound}, + rtac ctxt (ctrans OF @{thms ordLeq_csum2 ordLeq_cexp2}), rtac ctxt @{thm card_of_Card_order}, + rtac ctxt @{thm ordLeq_csum2}, rtac ctxt @{thm Card_order_ctwo}, rtac ctxt @{thm Card_order_csum}, + rtac ctxt @{thm ordIso_ordLeq_trans}, rtac ctxt @{thm cexp_cong1}, + if live = 1 then rtac ctxt @{thm ordIso_refl[OF Card_order_csum]} else - REPEAT_DETERM_N (live - 2) o rtac @{thm ordIso_transitive[OF csum_cong2]} THEN' - REPEAT_DETERM_N (live - 1) o rtac @{thm csum_csum}, - rtac bd_Card_order, rtac (@{thm cexp_mono2_Cnotzero} RS ctrans), rtac @{thm ordLeq_csum1}, - rtac bd_Card_order, rtac @{thm Card_order_csum}, rtac bd_Cnotzero, - rtac @{thm csum_Cfinite_cexp_Cinfinite}, - rtac (if live = 1 then @{thm card_of_Card_order} else @{thm Card_order_csum}), - CONJ_WRAP_GEN' (rtac @{thm Cfinite_csum}) (K (rtac @{thm Cfinite_cone})) set_maps, - rtac bd'_Cinfinite, rtac @{thm card_of_Card_order}, - rtac @{thm Card_order_cexp}, rtac @{thm Cinfinite_cexp}, rtac @{thm ordLeq_csum2}, - rtac @{thm Card_order_ctwo}, rtac bd'_Cinfinite, - rtac (Drule.rotate_prems 1 (@{thm cprod_mono2} RSN (2, ctrans))), + REPEAT_DETERM_N (live - 2) o rtac ctxt @{thm ordIso_transitive[OF csum_cong2]} THEN' + REPEAT_DETERM_N (live - 1) o rtac ctxt @{thm csum_csum}, + rtac ctxt bd_Card_order, rtac ctxt (@{thm cexp_mono2_Cnotzero} RS ctrans), rtac ctxt @{thm ordLeq_csum1}, + rtac ctxt bd_Card_order, rtac ctxt @{thm Card_order_csum}, rtac ctxt bd_Cnotzero, + rtac ctxt @{thm csum_Cfinite_cexp_Cinfinite}, + rtac ctxt (if live = 1 then @{thm card_of_Card_order} else @{thm Card_order_csum}), + CONJ_WRAP_GEN' (rtac ctxt @{thm Cfinite_csum}) (K (rtac ctxt @{thm Cfinite_cone})) set_maps, + rtac ctxt bd'_Cinfinite, rtac ctxt @{thm card_of_Card_order}, + rtac ctxt @{thm Card_order_cexp}, rtac ctxt @{thm Cinfinite_cexp}, rtac ctxt @{thm ordLeq_csum2}, + rtac ctxt @{thm Card_order_ctwo}, rtac ctxt bd'_Cinfinite, + rtac ctxt (Drule.rotate_prems 1 (@{thm cprod_mono2} RSN (2, ctrans))), REPEAT_DETERM_N (live - 1) o - (rtac (bd_Cinfinite RS @{thm cprod_cexp_csum_cexp_Cinfinite} RSN (2, ctrans)) THEN' - rtac @{thm ordLeq_ordIso_trans[OF cprod_mono2 ordIso_symmetric[OF cprod_cexp]]}), - rtac @{thm ordLeq_refl[OF Card_order_cexp]}] 1 THEN + (rtac ctxt (bd_Cinfinite RS @{thm cprod_cexp_csum_cexp_Cinfinite} RSN (2, ctrans)) THEN' + rtac ctxt @{thm ordLeq_ordIso_trans[OF cprod_mono2 ordIso_symmetric[OF cprod_cexp]]}), + rtac ctxt @{thm ordLeq_refl[OF Card_order_cexp]}] 1 THEN unfold_thms_tac ctxt [bd_card_order RS @{thm card_order_csum_cone_cexp_def}] THEN unfold_thms_tac ctxt @{thms cprod_def Field_card_of} THEN - EVERY' [rtac (Drule.rotate_prems 1 ctrans), rtac surj_imp_ordLeq_inst, rtac subsetI, - Method.insert_tac inserts, REPEAT_DETERM o dtac meta_spec, - REPEAT_DETERM o eresolve_tac ctxt [exE, Tactic.make_elim conjunct1], etac CollectE, + EVERY' [rtac ctxt (Drule.rotate_prems 1 ctrans), rtac ctxt surj_imp_ordLeq_inst, rtac ctxt subsetI, + Method.insert_tac inserts, REPEAT_DETERM o dtac ctxt meta_spec, + REPEAT_DETERM o eresolve_tac ctxt [exE, Tactic.make_elim conjunct1], etac ctxt CollectE, if live = 1 then K all_tac - else REPEAT_DETERM_N (live - 2) o (etac conjE THEN' rotate_tac ~1) THEN' etac conjE, - rtac (Drule.rotate_prems 1 @{thm image_eqI}), rtac @{thm SigmaI}, rtac @{thm UNIV_I}, - CONJ_WRAP_GEN' (rtac @{thm SigmaI}) - (K (etac @{thm If_the_inv_into_in_Func} THEN' atac)) set_maps, - rtac sym, - rtac (Drule.rotate_prems 1 + else REPEAT_DETERM_N (live - 2) o (etac ctxt conjE THEN' rotate_tac ~1) THEN' etac ctxt conjE, + rtac ctxt (Drule.rotate_prems 1 @{thm image_eqI}), rtac ctxt @{thm SigmaI}, rtac ctxt @{thm UNIV_I}, + CONJ_WRAP_GEN' (rtac ctxt @{thm SigmaI}) + (K (etac ctxt @{thm If_the_inv_into_in_Func} THEN' atac)) set_maps, + rtac ctxt sym, + rtac ctxt (Drule.rotate_prems 1 ((@{thm box_equals} OF [map_cong0 OF replicate live @{thm If_the_inv_into_f_f}, map_comp RS sym, map_id]) RSN (2, trans))), REPEAT_DETERM_N (2 * live) o atac, - REPEAT_DETERM_N live o rtac (@{thm prod.case} RS trans), - rtac refl, - rtac @{thm surj_imp_ordLeq}, rtac subsetI, rtac (Drule.rotate_prems 1 @{thm image_eqI}), - REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], rtac CollectI, + REPEAT_DETERM_N live o rtac ctxt (@{thm prod.case} RS trans), + rtac ctxt refl, + rtac ctxt @{thm surj_imp_ordLeq}, rtac ctxt subsetI, rtac ctxt (Drule.rotate_prems 1 @{thm image_eqI}), + REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], rtac ctxt CollectI, CONJ_WRAP' (fn thm => - rtac (thm RS ord_eq_le_trans) THEN' etac @{thm subset_trans[OF image_mono Un_upper1]}) + rtac ctxt (thm RS ord_eq_le_trans) THEN' etac ctxt @{thm subset_trans[OF image_mono Un_upper1]}) set_maps, - rtac sym, - rtac (@{thm box_equals} OF [map_cong0 OF replicate live @{thm fun_cong[OF case_sum_o_inj(1)]}, + rtac ctxt sym, + rtac ctxt (@{thm box_equals} OF [map_cong0 OF replicate live @{thm fun_cong[OF case_sum_o_inj(1)]}, map_comp RS sym, map_id])] 1 end; fun mk_trivial_wit_tac ctxt wit_defs set_maps = unfold_thms_tac ctxt wit_defs THEN HEADGOAL (EVERY' (map (fn thm => - dtac (thm RS equalityD1 RS set_mp) THEN' etac imageE THEN' atac) set_maps)) THEN ALLGOALS atac; + dtac ctxt (thm RS equalityD1 RS set_mp) THEN' etac ctxt imageE THEN' atac) set_maps)) THEN ALLGOALS atac; fun mk_set_transfer_tac ctxt in_rel set_maps = Goal.conjunction_tac 1 THEN - EVERY (map (fn set_map => HEADGOAL (rtac rel_funI) THEN + EVERY (map (fn set_map => HEADGOAL (rtac ctxt rel_funI) THEN REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt (Tactic.make_elim (in_rel RS iffD1) :: @{thms exE conjE CollectE}))) THEN - HEADGOAL (hyp_subst_tac ctxt THEN' rtac (@{thm iffD2[OF arg_cong2]} OF [set_map, set_map]) THEN' - rtac @{thm rel_setI}) THEN - REPEAT (HEADGOAL (etac imageE THEN' dtac @{thm set_mp} THEN' atac THEN' + HEADGOAL (hyp_subst_tac ctxt THEN' rtac ctxt (@{thm iffD2[OF arg_cong2]} OF [set_map, set_map]) THEN' + rtac ctxt @{thm rel_setI}) THEN + REPEAT (HEADGOAL (etac ctxt imageE THEN' dtac ctxt @{thm set_mp} THEN' atac THEN' REPEAT_DETERM o (eresolve_tac ctxt @{thms CollectE case_prodE}) THEN' hyp_subst_tac ctxt THEN' - rtac bexI THEN' etac @{thm subst_Pair[OF _ refl]} THEN' etac imageI))) set_maps); + rtac ctxt bexI THEN' etac ctxt @{thm subst_Pair[OF _ refl]} THEN' etac ctxt imageI))) set_maps); end; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_fp_def_sugar.ML --- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Thu Jul 16 12:23:22 2015 +0200 @@ -896,8 +896,8 @@ [] else Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) - (K (mk_disc_transfer_tac (the_single rel_sel_thms) (the_single exhaust_discs) - (flat (flat distinct_discsss)))) + (fn {context = ctxt, prems = _} => mk_disc_transfer_tac ctxt (the_single rel_sel_thms) + (the_single exhaust_discs) (flat (flat distinct_discsss))) |> Conjunction.elim_balanced (length goals) |> Proof_Context.export names_lthy lthy |> map Thm.close_derivation diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML --- a/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Thu Jul 16 12:23:22 2015 +0200 @@ -27,7 +27,7 @@ val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm -> tactic val mk_ctr_transfer_tac: Proof.context -> thm list -> thm list -> tactic - val mk_disc_transfer_tac: thm -> thm -> thm list -> tactic + val mk_disc_transfer_tac: Proof.context -> thm -> thm -> thm list -> tactic val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic val mk_half_distinct_tac: Proof.context -> thm -> thm -> thm list -> tactic val mk_induct_tac: Proof.context -> int -> int list -> int list list -> int list list list -> @@ -107,55 +107,55 @@ fun mk_case_transfer_tac ctxt rel_cases cases = let val n = length (tl (Thm.prems_of rel_cases)) in - REPEAT_DETERM (HEADGOAL (rtac rel_funI)) THEN - HEADGOAL (etac rel_cases) THEN + REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN + HEADGOAL (etac ctxt rel_cases) THEN ALLGOALS (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt cases THEN - ALLGOALS (fn k => (select_prem_tac n (dtac asm_rl) k) k) THEN - ALLGOALS (REPEAT_DETERM o (rotate_tac ~1 THEN' dtac rel_funD THEN' - (atac THEN' etac thin_rl ORELSE' rtac refl)) THEN' atac) + ALLGOALS (fn k => (select_prem_tac ctxt n (dtac ctxt asm_rl) k) k) THEN + ALLGOALS (REPEAT_DETERM o (rotate_tac ~1 THEN' dtac ctxt rel_funD THEN' + (atac THEN' etac ctxt thin_rl ORELSE' rtac ctxt refl)) THEN' atac) end; fun mk_ctr_transfer_tac ctxt rel_intros rel_eqs = HEADGOAL Goal.conjunction_tac THEN ALLGOALS (REPEAT o (resolve_tac ctxt (rel_funI :: rel_intros) THEN' TRY o (REPEAT_DETERM1 o (atac ORELSE' - K (unfold_thms_tac ctxt rel_eqs) THEN' hyp_subst_tac ctxt THEN' rtac refl)))); + K (unfold_thms_tac ctxt rel_eqs) THEN' hyp_subst_tac ctxt THEN' rtac ctxt refl)))); -fun mk_disc_transfer_tac rel_sel exhaust_disc distinct_disc = +fun mk_disc_transfer_tac ctxt rel_sel exhaust_disc distinct_disc = let fun last_disc_tac iffD = - HEADGOAL (rtac (rotate_prems ~1 exhaust_disc) THEN' atac THEN' - REPEAT_DETERM o (rotate_tac ~1 THEN' dtac (rotate_prems 1 iffD) THEN' atac THEN' - rotate_tac ~1 THEN' etac (rotate_prems 1 notE) THEN' eresolve0_tac distinct_disc)); + HEADGOAL (rtac ctxt (rotate_prems ~1 exhaust_disc) THEN' atac THEN' + REPEAT_DETERM o (rotate_tac ~1 THEN' dtac ctxt (rotate_prems 1 iffD) THEN' atac THEN' + rotate_tac ~1 THEN' etac ctxt (rotate_prems 1 notE) THEN' eresolve0_tac distinct_disc)); in HEADGOAL Goal.conjunction_tac THEN - REPEAT_DETERM (HEADGOAL (rtac rel_funI THEN' dtac (rel_sel RS iffD1) THEN' - REPEAT_DETERM o (etac conjE) THEN' (atac ORELSE' rtac iffI))) THEN + REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI THEN' dtac ctxt (rel_sel RS iffD1) THEN' + REPEAT_DETERM o (etac ctxt conjE) THEN' (atac ORELSE' rtac ctxt iffI))) THEN TRY (last_disc_tac iffD2) THEN TRY (last_disc_tac iffD1) end; fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' = - unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN HEADGOAL (rtac sumEN') THEN - HEADGOAL (EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k, - REPEAT_DETERM o dtac meta_spec, etac meta_mp, atac]) (1 upto n))); + unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN HEADGOAL (rtac ctxt sumEN') THEN + HEADGOAL (EVERY' (maps (fn k => [select_prem_tac ctxt n (rotate_tac 1) k, + REPEAT_DETERM o dtac ctxt meta_spec, etac ctxt meta_mp, atac]) (1 upto n))); fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor = - HEADGOAL (rtac iffI THEN' + HEADGOAL (rtac ctxt iffI THEN' EVERY' (@{map 3} (fn cTs => fn cx => fn th => - dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN' + dtac ctxt (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN' SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN' atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor])); fun mk_half_distinct_tac ctxt ctor_inject abs_inject ctr_defs = unfold_thms_tac ctxt (ctor_inject :: abs_inject :: @{thms sum.inject} @ ctr_defs) THEN - HEADGOAL (rtac @{thm sum.distinct(1)}); + HEADGOAL (rtac ctxt @{thm sum.distinct(1)}); fun mk_inject_tac ctxt ctr_def ctor_inject abs_inject = unfold_thms_tac ctxt [ctr_def] THEN - HEADGOAL (rtac (ctor_inject RS ssubst)) THEN + HEADGOAL (rtac ctxt (ctor_inject RS ssubst)) THEN unfold_thms_tac ctxt (abs_inject :: @{thms sum.inject prod.inject conj_assoc}) THEN - HEADGOAL (rtac refl); + HEADGOAL (rtac ctxt refl); val rec_unfold_thms = @{thms comp_def convol_def fst_conv id_def case_prod_Pair_iden snd_conv split_conv @@ -168,7 +168,7 @@ if_distrib[THEN sym]}; in HEADGOAL (subst_tac ctxt (SOME [1, 2]) [co_rec_def] THEN' - rtac (xtor_co_rec_o_map RS trans) THEN' + rtac ctxt (xtor_co_rec_o_map RS trans) THEN' CONVERSION Thm.eta_long_conversion THEN' asm_simp_tac (ss_only (pre_map_defs @ distinct Thm.eq_thm_prop (map_ident0s @ abs_inverses) @ rec_o_map_simps) ctxt)) @@ -178,7 +178,7 @@ HEADGOAL ((if is_def_looping ctr_def then subst_tac ctxt NONE else SELECT_GOAL o unfold_thms_tac ctxt) [ctr_def]) THEN unfold_thms_tac ctxt (ctor_rec :: fp_abs_inverse :: abs_inverse :: rec_defs @ - pre_map_defs @ map_ident0s @ rec_unfold_thms) THEN HEADGOAL (rtac refl); + pre_map_defs @ map_ident0s @ rec_unfold_thms) THEN HEADGOAL (rtac ctxt refl); fun mk_rec_transfer_tac ctxt nn ns actives passives xssss rec_defs ctor_rec_transfers rel_pre_T_defs rel_eqs = @@ -190,19 +190,19 @@ in HEADGOAL Goal.conjunction_tac THEN EVERY (map (fn ctor_rec_transfer => - REPEAT_DETERM (HEADGOAL (rtac rel_funI)) THEN + REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN unfold_thms_tac ctxt rec_defs THEN - HEADGOAL (etac (mk_rel_funDN_rotated (nn + 1) ctor_rec_transfer)) THEN + HEADGOAL (etac ctxt (mk_rel_funDN_rotated (nn + 1) ctor_rec_transfer)) THEN unfold_thms_tac ctxt rel_pre_T_defs THEN EVERY (fst (@{fold_map 2} (fn k => fn xsss => fn acc => rpair (k + acc) - (HEADGOAL (rtac (mk_rel_funDN_rotated 2 @{thm comp_transfer})) THEN - HEADGOAL (rtac @{thm vimage2p_rel_fun}) THEN + (HEADGOAL (rtac ctxt (mk_rel_funDN_rotated 2 @{thm comp_transfer})) THEN + HEADGOAL (rtac ctxt @{thm vimage2p_rel_fun}) THEN unfold_thms_tac ctxt rel_eqs THEN EVERY (@{map 2} (fn n => fn xss => REPEAT_DETERM (HEADGOAL (resolve_tac ctxt [mk_rel_funDN 2 case_sum_transfer_eq, mk_rel_funDN 2 case_sum_transfer])) THEN - HEADGOAL (select_prem_tac total_n (dtac asm_rl) (acc + n)) THEN + HEADGOAL (select_prem_tac ctxt total_n (dtac ctxt asm_rl) (acc + n)) THEN HEADGOAL (SELECT_GOAL (HEADGOAL (REPEAT_DETERM o (atac ORELSE' resolve_tac ctxt [mk_rel_funDN 1 case_prod_transfer_eq, @@ -212,7 +212,7 @@ let val thm = prems |> permute_like (op =) (True :: flat xss) (True :: flat_rec_arg_args xss) |> Library.foldl1 (fn (acc, elem) => elem RS (acc RS rel_funD)) - in HEADGOAL (rtac thm) end) ctxt))))) + in HEADGOAL (rtac ctxt thm) end) ctxt))))) (1 upto k) xsss))) ns xssss 0))) ctor_rec_transfers') @@ -226,16 +226,16 @@ @{thms o_apply vimage2p_def if_True if_False}) ctxt; in unfold_thms_tac ctxt (ctr_def :: corec_defs) THEN - HEADGOAL (rtac (ctor_dtor_corec RS trans) THEN' asm_simp_tac ss) THEN_MAYBE - HEADGOAL (rtac refl ORELSE' rtac (@{thm unit_eq} RS arg_cong)) + HEADGOAL (rtac ctxt (ctor_dtor_corec RS trans) THEN' asm_simp_tac ss) THEN_MAYBE + HEADGOAL (rtac ctxt refl ORELSE' rtac ctxt (@{thm unit_eq} RS arg_cong)) end; fun mk_corec_disc_iff_tac case_splits' corecs discs ctxt = EVERY (@{map 3} (fn case_split_tac => fn corec_thm => fn disc => HEADGOAL case_split_tac THEN unfold_thms_tac ctxt [corec_thm] THEN HEADGOAL (asm_simp_tac (ss_only basic_simp_thms ctxt)) THEN - (if is_refl disc then all_tac else HEADGOAL (rtac disc))) - (map rtac case_splits' @ [K all_tac]) corecs discs); + (if is_refl disc then all_tac else HEADGOAL (rtac ctxt disc))) + (map (rtac ctxt) case_splits' @ [K all_tac]) corecs discs); fun mk_corec_transfer_tac ctxt actives passives type_definitions corec_defs dtor_corec_transfers rel_pre_T_defs rel_eqs pgs pss qssss gssss = @@ -253,24 +253,24 @@ fun mk_unfold_If_tac total pos = HEADGOAL (Inl_Inr_Pair_tac THEN' - rtac (mk_rel_funDN 3 @{thm If_transfer}) THEN' - select_prem_tac total (dtac asm_rl) pos THEN' - dtac rel_funD THEN' atac THEN' atac); + rtac ctxt (mk_rel_funDN 3 @{thm If_transfer}) THEN' + select_prem_tac ctxt total (dtac ctxt asm_rl) pos THEN' + dtac ctxt rel_funD THEN' atac THEN' atac); fun mk_unfold_Inl_Inr_Pair_tac total pos = HEADGOAL (Inl_Inr_Pair_tac THEN' - select_prem_tac total (dtac asm_rl) pos THEN' - dtac rel_funD THEN' atac THEN' atac); + select_prem_tac ctxt total (dtac ctxt asm_rl) pos THEN' + dtac ctxt rel_funD THEN' atac THEN' atac); fun mk_unfold_arg_tac qs gs = EVERY (map (mk_unfold_If_tac num_pgs o prem_no_of) qs) THEN EVERY (map (mk_unfold_Inl_Inr_Pair_tac num_pgs o prem_no_of) gs); fun mk_unfold_ctr_tac type_definition qss gss = - HEADGOAL (rtac (mk_rel_funDN 1 (@{thm Abs_transfer} OF + HEADGOAL (rtac ctxt (mk_rel_funDN 1 (@{thm Abs_transfer} OF [type_definition, type_definition])) THEN' Inl_Inr_Pair_tac) THEN (case (qss, gss) of - ([], []) => HEADGOAL (rtac refl) + ([], []) => HEADGOAL (rtac ctxt refl) | _ => EVERY (map2 mk_unfold_arg_tac qss gss)); fun mk_unfold_type_tac type_definition ps qsss gsss = @@ -281,7 +281,7 @@ | mk_unfold_ty (p_tac :: p_tacs) (qg_tac :: qg_tacs) = p_tac THEN qg_tac THEN mk_unfold_ty p_tacs qg_tacs in - HEADGOAL (rtac rel_funI) THEN mk_unfold_ty p_tacs qg_tacs + HEADGOAL (rtac ctxt rel_funI) THEN mk_unfold_ty p_tacs qg_tacs end; fun mk_unfold_corec_type_tac dtor_corec_transfer corec_def = @@ -290,9 +290,9 @@ val dtor_corec_transfer' = cterm_instantiate_pos (SOME active :: map SOME passives @ map SOME actives') dtor_corec_transfer; in - HEADGOAL Goal.conjunction_tac THEN REPEAT_DETERM (HEADGOAL (rtac rel_funI)) THEN + HEADGOAL Goal.conjunction_tac THEN REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN unfold_thms_tac ctxt [corec_def] THEN - HEADGOAL (etac (mk_rel_funDN_rotated (1 + length actives) dtor_corec_transfer')) THEN + HEADGOAL (etac ctxt (mk_rel_funDN_rotated (1 + length actives) dtor_corec_transfer')) THEN unfold_thms_tac ctxt (rel_pre_T_defs @ rel_eqs) end; @@ -305,13 +305,13 @@ end; fun solve_prem_prem_tac ctxt = - REPEAT o (eresolve_tac ctxt @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE' + REPEAT o (eresolve_tac ctxt @{thms bexE rev_bexI} ORELSE' rtac ctxt @{thm rev_bexI[OF UNIV_I]} ORELSE' hyp_subst_tac ctxt ORELSE' resolve_tac ctxt @{thms disjI1 disjI2}) THEN' - (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI}); + (rtac ctxt refl ORELSE' atac ORELSE' rtac ctxt @{thm singletonI}); fun mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps pre_set_defs = - HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp, + HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac ctxt nn (dtac ctxt meta_spec) kk, etac ctxt meta_mp, SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ fp_abs_inverses @ abs_inverses @ set_maps @ sumprod_thms_set)), solve_prem_prem_tac ctxt]) (rev kks))); @@ -319,10 +319,10 @@ fun mk_induct_discharge_prem_tac ctxt nn n fp_abs_inverses abs_inverses set_maps pre_set_defs m k kks = let val r = length kks in - HEADGOAL (EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt, - REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)]) THEN + HEADGOAL (EVERY' [select_prem_tac ctxt n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt, + REPEAT_DETERM_N m o (dtac ctxt meta_spec THEN' rotate_tac ~1)]) THEN EVERY [REPEAT_DETERM_N r - (HEADGOAL (rotate_tac ~1 THEN' dtac meta_mp THEN' rotate_tac 1) THEN prefer_tac 2), + (HEADGOAL (rotate_tac ~1 THEN' dtac ctxt meta_mp THEN' rotate_tac 1) THEN prefer_tac 2), if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL atac, mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps pre_set_defs] @@ -335,7 +335,7 @@ EVERY (map (fn def => HEADGOAL (subst_asm_tac ctxt NONE [def])) ctr_defs) else unfold_thms_tac ctxt ctr_defs) THEN - HEADGOAL (rtac ctor_induct') THEN co_induct_inst_as_projs_tac ctxt 0 THEN + HEADGOAL (rtac ctxt ctor_induct') THEN co_induct_inst_as_projs_tac ctxt 0 THEN EVERY (@{map 4} (EVERY oooo @{map 3} o mk_induct_discharge_prem_tac ctxt nn n fp_abs_inverses abs_inverses set_maps) pre_set_defss mss (unflat mss (1 upto n)) kkss) @@ -349,9 +349,9 @@ SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels)) THEN' SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: fp_abs_inverse :: abs_inverse :: dtor_ctor :: sels @ sumprod_thms_rel @ @{thms o_apply vimage2p_def})) THEN' - (atac ORELSE' REPEAT o etac conjE THEN' + (atac ORELSE' REPEAT o etac ctxt conjE THEN' full_simp_tac (ss_only (no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN' - REPEAT o etac conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN' + REPEAT o etac ctxt conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN' REPEAT o (resolve_tac ctxt [refl, conjI] ORELSE' atac)); fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' = @@ -359,18 +359,18 @@ val discs'' = map (perhaps (try (fn th => th RS @{thm notnotD}))) (discs @ discs') |> distinct Thm.eq_thm_prop; in - hyp_subst_tac ctxt THEN' REPEAT o etac conjE THEN' + hyp_subst_tac ctxt THEN' REPEAT o etac ctxt conjE THEN' full_simp_tac (ss_only (refl :: no_refl discs'' @ basic_simp_thms) ctxt) end; fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def fp_abs_inverse abs_inverse dtor_ctor exhaust ctr_defs discss selss = let val ks = 1 upto n in - EVERY' ([rtac allI, rtac allI, rtac impI, select_prem_tac nn (dtac meta_spec) kk, - dtac meta_spec, dtac meta_mp, atac, rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 0), + EVERY' ([rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, select_prem_tac ctxt nn (dtac ctxt meta_spec) kk, + dtac ctxt meta_spec, dtac ctxt meta_mp, atac, rtac ctxt exhaust, K (co_induct_inst_as_projs_tac ctxt 0), hyp_subst_tac ctxt] @ @{map 4} (fn k => fn ctr_def => fn discs => fn sels => - EVERY' ([rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 1)] @ + EVERY' ([rtac ctxt exhaust, K (co_induct_inst_as_projs_tac ctxt 1)] @ map2 (fn k' => fn discs' => if k' = k then mk_coinduct_same_ctr_tac ctxt rel_eqs' pre_rel_def fp_abs_inverse abs_inverse @@ -381,82 +381,82 @@ fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs fp_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss discsss selsss = - HEADGOAL (rtac dtor_coinduct' THEN' + HEADGOAL (rtac ctxt dtor_coinduct' THEN' EVERY' (@{map 10} (mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn) (1 upto nn) ns pre_rel_defs fp_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss discsss selsss)); fun mk_map_disc_iff_tac ctxt ct exhaust discs maps = TRYALL Goal.conjunction_tac THEN - ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW + ALLGOALS (rtac ctxt (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW REPEAT_DETERM o hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt maps THEN unfold_thms_tac ctxt (map (fn thm => thm RS eqFalseI handle THM _ => thm RS eqTrueI) discs) THEN - ALLGOALS (rtac refl ORELSE' rtac TrueI); + ALLGOALS (rtac ctxt refl ORELSE' rtac ctxt TrueI); fun mk_map_sel_tac ctxt ct exhaust discs maps sels map_id0s = TRYALL Goal.conjunction_tac THEN - ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW + ALLGOALS (rtac ctxt (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW REPEAT_DETERM o hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @ @{thms not_True_eq_False not_False_eq_True}) THEN - TRYALL (etac FalseE ORELSE' etac @{thm TrueE}) THEN + TRYALL (etac ctxt FalseE ORELSE' etac ctxt @{thm TrueE}) THEN unfold_thms_tac ctxt (@{thm id_apply} :: maps @ sels @ map_id0s) THEN - ALLGOALS (rtac refl); + ALLGOALS (rtac ctxt refl); fun mk_rel_cases_tac ctxt ct1 ct2 exhaust injects rel_injects distincts rel_distincts rel_eqs= - HEADGOAL (rtac (cterm_instantiate_pos [SOME ct1] exhaust) THEN_ALL_NEW - rtac (cterm_instantiate_pos [SOME ct2] exhaust) THEN_ALL_NEW + HEADGOAL (rtac ctxt (cterm_instantiate_pos [SOME ct1] exhaust) THEN_ALL_NEW + rtac ctxt (cterm_instantiate_pos [SOME ct2] exhaust) THEN_ALL_NEW hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt (rel_eqs @ injects @ rel_injects @ @{thms conj_imp_eq_imp_imp simp_thms(6) True_implies_equals conj_imp_eq_imp_imp} @ map (fn thm => thm RS eqFalseI) (distincts @ rel_distincts) @ map (fn thm => thm RS eqTrueI) rel_injects) THEN - TRYALL (atac ORELSE' etac FalseE ORELSE' - (REPEAT_DETERM o dtac @{thm meta_spec} THEN' + TRYALL (atac ORELSE' etac ctxt FalseE ORELSE' + (REPEAT_DETERM o dtac ctxt @{thm meta_spec} THEN' TRY o filter_prems_tac ctxt (forall (curry (op <>) (HOLogic.mk_Trueprop @{term False})) o Logic.strip_imp_prems) THEN' - REPEAT_DETERM o (dtac @{thm meta_mp} THEN' rtac refl) THEN' Goal.assume_rule_tac ctxt)); + REPEAT_DETERM o (dtac ctxt @{thm meta_mp} THEN' rtac ctxt refl) THEN' Goal.assume_rule_tac ctxt)); fun mk_rel_coinduct0_tac ctxt dtor_rel_coinduct cts assms exhausts discss selss ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses nesting_rel_eqs = - rtac dtor_rel_coinduct 1 THEN + rtac ctxt dtor_rel_coinduct 1 THEN EVERY (@{map 11} (fn ct => fn assm => fn exhaust => fn discs => fn sels => fn ctor_defs => fn dtor_ctor => fn ctor_inject => fn abs_inject => fn rel_pre_def => fn abs_inverse => - (rtac exhaust THEN_ALL_NEW (rtac exhaust THEN_ALL_NEW - (dtac (rotate_prems ~1 (cterm_instantiate_pos [NONE, NONE, NONE, NONE, SOME ct] + (rtac ctxt exhaust THEN_ALL_NEW (rtac ctxt exhaust THEN_ALL_NEW + (dtac ctxt (rotate_prems ~1 (cterm_instantiate_pos [NONE, NONE, NONE, NONE, SOME ct] @{thm arg_cong2} RS iffD1)) THEN' - atac THEN' atac THEN' hyp_subst_tac ctxt THEN' dtac assm THEN' - REPEAT_DETERM o etac conjE))) 1 THEN + atac THEN' atac THEN' hyp_subst_tac ctxt THEN' dtac ctxt assm THEN' + REPEAT_DETERM o etac ctxt conjE))) 1 THEN unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @ sels @ simp_thms') THEN unfold_thms_tac ctxt (dtor_ctor :: rel_pre_def :: abs_inverse :: ctor_inject :: abs_inject :: ctor_defs @ nesting_rel_eqs @ simp_thms' @ @{thms id_bnf_def rel_sum_simps rel_prod_apply vimage2p_def Inl_Inr_False iffD2[OF eq_False Inr_not_Inl] sum.inject prod.inject}) THEN - REPEAT_DETERM (HEADGOAL ((REPEAT_DETERM o etac conjE) THEN' (REPEAT_DETERM o rtac conjI) THEN' - (rtac refl ORELSE' atac)))) + REPEAT_DETERM (HEADGOAL ((REPEAT_DETERM o etac ctxt conjE) THEN' (REPEAT_DETERM o rtac ctxt conjI) THEN' + (rtac ctxt refl ORELSE' atac)))) cts assms exhausts discss selss ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses); fun mk_rel_induct0_tac ctxt ctor_rel_induct assms cterms exhausts ctor_defss ctor_injects rel_pre_list_defs Abs_inverses nesting_rel_eqs = - rtac ctor_rel_induct 1 THEN EVERY (@{map 6} (fn cterm => fn exhaust => fn ctor_defs => + rtac ctxt ctor_rel_induct 1 THEN EVERY (@{map 6} (fn cterm => fn exhaust => fn ctor_defs => fn ctor_inject => fn rel_pre_list_def => fn Abs_inverse => - HEADGOAL (rtac exhaust THEN_ALL_NEW (rtac exhaust THEN_ALL_NEW - (rtac (cterm_instantiate_pos (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2} RS iffD2) + HEADGOAL (rtac ctxt exhaust THEN_ALL_NEW (rtac ctxt exhaust THEN_ALL_NEW + (rtac ctxt (cterm_instantiate_pos (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2} RS iffD2) THEN' atac THEN' atac THEN' TRY o resolve_tac ctxt assms))) THEN unfold_thms_tac ctxt (ctor_inject :: rel_pre_list_def :: ctor_defs @ nesting_rel_eqs @ @{thms id_bnf_def vimage2p_def}) THEN TRYALL (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt (Abs_inverse :: @{thms rel_sum_simps rel_prod_apply Inl_Inr_False Inr_Inl_False sum.inject prod.inject}) THEN - TRYALL (rtac refl ORELSE' etac FalseE ORELSE' (REPEAT_DETERM o etac conjE) THEN' atac)) + TRYALL (rtac ctxt refl ORELSE' etac ctxt FalseE ORELSE' (REPEAT_DETERM o etac ctxt conjE) THEN' atac)) cterms exhausts ctor_defss ctor_injects rel_pre_list_defs Abs_inverses); fun mk_rel_sel_tac ctxt ct1 ct2 exhaust discs sels rel_injects distincts rel_distincts rel_eqs = - HEADGOAL (rtac (cterm_instantiate_pos [SOME ct1] exhaust) THEN_ALL_NEW - rtac (cterm_instantiate_pos [SOME ct2] exhaust) THEN_ALL_NEW + HEADGOAL (rtac ctxt (cterm_instantiate_pos [SOME ct1] exhaust) THEN_ALL_NEW + rtac ctxt (cterm_instantiate_pos [SOME ct2] exhaust) THEN_ALL_NEW hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt (sels @ rel_injects @ rel_eqs @ @{thms simp_thms(6,7,8,11,12,15,16,21,22,24)} @ ((discs @ distincts) RL [eqTrueI, eqFalseI]) @ @@ -466,24 +466,24 @@ fun mk_sel_transfer_tac ctxt n sel_defs case_transfer = TRYALL Goal.conjunction_tac THEN unfold_thms_tac ctxt (map (Drule.abs_def o Local_Defs.meta_rewrite_rule ctxt) sel_defs) THEN - ALLGOALS (rtac (mk_rel_funDN n case_transfer) THEN_ALL_NEW - REPEAT_DETERM o (atac ORELSE' rtac rel_funI)); + ALLGOALS (rtac ctxt (mk_rel_funDN n case_transfer) THEN_ALL_NEW + REPEAT_DETERM o (atac ORELSE' rtac ctxt rel_funI)); fun mk_set_sel_tac ctxt ct exhaust discs sels sets = TRYALL Goal.conjunction_tac THEN - ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW + ALLGOALS (rtac ctxt (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW REPEAT_DETERM o hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @ @{thms not_True_eq_False not_False_eq_True}) THEN - TRYALL (etac FalseE ORELSE' etac @{thm TrueE}) THEN + TRYALL (etac ctxt FalseE ORELSE' etac ctxt @{thm TrueE}) THEN unfold_thms_tac ctxt (sels @ sets) THEN ALLGOALS (REPEAT o (resolve_tac ctxt @{thms UnI1 UnI2 imageI} ORELSE' eresolve_tac ctxt @{thms UN_I UN_I[rotated] imageE} ORELSE' hyp_subst_tac ctxt) THEN' - (rtac @{thm singletonI} ORELSE' atac)); + (rtac ctxt @{thm singletonI} ORELSE' atac)); fun mk_set_cases_tac ctxt ct assms exhaust sets = - HEADGOAL (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW hyp_subst_tac ctxt) THEN + HEADGOAL (rtac ctxt (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt sets THEN REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt @{thms FalseE emptyE singletonE UnE UN_E insertE} ORELSE' @@ -494,7 +494,7 @@ TRYALL Goal.conjunction_tac THEN unfold_thms_tac ctxt sets THEN TRYALL (REPEAT o (resolve_tac ctxt @{thms UnI1 UnI2} ORELSE' - eresolve_tac ctxt @{thms UN_I UN_I[rotated]}) THEN' (rtac @{thm singletonI} ORELSE' atac)); + eresolve_tac ctxt @{thms UN_I UN_I[rotated]}) THEN' (rtac ctxt @{thm singletonI} ORELSE' atac)); fun mk_set_induct0_tac ctxt cts assms dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses = @@ -502,8 +502,8 @@ val assms_tac = let val assms' = map (unfold_thms ctxt (@{thm id_bnf_def} :: ctor_defs)) assms in fold (curry (op ORELSE')) (map (fn thm => - funpow (length (Thm.prems_of thm)) (fn tac => tac THEN' atac) (rtac thm)) assms') - (etac FalseE) + funpow (length (Thm.prems_of thm)) (fn tac => tac THEN' atac) (rtac ctxt thm)) assms') + (etac ctxt FalseE) end; val exhausts' = map (fn thm => thm RS @{thm asm_rl[of "P x y" for P x y]}) exhausts |> map2 (fn ct => cterm_instantiate_pos [NONE, SOME ct]) cts; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_fp_n2m.ML --- a/src/HOL/Tools/BNF/bnf_fp_n2m.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_fp_n2m.ML Thu Jul 16 12:23:22 2015 +0200 @@ -462,7 +462,7 @@ unfold_thms_tac ctxt (flat [rec_thms, raw_co_rec_defs, pre_map_defs, fp_pre_map_defs, fp_xtor_co_recs, fp_rec_o_maps, map_thms, fp_Rep_o_Abss, Rep_o_Abss]) THEN - CONJ_WRAP (K (HEADGOAL (rtac refl))) bnfs; + CONJ_WRAP (K (HEADGOAL (rtac ctxt refl))) bnfs; in Library.foldr1 HOLogic.mk_conj goals |> HOLogic.mk_Trueprop diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML --- a/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML Thu Jul 16 12:23:22 2015 +0200 @@ -40,10 +40,10 @@ folded_C_IHs rel_monos unfolds; in unfold_thms_tac ctxt vimage2p_unfolds THEN - HEADGOAL (CONJ_WRAP_GEN' (rtac @{thm context_conjI}) - (fn thm => rtac thm THEN_ALL_NEW (rotate_tac ~1 THEN' + HEADGOAL (CONJ_WRAP_GEN' (rtac ctxt @{thm context_conjI}) + (fn thm => rtac ctxt thm THEN_ALL_NEW (rotate_tac ~1 THEN' REPEAT_ALL_NEW (FIRST' [eresolve_tac ctxt C_IHs, eresolve_tac ctxt C_IH_monos, - SELECT_GOAL (unfold_thms_tac ctxt nesting_rel_eqs) THEN' rtac @{thm order_refl}, + SELECT_GOAL (unfold_thms_tac ctxt nesting_rel_eqs) THEN' rtac ctxt @{thm order_refl}, assume_tac ctxt, resolve_tac ctxt co_inducts, resolve_tac ctxt C_IH_monos THEN' REPEAT_ALL_NEW (eresolve_tac ctxt nesting_rel_monos)]))) co_inducts) diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_gfp.ML --- a/src/HOL/Tools/BNF/bnf_gfp.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_gfp.ML Thu Jul 16 12:23:22 2015 +0200 @@ -257,7 +257,7 @@ val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x); in Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, goal)) - (K (mk_map_cong0L_tac m map_cong0 map_id)) + (fn {context = ctxt, prems = _} => mk_map_cong0L_tac ctxt m map_cong0 map_id) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -276,7 +276,8 @@ val goals = map2 (fn prem => fn concl => Logic.mk_implies (prem, concl)) prems concls; in map (fn goal => - Goal.prove_sorry lthy [] [] goal (K ((hyp_subst_tac lthy THEN' rtac refl) 1)) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + (hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) goals end; @@ -335,7 +336,8 @@ Logic.list_implies (coalg_prem :: [prem], concl)) concls) prems conclss; in map (fn goals => map (fn goal => - Goal.prove_sorry lthy [] [] goal (K (mk_coalg_set_tac coalg_def)) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_coalg_set_tac ctxt coalg_def) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) goals) goalss end; @@ -347,9 +349,9 @@ val goal = HOLogic.mk_Trueprop (mk_tcoalg activeAs ss) in Goal.prove_sorry lthy [] [] goal - (K (rtac (coalg_def RS iffD2) 1 THEN CONJ_WRAP - (K (EVERY' [rtac ballI, rtac CollectI, - CONJ_WRAP' (K (EVERY' [rtac @{thm subset_UNIV}])) allAs] 1)) ss)) + (fn {context = ctxt, prems = _} => (rtac ctxt (coalg_def RS iffD2) 1 THEN CONJ_WRAP + (K (EVERY' [rtac ctxt ballI, rtac ctxt CollectI, + CONJ_WRAP' (K (EVERY' [rtac ctxt @{thm subset_UNIV}])) allAs] 1)) ss)) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -409,7 +411,8 @@ mk_Trueprop_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ x]), s' $ (f $ x))); val elim_goals = @{map 6} mk_elim_goal Bs mapsAsBs fs ss s's zs; fun prove goal = - Goal.prove_sorry lthy [] [] goal (K (mk_mor_elim_tac mor_def)) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_mor_elim_tac ctxt mor_def) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); in @@ -424,7 +427,7 @@ val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids); in Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl)) - (K (mk_mor_incl_tac mor_def map_ids)) + (fn {context = ctxt, prems = _} => mk_mor_incl_tac ctxt mor_def map_ids) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -466,7 +469,7 @@ val rhs = Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct mapsAsBs fs ss s's); in Goal.prove_sorry lthy [] [] (mk_Trueprop_eq (lhs, rhs)) - (K (mk_mor_UNIV_tac morE_thms mor_def)) + (fn {context = ctxt, prems = _} => mk_mor_UNIV_tac ctxt morE_thms mor_def) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -478,7 +481,7 @@ in Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor active_UNIVs ss (map HOLogic.mk_UNIV FTsAs) maps ss)) - (K (mk_mor_str_tac ks mor_UNIV_thm)) + (fn {context = ctxt, prems = _} => mk_mor_str_tac ctxt ks mor_UNIV_thm) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -491,7 +494,7 @@ in Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor (map HOLogic.mk_UNIV activeBs) s's sum_UNIVs maps Inls)) - (K (mk_mor_case_sum_tac ks mor_UNIV_thm)) + (fn {context = ctxt, prems = _} => mk_mor_case_sum_tac ctxt ks mor_UNIV_thm) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -567,7 +570,8 @@ (@{map 6} mk_conjunct Rs ss s's zs z's relsAsBs)) in Goal.prove_sorry lthy [] [] (mk_Trueprop_eq (mk_bis Bs ss B's s's Rs, rhs)) - (K (mk_bis_rel_tac m bis_def in_rels map_comps map_cong0s set_mapss)) + (fn {context = ctxt, prems = _} => mk_bis_rel_tac ctxt m bis_def in_rels map_comps + map_cong0s set_mapss) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -576,7 +580,8 @@ Goal.prove_sorry lthy [] [] (Logic.mk_implies (HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs), HOLogic.mk_Trueprop (mk_bis B's s's Bs ss (map mk_converse Rs)))) - (K (mk_bis_converse_tac m bis_rel_thm rel_congs rel_converseps)) + (fn {context = ctxt, prems = _} => mk_bis_converse_tac ctxt m bis_rel_thm rel_congs + rel_converseps) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); @@ -686,7 +691,8 @@ val goals = map2 (fn i => fn R => Logic.mk_implies (sbis_prem, mk_concl i R)) ks sRs; in @{map 3} (fn goal => fn i => fn def => - Goal.prove_sorry lthy [] [] goal (K (mk_incl_lsbis_tac n i def)) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_incl_lsbis_tac ctxt n i def) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) goals ks lsbis_defs end; @@ -698,8 +704,8 @@ in @{map 3} (fn goal => fn l_incl => fn incl_l => Goal.prove_sorry lthy [] [] goal - (K (mk_equiv_lsbis_tac sbis_lsbis_thm l_incl incl_l - bis_Id_on_thm bis_converse_thm bis_O_thm)) + (fn {context = ctxt, prems = _} => mk_equiv_lsbis_tac ctxt sbis_lsbis_thm l_incl incl_l + bis_Id_on_thm bis_converse_thm bis_O_thm) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) goals lsbis_incl_thms incl_lsbis_thms @@ -719,7 +725,8 @@ val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) = typedef (sbdT_bind, sum_bdT_params', NoSyn) - (HOLogic.mk_UNIV sum_bdT) NONE (fn _ => EVERY' [rtac exI, rtac UNIV_I] 1) lthy; + (HOLogic.mk_UNIV sum_bdT) NONE (fn ctxt => + EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val sbdT = Type (sbdT_name, sum_bdT_params); val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT); @@ -1097,7 +1104,8 @@ val goals = map2 mk_goal ks zs; val length_Levs' = map2 (fn goal => fn length_Lev => - Goal.prove_sorry lthy [] [] goal (K (mk_length_Lev'_tac length_Lev)) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_length_Lev'_tac ctxt length_Lev) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) goals length_Levs; in @@ -1120,7 +1128,7 @@ val rv_last = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal) - (K (mk_rv_last_tac cTs cts rv_Nils rv_Conss)) + (fn {context = ctxt, prems = _} => mk_rv_last_tac ctxt cTs cts rv_Nils rv_Conss) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); @@ -1244,20 +1252,22 @@ in @{map 4} (fn goal => fn lsbisE => fn map_comp_id => fn map_cong0 => Goal.prove_sorry lthy [] [] goal - (K (mk_congruent_str_final_tac m lsbisE map_comp_id map_cong0 equiv_LSBIS_thms)) + (fn {context = ctxt, prems = _} => mk_congruent_str_final_tac ctxt m lsbisE map_comp_id + map_cong0 equiv_LSBIS_thms) |> Thm.close_derivation) goals lsbisE_thms map_comp_id_thms map_cong0s end; val coalg_final_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_coalg car_finals str_finals)) - (K (mk_coalg_final_tac m coalg_def congruent_str_final_thms equiv_LSBIS_thms - set_mapss coalgT_set_thmss)) + (fn {context = ctxt, prems = _} => mk_coalg_final_tac ctxt m coalg_def + congruent_str_final_thms equiv_LSBIS_thms set_mapss coalgT_set_thmss) |> Thm.close_derivation; val mor_T_final_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor carTAs strTAs car_finals str_finals (map mk_proj lsbisAs))) - (K (mk_mor_T_final_tac mor_def congruent_str_final_thms equiv_LSBIS_thms)) + (fn {context = ctxt, prems = _} => mk_mor_T_final_tac ctxt mor_def congruent_str_final_thms + equiv_LSBIS_thms) |> Thm.close_derivation; val mor_final_thm = mor_comp_thm OF [mor_beh_thm, mor_T_final_thm]; @@ -1268,8 +1278,9 @@ val ((T_names, (T_glob_infos, T_loc_infos)), lthy) = lthy |> @{fold_map 4} (fn b => fn mx => fn car_final => fn in_car_final => - typedef (b, params, mx) car_final NONE - (fn _ => EVERY' [rtac exI, rtac in_car_final] 1)) bs mixfixes car_finals in_car_final_thms + typedef (b, params, mx) car_final NONE + (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt in_car_final] 1)) + bs mixfixes car_finals in_car_final_thms |>> apsnd split_list o split_list; val Ts = map (fn name => Type (name, params')) T_names; @@ -1397,8 +1408,8 @@ in Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors (map (mk_unfold Ts ss) ks))) - (K (mk_mor_unfold_tac m mor_UNIV_thm dtor_defs unfold_defs Abs_inverses' morEs' - map_comp_id_thms map_cong0s)) + (fn {context = ctxt, prems = _} => mk_mor_unfold_tac ctxt m mor_UNIV_thm dtor_defs + unfold_defs Abs_inverses' morEs' map_comp_id_thms map_cong0s) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -1411,9 +1422,9 @@ (map2 (fn R => fn T => mk_leq R (Id_const T)) TRs Ts)); in `split_conj_thm (Goal.prove_sorry lthy [] [] (Logic.mk_implies (prem, concl)) - (K (mk_raw_coind_tac bis_def bis_cong_thm bis_O_thm bis_converse_thm bis_Gr_thm - tcoalg_thm coalgT_thm mor_T_final_thm sbis_lsbis_thm - lsbis_incl_thms incl_lsbis_thms equiv_LSBIS_thms mor_Rep_thm Rep_injects)) + (fn {context = ctxt, prems = _} => mk_raw_coind_tac ctxt bis_def bis_cong_thm bis_O_thm + bis_converse_thm bis_Gr_thm tcoalg_thm coalgT_thm mor_T_final_thm sbis_lsbis_thm + lsbis_incl_thms incl_lsbis_thms equiv_LSBIS_thms mor_Rep_thm Rep_injects) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) end; @@ -1429,7 +1440,8 @@ val mor_thm = mor_comp_thm OF [mor_final_thm, mor_Abs_thm]; val unique_mor = Goal.prove_sorry lthy [] [] (Logic.mk_implies (prem, unique)) - (K (mk_unfold_unique_mor_tac raw_coind_thms bis_thm mor_thm unfold_defs)) + (fn {context = ctxt, prems = _} => mk_unfold_unique_mor_tac ctxt raw_coind_thms + bis_thm mor_thm unfold_defs) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); in @@ -1629,7 +1641,8 @@ val dtor_coinduct = Goal.prove_sorry lthy [] [] dtor_coinduct_goal - (K (mk_dtor_coinduct_tac m raw_coind_thm bis_rel_thm rel_congs)) + (fn {context = ctxt, prems = _} => mk_dtor_coinduct_tac ctxt m raw_coind_thm bis_rel_thm + rel_congs) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); in @@ -1664,7 +1677,7 @@ mk_xtor_rel_co_induct_thm Greatest_FP rels activeJphis (map HOLogic.eq_const Ts) Jphis Jzs Jz's dtors dtor's (fn {context = ctxt, prems} => (unfold_thms_tac ctxt @{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} THEN - REPEAT_DETERM (rtac allI 1) THEN rtac (dtor_coinduct_thm OF prems) 1)) lthy; + REPEAT_DETERM (rtac ctxt allI 1) THEN rtac ctxt (dtor_coinduct_thm OF prems) 1)) lthy; (*register new codatatypes as BNFs*) val (timer, Jbnfs, (dtor_Jmap_o_thms, dtor_Jmap_thms), dtor_Jmap_unique_thms, dtor_Jset_thmss', @@ -1800,7 +1813,7 @@ @{map 5} (fn goal => fn unfold => fn map_comp => fn map_cong0 => fn map_arg_cong => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN - mk_map_tac m n map_arg_cong unfold map_comp map_cong0) + mk_map_tac ctxt m n map_arg_cong unfold map_comp map_cong0) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) goals dtor_unfold_thms map_comps map_cong0s map_arg_cong_thms; @@ -1833,7 +1846,8 @@ fs_Jmaps gs_Jmaps fgs_Jmaps)) in split_conj_thm (Goal.prove_sorry lthy [] [] goal - (K (mk_map_comp0_tac Jmap_thms map_comp0s dtor_Jmap_unique_thm)) + (fn {context = ctxt, prems = _} => + mk_map_comp0_tac ctxt Jmap_thms map_comp0s dtor_Jmap_unique_thm) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) end; @@ -1920,7 +1934,7 @@ map2 (fn goal => fn rec_Suc => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN - mk_set_incl_Jset_tac rec_Suc) + mk_set_incl_Jset_tac ctxt rec_Suc) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) goals rec_Sucs) @@ -1930,7 +1944,7 @@ map2 (fn goal => fn rec_Suc => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN - mk_set_Jset_incl_Jset_tac n rec_Suc k) + mk_set_Jset_incl_Jset_tac ctxt n rec_Suc k) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) goals rec_Sucs) @@ -1966,7 +1980,7 @@ (replicate n ballI @ maps (map (fn thm => thm RS @{thm subset_CollectI})) set_set_inclss)) |> singleton (Proof_Context.export names_lthy lthy) - |> rule_by_tactic lthy (ALLGOALS (TRY o etac asm_rl))) + |> rule_by_tactic lthy (ALLGOALS (TRY o etac lthy asm_rl))) Jset_minimal_thms set_Jset_incl_Jset_thmsss' Jsetss_by_range ys ys' dtor_set_induct_phiss end; @@ -1988,7 +2002,8 @@ val set_le_thmss = map split_conj_thm (@{map 4} (fn goal => fn Jset_minimal => fn set_Jsets => fn set_Jset_Jsetss => Goal.prove_sorry lthy [] [] goal - (K (mk_set_le_tac n Jset_minimal set_Jsets set_Jset_Jsetss)) + (fn {context = ctxt, prems = _} => + mk_set_le_tac ctxt n Jset_minimal set_Jsets set_Jset_Jsetss) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) le_goals Jset_minimal_thms set_Jset_thmss' set_Jset_Jset_thmsss'); @@ -1997,7 +2012,8 @@ val set_ge_thmss = @{map 3} (@{map 3} (fn goal => fn set_incl_Jset => fn set_Jset_incl_Jsets => Goal.prove_sorry lthy [] [] goal - (K (mk_set_ge_tac n set_incl_Jset set_Jset_incl_Jsets)) + (fn {context = ctxt, prems = _} => + mk_set_ge_tac ctxt n set_incl_Jset set_Jset_incl_Jsets) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy))) ge_goalss set_incl_Jset_thmss' set_Jset_incl_Jset_thmsss' @@ -2054,7 +2070,7 @@ @{map 5} (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs => Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jbd_defs THEN - mk_col_bd_tac m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss) + mk_col_bd_tac ctxt m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) ls goals ctss col_0ss' col_Sucss'; @@ -2099,8 +2115,9 @@ val thm = Goal.prove_sorry lthy [] [] goal - (K (mk_mcong_tac lthy m (rtac coinduct) map_comps dtor_Jmap_thms map_cong0s - set_mapss set_Jset_thmss set_Jset_Jset_thmsss in_rels)) + (fn {context = ctxt, prems = _} => mk_mcong_tac ctxt m (rtac ctxt coinduct) map_comps + dtor_Jmap_thms map_cong0s + set_mapss set_Jset_thmss set_Jset_Jset_thmsss in_rels) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); in @@ -2268,8 +2285,8 @@ val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals); in - Goal.prove_sorry lthy [] [] goal - (K (mk_le_rel_OO_tac Jrel_coinduct_thm dtor_Jrel_thms le_rel_OOs)) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_le_rel_OO_tac ctxt Jrel_coinduct_thm dtor_Jrel_thms le_rel_OOs) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -2412,27 +2429,31 @@ val timer = time (timer "witnesses"); val map_id0_tacs = - map2 (K oo mk_map_id0_tac Jmap_thms) dtor_unfold_unique_thms unfold_dtor_thms; - val map_comp0_tacs = map (fn thm => K (rtac (thm RS sym) 1)) Jmap_comp0_thms; + map2 (fn thm => fn thm' => fn ctxt => + mk_map_id0_tac ctxt Jmap_thms thm thm') + dtor_unfold_unique_thms unfold_dtor_thms; + val map_comp0_tacs = map (fn thm => fn ctxt => rtac ctxt (thm RS sym) 1) Jmap_comp0_thms; val map_cong0_tacs = map (fn thm => fn ctxt => mk_map_cong0_tac ctxt m thm) map_cong0_thms; val set_map0_tacss = - map (map (fn col => fn ctxt => unfold_thms_tac ctxt Jset_defs THEN mk_set_map0_tac col)) - (transpose col_natural_thmss); + map (map (fn col => fn ctxt => + unfold_thms_tac ctxt Jset_defs THEN mk_set_map0_tac ctxt col)) + (transpose col_natural_thmss); val Jbd_card_orders = map (fn def => fold_thms lthy [def] sbd_card_order) Jbd_defs; val Jbd_Cinfinites = map (fn def => fold_thms lthy [def] sbd_Cinfinite) Jbd_defs; - val bd_co_tacs = map (fn thm => K (rtac thm 1)) Jbd_card_orders; - val bd_cinf_tacs = map (fn thm => K (rtac (thm RS conjunct1) 1)) Jbd_Cinfinites; + val bd_co_tacs = map (fn thm => fn ctxt => rtac ctxt thm 1) Jbd_card_orders; + val bd_cinf_tacs = map (fn thm => fn ctxt => rtac ctxt (thm RS conjunct1) 1) Jbd_Cinfinites; val set_bd_tacss = map2 (fn Cinf => map (fn col => fn ctxt => - unfold_thms_tac ctxt Jset_defs THEN mk_set_bd_tac Cinf col)) + unfold_thms_tac ctxt Jset_defs THEN mk_set_bd_tac ctxt Cinf col)) Jbd_Cinfinites (transpose col_bd_thmss); - val le_rel_OO_tacs = map (fn i => K (rtac (le_Jrel_OO_thm RS mk_conjunctN n i) 1)) ks; + val le_rel_OO_tacs = map (fn i => fn ctxt => + rtac ctxt (le_Jrel_OO_thm RS mk_conjunctN n i) 1) ks; - val rel_OO_Grp_tacs = map (fn def => K (rtac def 1)) Jrel_unabs_defs; + val rel_OO_Grp_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Jrel_unabs_defs; val tacss = @{map 9} zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML --- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Thu Jul 16 12:23:22 2015 +0200 @@ -65,9 +65,9 @@ cut_tac nchotomy THEN' K (exhaust_inst_as_projs_tac ctxt frees) THEN' EVERY' (map (fn k => - (if k < n then etac disjE else K all_tac) THEN' - REPEAT o (dtac meta_mp THEN' atac ORELSE' - etac conjE THEN' dtac meta_mp THEN' atac ORELSE' + (if k < n then etac ctxt disjE else K all_tac) THEN' + REPEAT o (dtac ctxt meta_mp THEN' atac ORELSE' + etac ctxt conjE THEN' dtac ctxt meta_mp THEN' atac ORELSE' atac)) ks)) end; @@ -75,26 +75,26 @@ fun mk_primcorec_assumption_tac ctxt discIs = SELECT_GOAL (unfold_thms_tac ctxt @{thms fst_conv snd_conv not_not not_False_eq_True not_True_eq_False de_Morgan_conj de_Morgan_disj} THEN - SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE' etac conjE ORELSE' + SOLVE (HEADGOAL (REPEAT o (rtac ctxt refl ORELSE' atac ORELSE' etac ctxt conjE ORELSE' eresolve_tac ctxt falseEs ORELSE' resolve_tac ctxt @{thms TrueI conjI disjI1 disjI2} ORELSE' dresolve_tac ctxt discIs THEN' atac ORELSE' - etac notE THEN' atac ORELSE' - etac disjE)))); + etac ctxt notE THEN' atac ORELSE' + etac ctxt disjE)))); fun ss_fst_snd_conv ctxt = simpset_of (ss_only @{thms fst_conv snd_conv} ctxt); fun case_atac ctxt = simp_tac (put_simpset (ss_fst_snd_conv ctxt) ctxt); fun same_case_tac ctxt m = - HEADGOAL (if m = 0 then rtac TrueI - else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' case_atac ctxt) THEN' case_atac ctxt); + HEADGOAL (if m = 0 then rtac ctxt TrueI + else REPEAT_DETERM_N (m - 1) o (rtac ctxt conjI THEN' case_atac ctxt) THEN' case_atac ctxt); fun different_case_tac ctxt m exclude = HEADGOAL (if m = 0 then mk_primcorec_assumption_tac ctxt [] else - dtac exclude THEN' (REPEAT_DETERM_N (m - 1) o case_atac ctxt) THEN' + dtac ctxt exclude THEN' (REPEAT_DETERM_N (m - 1) o case_atac ctxt) THEN' mk_primcorec_assumption_tac ctxt []); fun cases_tac ctxt k m excludesss = @@ -105,51 +105,51 @@ end; fun prelude_tac ctxt defs thm = - unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN unfold_thms_tac ctxt unfold_lets; + unfold_thms_tac ctxt defs THEN HEADGOAL (rtac ctxt thm) THEN unfold_thms_tac ctxt unfold_lets; fun mk_primcorec_disc_tac ctxt defs corec_disc k m excludesss = prelude_tac ctxt defs corec_disc THEN cases_tac ctxt k m excludesss; fun mk_primcorec_disc_iff_tac ctxt fun_exhaust_frees fun_exhaust fun_discs fun_discss distinct_discs = - HEADGOAL (rtac iffI THEN' - rtac fun_exhaust THEN' + HEADGOAL (rtac ctxt iffI THEN' + rtac ctxt fun_exhaust THEN' K (exhaust_inst_as_projs_tac ctxt fun_exhaust_frees) THEN' - EVERY' (map (fn [] => etac FalseE + EVERY' (map (fn [] => etac ctxt FalseE | fun_discs' as [fun_disc'] => if eq_list Thm.eq_thm (fun_discs', fun_discs) then - REPEAT_DETERM o etac conjI THEN' (atac ORELSE' rtac TrueI) + REPEAT_DETERM o etac ctxt conjI THEN' (atac ORELSE' rtac ctxt TrueI) else - rtac FalseE THEN' - (rotate_tac 1 THEN' dtac fun_disc' THEN' REPEAT o atac ORELSE' + rtac ctxt FalseE THEN' + (rotate_tac 1 THEN' dtac ctxt fun_disc' THEN' REPEAT o atac ORELSE' cut_tac fun_disc') THEN' - dresolve_tac ctxt distinct_discs THEN' etac notE THEN' atac) + dresolve_tac ctxt distinct_discs THEN' etac ctxt notE THEN' atac) fun_discss) THEN' - (etac FalseE ORELSE' + (etac ctxt FalseE ORELSE' resolve_tac ctxt (map (unfold_thms ctxt [atomize_conjL]) fun_discs) THEN_MAYBE' atac)); fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms mapsx map_ident0s map_comps fun_sel k m excludesss = prelude_tac ctxt defs (fun_sel RS trans) THEN cases_tac ctxt k m excludesss THEN - HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE' + HEADGOAL (REPEAT_DETERM o (rtac ctxt refl ORELSE' eresolve_tac ctxt falseEs ORELSE' resolve_tac ctxt split_connectI ORELSE' Splitter.split_asm_tac ctxt (split_if_asm :: split_asms) ORELSE' Splitter.split_tac ctxt (split_if :: splits) ORELSE' eresolve_tac ctxt (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE' - etac notE THEN' atac ORELSE' + etac ctxt notE THEN' atac ORELSE' (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt (@{thms fst_conv snd_conv id_def comp_def split_def sum.case sum.sel sum.distinct[THEN eq_False[THEN iffD2]]} @ mapsx @ map_ident0s @ map_comps))) ORELSE' - fo_rtac @{thm cong} ctxt ORELSE' - rtac @{thm ext} ORELSE' + fo_rtac ctxt @{thm cong} ORELSE' + rtac ctxt @{thm ext} ORELSE' mk_primcorec_assumption_tac ctxt [])); fun mk_primcorec_ctr_tac ctxt m collapse disc_fun_opt sel_funs = - HEADGOAL (rtac ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN' - (the_default (K all_tac) (Option.map rtac disc_fun_opt)) THEN' REPEAT_DETERM_N m o atac) THEN - unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac refl); + HEADGOAL (rtac ctxt ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN' + (the_default (K all_tac) (Option.map (rtac ctxt) disc_fun_opt)) THEN' REPEAT_DETERM_N m o atac) THEN + unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac ctxt refl); fun inst_split_eq ctxt split = (case Thm.prop_of split of @@ -172,13 +172,13 @@ prelude_tac ctxt [] (fun_ctr RS trans) THEN HEADGOAL ((REPEAT_DETERM_N m o mk_primcorec_assumption_tac ctxt discIs) THEN' SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o - (rtac refl ORELSE' atac ORELSE' + (rtac ctxt refl ORELSE' atac ORELSE' resolve_tac ctxt (@{thm Code.abort_def} :: split_connectI) ORELSE' Splitter.split_tac ctxt (split_if :: splits) ORELSE' Splitter.split_asm_tac ctxt (split_if_asm :: split_asms) ORELSE' mk_primcorec_assumption_tac ctxt discIs ORELSE' distinct_in_prems_tac distincts ORELSE' - (TRY o dresolve_tac ctxt discIs) THEN' etac notE THEN' atac))))) + (TRY o dresolve_tac ctxt discIs) THEN' etac ctxt notE THEN' atac))))) end; fun rulify_nchotomy n = funpow (n - 1) (fn thm => thm RS @{thm Meson.make_pos_rule'}); @@ -203,13 +203,13 @@ end; fun mk_primcorec_code_tac ctxt distincts splits raw = - HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN' + HEADGOAL (rtac ctxt raw ORELSE' rtac ctxt (raw RS trans) THEN' SELECT_GOAL (unfold_thms_tac ctxt unfold_lets) THEN' REPEAT_DETERM o - (rtac refl ORELSE' atac ORELSE' + (rtac ctxt refl ORELSE' atac ORELSE' resolve_tac ctxt split_connectI ORELSE' Splitter.split_tac ctxt (split_if :: splits) ORELSE' distinct_in_prems_tac distincts ORELSE' - rtac sym THEN' atac ORELSE' - etac notE THEN' atac)); + rtac ctxt sym THEN' atac ORELSE' + etac ctxt notE THEN' atac)); end; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_gfp_tactics.ML --- a/src/HOL/Tools/BNF/bnf_gfp_tactics.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_gfp_tactics.ML Thu Jul 16 12:23:22 2015 +0200 @@ -12,82 +12,83 @@ val mk_bis_Gr_tac: Proof.context -> thm -> thm list -> thm list -> thm list -> thm list -> tactic val mk_bis_O_tac: Proof.context -> int -> thm -> thm list -> thm list -> tactic val mk_bis_Union_tac: Proof.context -> thm -> thm list -> tactic - val mk_bis_converse_tac: int -> thm -> thm list -> thm list -> tactic - val mk_bis_rel_tac: int -> thm -> thm list -> thm list -> thm list -> + val mk_bis_converse_tac: Proof.context -> int -> thm -> thm list -> thm list -> tactic + val mk_bis_rel_tac: Proof.context -> int -> thm -> thm list -> thm list -> thm list -> thm list list -> tactic val mk_coalgT_tac: Proof.context -> int -> thm list -> thm list -> thm list list -> tactic - val mk_coalg_final_tac: int -> thm -> thm list -> thm list -> thm list list -> thm list list -> - tactic - val mk_coalg_set_tac: thm -> tactic + val mk_coalg_final_tac: Proof.context -> int -> thm -> thm list -> thm list -> thm list list -> + thm list list -> tactic + val mk_coalg_set_tac: Proof.context -> thm -> tactic val mk_coind_wit_tac: Proof.context -> thm -> thm list -> thm list -> thm list -> tactic - val mk_col_bd_tac: int -> int -> cterm option list -> thm list -> thm list -> thm -> thm -> - thm list list -> tactic + val mk_col_bd_tac: Proof.context -> int -> int -> cterm option list -> thm list -> thm list -> + thm -> thm -> thm list list -> tactic val mk_col_natural_tac: Proof.context -> cterm option list -> thm list -> thm list -> thm list -> thm list list -> tactic - val mk_congruent_str_final_tac: int -> thm -> thm -> thm -> thm list -> tactic + val mk_congruent_str_final_tac: Proof.context -> int -> thm -> thm -> thm -> thm list -> tactic val mk_corec_tac: Proof.context -> int -> thm list -> thm -> thm -> thm list -> tactic val mk_corec_unique_mor_tac: Proof.context -> thm list -> thm list -> thm -> tactic - val mk_dtor_coinduct_tac: int -> thm -> thm -> thm list -> tactic + val mk_dtor_coinduct_tac: Proof.context -> int -> thm -> thm -> thm list -> tactic val mk_dtor_corec_transfer_tac: Proof.context -> int -> int -> thm list -> thm list -> thm list -> thm list -> tactic val mk_dtor_rel_tac: Proof.context -> thm list -> int -> thm -> thm -> thm -> thm -> thm list -> thm -> thm -> thm list -> thm list -> thm list list -> tactic val mk_dtor_o_ctor_tac: Proof.context -> thm -> thm -> thm -> thm -> thm list -> tactic - val mk_equiv_lsbis_tac: thm -> thm -> thm -> thm -> thm -> thm -> tactic + val mk_equiv_lsbis_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> tactic val mk_Jset_minimal_tac: Proof.context -> int -> thm -> tactic val mk_col_minimal_tac: Proof.context -> int -> cterm option list -> thm list -> thm list -> tactic - val mk_incl_lsbis_tac: int -> int -> thm -> tactic - val mk_length_Lev'_tac: thm -> tactic + val mk_incl_lsbis_tac: Proof.context -> int -> int -> thm -> tactic + val mk_length_Lev'_tac: Proof.context -> thm -> tactic val mk_length_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> tactic - val mk_map_comp0_tac: thm list -> thm list -> thm -> tactic + val mk_map_comp0_tac: Proof.context -> thm list -> thm list -> thm -> tactic val mk_mcong_tac: Proof.context -> int -> (int -> tactic) -> thm list -> thm list -> thm list -> thm list list -> thm list list -> thm list list list -> thm list -> tactic - val mk_map_id0_tac: thm list -> thm -> thm -> tactic - val mk_map_tac: int -> int -> thm -> thm -> thm -> thm -> tactic + val mk_map_id0_tac: Proof.context -> thm list -> thm -> thm -> tactic + val mk_map_tac: Proof.context -> int -> int -> thm -> thm -> thm -> thm -> tactic val mk_dtor_map_unique_tac: Proof.context -> thm -> thm list -> tactic val mk_mor_Abs_tac: Proof.context -> thm list -> thm list -> tactic val mk_mor_Rep_tac: Proof.context -> thm list -> thm list -> thm list -> thm list list -> thm list -> thm list -> tactic - val mk_mor_T_final_tac: thm -> thm list -> thm list -> tactic - val mk_mor_UNIV_tac: thm list -> thm -> tactic + val mk_mor_T_final_tac: Proof.context -> thm -> thm list -> thm list -> tactic + val mk_mor_UNIV_tac: Proof.context -> thm list -> thm -> tactic val mk_mor_beh_tac: Proof.context -> int -> thm -> thm -> thm list -> thm list -> thm list -> thm list -> thm list list -> thm list list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list list -> thm list list list -> thm list list list -> thm list list -> thm list -> thm list -> tactic - val mk_mor_case_sum_tac: 'a list -> thm -> tactic + val mk_mor_case_sum_tac: Proof.context -> 'a list -> thm -> tactic val mk_mor_comp_tac: Proof.context -> thm -> thm list -> thm list -> thm list -> tactic - val mk_mor_elim_tac: thm -> tactic - val mk_mor_col_tac: int -> int -> cterm option list -> int -> thm list -> thm list -> - thm list -> thm list list -> thm list list -> tactic - val mk_mor_incl_tac: thm -> thm list -> tactic - val mk_mor_str_tac: 'a list -> thm -> tactic - val mk_mor_unfold_tac: int -> thm -> thm list -> thm list -> thm list -> thm list -> thm list -> - thm list -> tactic - val mk_raw_coind_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm list -> - thm list -> thm list -> thm -> thm list -> tactic + val mk_mor_elim_tac: Proof.context -> thm -> tactic + val mk_mor_col_tac: Proof.context -> int -> int -> cterm option list -> int -> thm list -> + thm list -> thm list -> thm list list -> thm list list -> tactic + val mk_mor_incl_tac: Proof.context -> thm -> thm list -> tactic + val mk_mor_str_tac: Proof.context -> 'a list -> thm -> tactic + val mk_mor_unfold_tac: Proof.context -> int -> thm -> thm list -> thm list -> thm list -> + thm list -> thm list -> thm list -> tactic + val mk_raw_coind_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> + thm -> thm list -> thm list -> thm list -> thm -> thm list -> tactic val mk_rel_coinduct_tac: Proof.context -> thm list -> thm list -> thm list -> thm list list -> thm list -> thm list -> tactic val mk_rel_coinduct_coind_tac: Proof.context -> bool -> int -> thm -> int list -> thm list -> thm list -> thm list -> thm list list -> thm list -> thm list -> thm list -> tactic val mk_rel_coinduct_ind_tac: Proof.context -> int -> int list -> thm list -> thm list list -> int -> thm -> tactic - val mk_rv_last_tac: ctyp option list -> cterm option list -> thm list -> thm list -> tactic + val mk_rv_last_tac: Proof.context -> ctyp option list -> cterm option list -> thm list -> + thm list -> tactic val mk_sbis_lsbis_tac: Proof.context -> thm list -> thm -> thm -> tactic val mk_set_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> thm list -> thm list -> thm list list -> tactic - val mk_set_bd_tac: thm -> thm -> tactic - val mk_set_Jset_incl_Jset_tac: int -> thm -> int -> tactic + val mk_set_bd_tac: Proof.context -> thm -> thm -> tactic + val mk_set_Jset_incl_Jset_tac: Proof.context -> int -> thm -> int -> tactic val mk_set_image_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> thm list -> thm list -> thm list list -> thm list list -> tactic - val mk_set_incl_Jset_tac: thm -> tactic - val mk_set_ge_tac: int -> thm -> thm list -> tactic - val mk_set_le_tac: int -> thm -> thm list -> thm list list -> tactic - val mk_set_map0_tac: thm -> tactic - val mk_unfold_unique_mor_tac: thm list -> thm -> thm -> thm list -> tactic + val mk_set_incl_Jset_tac: Proof.context -> thm -> tactic + val mk_set_ge_tac: Proof.context -> int -> thm -> thm list -> tactic + val mk_set_le_tac: Proof.context -> int -> thm -> thm list -> thm list list -> tactic + val mk_set_map0_tac: Proof.context -> thm -> tactic + val mk_unfold_unique_mor_tac: Proof.context -> thm list -> thm -> thm -> thm list -> tactic val mk_unfold_transfer_tac: Proof.context -> int -> thm -> thm list -> thm list -> tactic val mk_wit_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm list -> tactic - val mk_le_rel_OO_tac: thm -> thm list -> thm list -> tactic + val mk_le_rel_OO_tac: Proof.context -> thm -> thm list -> thm list -> tactic end; structure BNF_GFP_Tactics : BNF_GFP_TACTICS = @@ -112,196 +113,196 @@ val Un_cong = @{thm arg_cong2[of _ _ _ _ "op \"]}; val converse_shift = @{thm converse_subset_swap} RS iffD1; -fun mk_coalg_set_tac coalg_def = - dtac (coalg_def RS iffD1) 1 THEN - REPEAT_DETERM (etac conjE 1) THEN - EVERY' [dtac rev_bspec, atac] 1 THEN +fun mk_coalg_set_tac ctxt coalg_def = + dtac ctxt (coalg_def RS iffD1) 1 THEN + REPEAT_DETERM (etac ctxt conjE 1) THEN + EVERY' [dtac ctxt rev_bspec, atac] 1 THEN REPEAT_DETERM (eresolve0_tac [CollectE, conjE] 1) THEN atac 1; -fun mk_mor_elim_tac mor_def = - (dtac (mor_def RS iffD1) THEN' - REPEAT o etac conjE THEN' - TRY o rtac @{thm image_subsetI} THEN' - etac bspec THEN' +fun mk_mor_elim_tac ctxt mor_def = + (dtac ctxt (mor_def RS iffD1) THEN' + REPEAT o etac ctxt conjE THEN' + TRY o rtac ctxt @{thm image_subsetI} THEN' + etac ctxt bspec THEN' atac) 1; -fun mk_mor_incl_tac mor_def map_ids = - (rtac (mor_def RS iffD2) THEN' - rtac conjI THEN' - CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, etac (id_apply RS @{thm ssubst_mem})])) +fun mk_mor_incl_tac ctxt mor_def map_ids = + (rtac ctxt (mor_def RS iffD2) THEN' + rtac ctxt conjI THEN' + CONJ_WRAP' (K (EVERY' [rtac ctxt ballI, etac ctxt set_mp, etac ctxt (id_apply RS @{thm ssubst_mem})])) map_ids THEN' CONJ_WRAP' (fn thm => - (EVERY' [rtac ballI, rtac (thm RS trans), rtac sym, rtac (id_apply RS arg_cong)])) map_ids) 1; + (EVERY' [rtac ctxt ballI, rtac ctxt (thm RS trans), rtac ctxt sym, rtac ctxt (id_apply RS arg_cong)])) map_ids) 1; fun mk_mor_comp_tac ctxt mor_def mor_images morEs map_comp_ids = let - fun fbetw_tac image = EVERY' [rtac ballI, rtac (o_apply RS @{thm ssubst_mem}), etac image, - etac image, atac]; + fun fbetw_tac image = EVERY' [rtac ctxt ballI, rtac ctxt (o_apply RS @{thm ssubst_mem}), etac ctxt image, + etac ctxt image, atac]; fun mor_tac ((mor_image, morE), map_comp_id) = - EVERY' [rtac ballI, stac ctxt o_apply, rtac trans, rtac (map_comp_id RS sym), rtac trans, - etac (morE RS arg_cong), atac, etac morE, etac mor_image, atac]; + EVERY' [rtac ctxt ballI, stac ctxt o_apply, rtac ctxt trans, rtac ctxt (map_comp_id RS sym), rtac ctxt trans, + etac ctxt (morE RS arg_cong), atac, etac ctxt morE, etac ctxt mor_image, atac]; in - (rtac (mor_def RS iffD2) THEN' rtac conjI THEN' + (rtac ctxt (mor_def RS iffD2) THEN' rtac ctxt conjI THEN' CONJ_WRAP' fbetw_tac mor_images THEN' CONJ_WRAP' mor_tac ((mor_images ~~ morEs) ~~ map_comp_ids)) 1 end; -fun mk_mor_UNIV_tac morEs mor_def = +fun mk_mor_UNIV_tac ctxt morEs mor_def = let val n = length morEs; - fun mor_tac morE = EVERY' [rtac @{thm ext}, rtac trans, rtac o_apply, rtac trans, etac morE, - rtac UNIV_I, rtac sym, rtac o_apply]; + fun mor_tac morE = EVERY' [rtac ctxt @{thm ext}, rtac ctxt trans, rtac ctxt o_apply, rtac ctxt trans, etac ctxt morE, + rtac ctxt UNIV_I, rtac ctxt sym, rtac ctxt o_apply]; in - EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs, - rtac (mor_def RS iffD2), rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs, + EVERY' [rtac ctxt iffI, CONJ_WRAP' mor_tac morEs, + rtac ctxt (mor_def RS iffD2), rtac ctxt conjI, CONJ_WRAP' (K (rtac ctxt ballI THEN' rtac ctxt UNIV_I)) morEs, CONJ_WRAP' (fn i => - EVERY' [dtac (mk_conjunctN n i), rtac ballI, etac @{thm comp_eq_dest}]) (1 upto n)] 1 + EVERY' [dtac ctxt (mk_conjunctN n i), rtac ctxt ballI, etac ctxt @{thm comp_eq_dest}]) (1 upto n)] 1 end; -fun mk_mor_str_tac ks mor_UNIV = - (rtac (mor_UNIV RS iffD2) THEN' CONJ_WRAP' (K (rtac refl)) ks) 1; +fun mk_mor_str_tac ctxt ks mor_UNIV = + (rtac ctxt (mor_UNIV RS iffD2) THEN' CONJ_WRAP' (K (rtac ctxt refl)) ks) 1; -fun mk_mor_case_sum_tac ks mor_UNIV = - (rtac (mor_UNIV RS iffD2) THEN' CONJ_WRAP' (K (rtac @{thm case_sum_o_inj(1)[symmetric]})) ks) 1; +fun mk_mor_case_sum_tac ctxt ks mor_UNIV = + (rtac ctxt (mor_UNIV RS iffD2) THEN' CONJ_WRAP' (K (rtac ctxt @{thm case_sum_o_inj(1)[symmetric]})) ks) 1; -fun mk_set_incl_Jset_tac rec_Suc = - EVERY' (map rtac [@{thm SUP_upper2}, UNIV_I, @{thm ord_le_eq_trans}, @{thm Un_upper1}, sym, +fun mk_set_incl_Jset_tac ctxt rec_Suc = + EVERY' (map (rtac ctxt) [@{thm SUP_upper2}, UNIV_I, @{thm ord_le_eq_trans}, @{thm Un_upper1}, sym, rec_Suc]) 1; -fun mk_set_Jset_incl_Jset_tac n rec_Suc i = - EVERY' (map rtac [@{thm UN_least}, subsetI, @{thm UN_I}, UNIV_I, set_mp, equalityD2, rec_Suc, - UnI2, mk_UnIN n i] @ [etac @{thm UN_I}, atac]) 1; +fun mk_set_Jset_incl_Jset_tac ctxt n rec_Suc i = + EVERY' (map (rtac ctxt) [@{thm UN_least}, subsetI, @{thm UN_I}, UNIV_I, set_mp, equalityD2, + rec_Suc, UnI2, mk_UnIN n i] @ [etac ctxt @{thm UN_I}, atac]) 1; fun mk_col_minimal_tac ctxt m cts rec_0s rec_Sucs = - EVERY' [rtac (Drule.instantiate' [] cts nat_induct), - REPEAT_DETERM o rtac allI, + EVERY' [rtac ctxt (Drule.instantiate' [] cts nat_induct), + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn thm => EVERY' - [rtac ord_eq_le_trans, rtac thm, rtac @{thm empty_subsetI}]) rec_0s, - REPEAT_DETERM o rtac allI, + [rtac ctxt ord_eq_le_trans, rtac ctxt thm, rtac ctxt @{thm empty_subsetI}]) rec_0s, + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn rec_Suc => EVERY' - [rtac ord_eq_le_trans, rtac rec_Suc, + [rtac ctxt ord_eq_le_trans, rtac ctxt rec_Suc, if m = 0 then K all_tac - else (rtac @{thm Un_least} THEN' Goal.assume_rule_tac ctxt), - CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least})) - (K (EVERY' [rtac @{thm UN_least}, REPEAT_DETERM o eresolve_tac ctxt [allE, conjE], - rtac subset_trans, atac, Goal.assume_rule_tac ctxt])) rec_0s]) + else (rtac ctxt @{thm Un_least} THEN' Goal.assume_rule_tac ctxt), + CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 @{thm Un_least})) + (K (EVERY' [rtac ctxt @{thm UN_least}, REPEAT_DETERM o eresolve_tac ctxt [allE, conjE], + rtac ctxt subset_trans, atac, Goal.assume_rule_tac ctxt])) rec_0s]) rec_Sucs] 1; fun mk_Jset_minimal_tac ctxt n col_minimal = - (CONJ_WRAP' (K (EVERY' [rtac @{thm UN_least}, rtac rev_mp, rtac col_minimal, - EVERY' (replicate ((n + 1) * n) (Goal.assume_rule_tac ctxt)), rtac impI, + (CONJ_WRAP' (K (EVERY' [rtac ctxt @{thm UN_least}, rtac ctxt rev_mp, rtac ctxt col_minimal, + EVERY' (replicate ((n + 1) * n) (Goal.assume_rule_tac ctxt)), rtac ctxt impI, REPEAT_DETERM o eresolve_tac ctxt [allE, conjE], atac])) (1 upto n)) 1 -fun mk_mor_col_tac m n cts j rec_0s rec_Sucs morEs set_map0ss coalg_setss = - EVERY' [rtac (Drule.instantiate' [] cts nat_induct), - REPEAT_DETERM o rtac allI, - CONJ_WRAP' (fn thm => EVERY' (map rtac [impI, thm RS trans, thm RS sym])) rec_0s, - REPEAT_DETERM o rtac allI, +fun mk_mor_col_tac ctxt m n cts j rec_0s rec_Sucs morEs set_map0ss coalg_setss = + EVERY' [rtac ctxt (Drule.instantiate' [] cts nat_induct), + REPEAT_DETERM o rtac ctxt allI, + CONJ_WRAP' (fn thm => EVERY' (map (rtac ctxt) [impI, thm RS trans, thm RS sym])) rec_0s, + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn (rec_Suc, (morE, ((passive_set_map0s, active_set_map0s), coalg_sets))) => - EVERY' [rtac impI, rtac (rec_Suc RS trans), rtac (rec_Suc RS trans RS sym), + EVERY' [rtac ctxt impI, rtac ctxt (rec_Suc RS trans), rtac ctxt (rec_Suc RS trans RS sym), if m = 0 then K all_tac - else EVERY' [rtac Un_cong, rtac @{thm box_equals}, - rtac (nth passive_set_map0s (j - 1) RS sym), - rtac trans_fun_cong_image_id_id_apply, etac (morE RS arg_cong), atac], - CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 Un_cong)) + else EVERY' [rtac ctxt Un_cong, rtac ctxt @{thm box_equals}, + rtac ctxt (nth passive_set_map0s (j - 1) RS sym), + rtac ctxt trans_fun_cong_image_id_id_apply, etac ctxt (morE RS arg_cong), atac], + CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 Un_cong)) (fn (i, (set_map0, coalg_set)) => - EVERY' [rtac sym, rtac trans, rtac (refl RSN (2, @{thm SUP_cong})), - etac (morE RS sym RS arg_cong RS trans), atac, rtac set_map0, - rtac (@{thm UN_simps(10)} RS trans), rtac (refl RS @{thm SUP_cong}), - ftac coalg_set, atac, dtac set_mp, atac, rtac mp, rtac (mk_conjunctN n i), - REPEAT_DETERM o etac allE, atac, atac]) + EVERY' [rtac ctxt sym, rtac ctxt trans, rtac ctxt (refl RSN (2, @{thm SUP_cong})), + etac ctxt (morE RS sym RS arg_cong RS trans), atac, rtac ctxt set_map0, + rtac ctxt (@{thm UN_simps(10)} RS trans), rtac ctxt (refl RS @{thm SUP_cong}), + ftac ctxt coalg_set, atac, dtac ctxt set_mp, atac, rtac ctxt mp, rtac ctxt (mk_conjunctN n i), + REPEAT_DETERM o etac ctxt allE, atac, atac]) (rev ((1 upto n) ~~ (active_set_map0s ~~ coalg_sets)))]) (rec_Sucs ~~ (morEs ~~ (map (chop m) set_map0ss ~~ map (drop m) coalg_setss)))] 1; -fun mk_bis_rel_tac m bis_def in_rels map_comp0s map_cong0s set_map0ss = +fun mk_bis_rel_tac ctxt m bis_def in_rels map_comp0s map_cong0s set_map0ss = let val n = length in_rels; val thms = ((1 upto n) ~~ map_comp0s ~~ map_cong0s ~~ set_map0ss ~~ in_rels); fun mk_if_tac ((((i, map_comp0), map_cong0), set_map0s), in_rel) = - EVERY' [rtac allI, rtac allI, rtac impI, dtac (mk_conjunctN n i), - etac allE, etac allE, etac impE, atac, etac bexE, + EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, dtac ctxt (mk_conjunctN n i), + etac ctxt allE, etac ctxt allE, etac ctxt impE, atac, etac ctxt bexE, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], - rtac (in_rel RS iffD2), rtac exI, rtac (Drule.rotate_prems 1 conjI), - CONJ_WRAP' (fn thm => EVERY' [rtac trans, rtac trans, rtac map_comp0, rtac map_cong0, - REPEAT_DETERM_N m o rtac thm, REPEAT_DETERM_N n o rtac (@{thm comp_id} RS fun_cong), + rtac ctxt (in_rel RS iffD2), rtac ctxt exI, rtac ctxt (Drule.rotate_prems 1 conjI), + CONJ_WRAP' (fn thm => EVERY' [rtac ctxt trans, rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt map_cong0, + REPEAT_DETERM_N m o rtac ctxt thm, REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong), atac]) @{thms fst_diag_id snd_diag_id}, - rtac CollectI, + rtac ctxt CollectI, CONJ_WRAP' (fn (i, thm) => if i <= m - then EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans, - etac @{thm image_mono}, rtac @{thm image_subsetI}, rtac CollectI, - rtac @{thm case_prodI}, rtac refl] - else EVERY' [rtac ord_eq_le_trans, rtac trans, rtac thm, - rtac trans_fun_cong_image_id_id_apply, etac @{thm Collect_split_in_rel_leI}]) + then EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt thm, rtac ctxt subset_trans, + etac ctxt @{thm image_mono}, rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, + rtac ctxt @{thm case_prodI}, rtac ctxt refl] + else EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt trans, rtac ctxt thm, + rtac ctxt trans_fun_cong_image_id_id_apply, etac ctxt @{thm Collect_split_in_rel_leI}]) (1 upto (m + n) ~~ set_map0s)]; fun mk_only_if_tac ((((i, map_comp0), map_cong0), set_map0s), in_rel) = - EVERY' [dtac (mk_conjunctN n i), rtac allI, rtac allI, rtac impI, - etac allE, etac allE, etac impE, atac, - dtac (in_rel RS @{thm iffD1}), + EVERY' [dtac ctxt (mk_conjunctN n i), rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, + etac ctxt allE, etac ctxt allE, etac ctxt impE, atac, + dtac ctxt (in_rel RS @{thm iffD1}), REPEAT_DETERM o eresolve0_tac ([CollectE, conjE, exE] @ @{thms CollectE Collect_split_in_rel_leE}), - rtac bexI, rtac conjI, rtac trans, rtac map_comp0, rtac trans, rtac map_cong0, - REPEAT_DETERM_N m o rtac (@{thm id_comp} RS fun_cong), - REPEAT_DETERM_N n o rtac (@{thm comp_id} RS fun_cong), - atac, rtac trans, rtac map_comp0, rtac trans, rtac map_cong0, - REPEAT_DETERM_N m o rtac (@{thm id_comp} RS fun_cong), - REPEAT_DETERM_N n o rtac (@{thm comp_id} RS fun_cong), - rtac trans, rtac map_cong0, - REPEAT_DETERM_N m o EVERY' [rtac @{thm Collect_splitD}, etac set_mp, atac], - REPEAT_DETERM_N n o rtac refl, - atac, rtac CollectI, + rtac ctxt bexI, rtac ctxt conjI, rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0, + REPEAT_DETERM_N m o rtac ctxt (@{thm id_comp} RS fun_cong), + REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong), + atac, rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0, + REPEAT_DETERM_N m o rtac ctxt (@{thm id_comp} RS fun_cong), + REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong), + rtac ctxt trans, rtac ctxt map_cong0, + REPEAT_DETERM_N m o EVERY' [rtac ctxt @{thm Collect_splitD}, etac ctxt set_mp, atac], + REPEAT_DETERM_N n o rtac ctxt refl, + atac, rtac ctxt CollectI, CONJ_WRAP' (fn (i, thm) => - if i <= m then rtac subset_UNIV - else EVERY' [rtac ord_eq_le_trans, rtac trans, rtac thm, - rtac trans_fun_cong_image_id_id_apply, atac]) + if i <= m then rtac ctxt subset_UNIV + else EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt trans, rtac ctxt thm, + rtac ctxt trans_fun_cong_image_id_id_apply, atac]) (1 upto (m + n) ~~ set_map0s)]; in - EVERY' [rtac (bis_def RS trans), - rtac iffI, etac conjE, etac conjI, CONJ_WRAP' mk_if_tac thms, - etac conjE, etac conjI, CONJ_WRAP' mk_only_if_tac thms] 1 + EVERY' [rtac ctxt (bis_def RS trans), + rtac ctxt iffI, etac ctxt conjE, etac ctxt conjI, CONJ_WRAP' mk_if_tac thms, + etac ctxt conjE, etac ctxt conjI, CONJ_WRAP' mk_only_if_tac thms] 1 end; -fun mk_bis_converse_tac m bis_rel rel_congs rel_converseps = - EVERY' [rtac (bis_rel RS iffD2), dtac (bis_rel RS iffD1), - REPEAT_DETERM o etac conjE, rtac conjI, - CONJ_WRAP' (K (EVERY' [rtac converse_shift, etac subset_trans, - rtac equalityD2, rtac @{thm converse_Times}])) rel_congs, +fun mk_bis_converse_tac ctxt m bis_rel rel_congs rel_converseps = + EVERY' [rtac ctxt (bis_rel RS iffD2), dtac ctxt (bis_rel RS iffD1), + REPEAT_DETERM o etac ctxt conjE, rtac ctxt conjI, + CONJ_WRAP' (K (EVERY' [rtac ctxt converse_shift, etac ctxt subset_trans, + rtac ctxt equalityD2, rtac ctxt @{thm converse_Times}])) rel_congs, CONJ_WRAP' (fn (rel_cong, rel_conversep) => - EVERY' [rtac allI, rtac allI, rtac impI, - rtac (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}), - REPEAT_DETERM_N m o rtac @{thm conversep_eq}, - REPEAT_DETERM_N (length rel_congs) o rtac @{thm conversep_in_rel}, - rtac (rel_conversep RS sym RS @{thm eq_refl} RS @{thm predicate2D}), - REPEAT_DETERM o etac allE, - rtac @{thm conversepI}, etac mp, etac @{thm converseD}]) (rel_congs ~~ rel_converseps)] 1; + EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, + rtac ctxt (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}), + REPEAT_DETERM_N m o rtac ctxt @{thm conversep_eq}, + REPEAT_DETERM_N (length rel_congs) o rtac ctxt @{thm conversep_in_rel}, + rtac ctxt (rel_conversep RS sym RS @{thm eq_refl} RS @{thm predicate2D}), + REPEAT_DETERM o etac ctxt allE, + rtac ctxt @{thm conversepI}, etac ctxt mp, etac ctxt @{thm converseD}]) (rel_congs ~~ rel_converseps)] 1; fun mk_bis_O_tac ctxt m bis_rel rel_congs le_rel_OOs = - EVERY' [rtac (bis_rel RS iffD2), REPEAT_DETERM o dtac (bis_rel RS iffD1), - REPEAT_DETERM o etac conjE, rtac conjI, - CONJ_WRAP' (K (EVERY' [etac @{thm relcomp_subset_Sigma}, atac])) rel_congs, + EVERY' [rtac ctxt (bis_rel RS iffD2), REPEAT_DETERM o dtac ctxt (bis_rel RS iffD1), + REPEAT_DETERM o etac ctxt conjE, rtac ctxt conjI, + CONJ_WRAP' (K (EVERY' [etac ctxt @{thm relcomp_subset_Sigma}, atac])) rel_congs, CONJ_WRAP' (fn (rel_cong, le_rel_OO) => - EVERY' [rtac allI, rtac allI, rtac impI, - rtac (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}), - REPEAT_DETERM_N m o rtac @{thm eq_OO}, - REPEAT_DETERM_N (length rel_congs) o rtac @{thm relcompp_in_rel}, - rtac (le_rel_OO RS @{thm predicate2D}), - etac @{thm relcompE}, - REPEAT_DETERM o dtac prod_injectD, - etac conjE, hyp_subst_tac ctxt, - REPEAT_DETERM o etac allE, rtac @{thm relcomppI}, - etac mp, atac, etac mp, atac]) (rel_congs ~~ le_rel_OOs)] 1; + EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, + rtac ctxt (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}), + REPEAT_DETERM_N m o rtac ctxt @{thm eq_OO}, + REPEAT_DETERM_N (length rel_congs) o rtac ctxt @{thm relcompp_in_rel}, + rtac ctxt (le_rel_OO RS @{thm predicate2D}), + etac ctxt @{thm relcompE}, + REPEAT_DETERM o dtac ctxt prod_injectD, + etac ctxt conjE, hyp_subst_tac ctxt, + REPEAT_DETERM o etac ctxt allE, rtac ctxt @{thm relcomppI}, + etac ctxt mp, atac, etac ctxt mp, atac]) (rel_congs ~~ le_rel_OOs)] 1; fun mk_bis_Gr_tac ctxt bis_rel rel_Grps mor_images morEs coalg_ins = unfold_thms_tac ctxt (bis_rel :: @{thm eq_alt} :: @{thm in_rel_Gr} :: rel_Grps) THEN - EVERY' [rtac conjI, - CONJ_WRAP' (fn thm => rtac (@{thm Gr_incl} RS iffD2) THEN' etac thm) mor_images, + EVERY' [rtac ctxt conjI, + CONJ_WRAP' (fn thm => rtac ctxt (@{thm Gr_incl} RS iffD2) THEN' etac ctxt thm) mor_images, CONJ_WRAP' (fn (coalg_in, morE) => - EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm GrpI}, etac (morE RS trans), - etac @{thm GrD1}, etac (@{thm GrD2} RS arg_cong), etac coalg_in, etac @{thm GrD1}]) + EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, rtac ctxt @{thm GrpI}, etac ctxt (morE RS trans), + etac ctxt @{thm GrD1}, etac ctxt (@{thm GrD2} RS arg_cong), etac ctxt coalg_in, etac ctxt @{thm GrD1}]) (coalg_ins ~~ morEs)] 1; fun mk_bis_Union_tac ctxt bis_def in_monos = @@ -310,15 +311,15 @@ val ks = 1 upto n; in unfold_thms_tac ctxt [bis_def] THEN - EVERY' [rtac conjI, + EVERY' [rtac ctxt conjI, CONJ_WRAP' (fn i => - EVERY' [rtac @{thm UN_least}, dtac bspec, atac, - dtac conjunct1, etac (mk_conjunctN n i)]) ks, + EVERY' [rtac ctxt @{thm UN_least}, dtac ctxt bspec, atac, + dtac ctxt conjunct1, etac ctxt (mk_conjunctN n i)]) ks, CONJ_WRAP' (fn (i, in_mono) => - EVERY' [rtac allI, rtac allI, rtac impI, etac @{thm UN_E}, dtac bspec, atac, - dtac conjunct2, dtac (mk_conjunctN n i), etac allE, etac allE, dtac mp, - atac, etac bexE, rtac bexI, atac, rtac in_mono, - REPEAT_DETERM_N n o etac @{thm SUP_upper2[OF _ subset_refl]}, + EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, etac ctxt @{thm UN_E}, dtac ctxt bspec, atac, + dtac ctxt conjunct2, dtac ctxt (mk_conjunctN n i), etac ctxt allE, etac ctxt allE, dtac ctxt mp, + atac, etac ctxt bexE, rtac ctxt bexI, atac, rtac ctxt in_mono, + REPEAT_DETERM_N n o etac ctxt @{thm SUP_upper2[OF _ subset_refl]}, atac]) (ks ~~ in_monos)] 1 end; @@ -326,65 +327,65 @@ let val n = length lsbis_defs; in - EVERY' [rtac (Thm.permute_prems 0 1 bis_cong), EVERY' (map rtac lsbis_defs), - rtac bis_Union, rtac ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE, exE], - hyp_subst_tac ctxt, etac bis_cong, EVERY' (map (rtac o mk_nth_conv n) (1 upto n))] 1 + EVERY' [rtac ctxt (Thm.permute_prems 0 1 bis_cong), EVERY' (map (rtac ctxt) lsbis_defs), + rtac ctxt bis_Union, rtac ctxt ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE, exE], + hyp_subst_tac ctxt, etac ctxt bis_cong, EVERY' (map (rtac ctxt o mk_nth_conv n) (1 upto n))] 1 end; -fun mk_incl_lsbis_tac n i lsbis_def = - EVERY' [rtac @{thm xt1(3)}, rtac lsbis_def, rtac @{thm SUP_upper2}, rtac CollectI, - REPEAT_DETERM_N n o rtac exI, rtac conjI, rtac refl, atac, rtac equalityD2, - rtac (mk_nth_conv n i)] 1; +fun mk_incl_lsbis_tac ctxt n i lsbis_def = + EVERY' [rtac ctxt @{thm xt1(3)}, rtac ctxt lsbis_def, rtac ctxt @{thm SUP_upper2}, rtac ctxt CollectI, + REPEAT_DETERM_N n o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, atac, rtac ctxt equalityD2, + rtac ctxt (mk_nth_conv n i)] 1; -fun mk_equiv_lsbis_tac sbis_lsbis lsbis_incl incl_lsbis bis_Id_on bis_converse bis_O = - EVERY' [rtac (@{thm equiv_def} RS iffD2), +fun mk_equiv_lsbis_tac ctxt sbis_lsbis lsbis_incl incl_lsbis bis_Id_on bis_converse bis_O = + EVERY' [rtac ctxt (@{thm equiv_def} RS iffD2), - rtac conjI, rtac (@{thm refl_on_def} RS iffD2), - rtac conjI, rtac lsbis_incl, rtac ballI, rtac set_mp, - rtac incl_lsbis, rtac bis_Id_on, atac, etac @{thm Id_onI}, + rtac ctxt conjI, rtac ctxt (@{thm refl_on_def} RS iffD2), + rtac ctxt conjI, rtac ctxt lsbis_incl, rtac ctxt ballI, rtac ctxt set_mp, + rtac ctxt incl_lsbis, rtac ctxt bis_Id_on, atac, etac ctxt @{thm Id_onI}, - rtac conjI, rtac (@{thm sym_def} RS iffD2), - rtac allI, rtac allI, rtac impI, rtac set_mp, - rtac incl_lsbis, rtac bis_converse, rtac sbis_lsbis, etac @{thm converseI}, + rtac ctxt conjI, rtac ctxt (@{thm sym_def} RS iffD2), + rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, rtac ctxt set_mp, + rtac ctxt incl_lsbis, rtac ctxt bis_converse, rtac ctxt sbis_lsbis, etac ctxt @{thm converseI}, - rtac (@{thm trans_def} RS iffD2), - rtac allI, rtac allI, rtac allI, rtac impI, rtac impI, rtac set_mp, - rtac incl_lsbis, rtac bis_O, rtac sbis_lsbis, rtac sbis_lsbis, - etac @{thm relcompI}, atac] 1; + rtac ctxt (@{thm trans_def} RS iffD2), + rtac ctxt allI, rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, rtac ctxt impI, rtac ctxt set_mp, + rtac ctxt incl_lsbis, rtac ctxt bis_O, rtac ctxt sbis_lsbis, rtac ctxt sbis_lsbis, + etac ctxt @{thm relcompI}, atac] 1; fun mk_coalgT_tac ctxt m defs strT_defs set_map0ss = let val n = length strT_defs; val ks = 1 upto n; fun coalg_tac (i, (active_sets, def)) = - EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], - hyp_subst_tac ctxt, rtac (def RS trans RS @{thm ssubst_mem}), etac (arg_cong RS trans), - rtac (mk_sum_caseN n i), rtac CollectI, - REPEAT_DETERM_N m o EVERY' [rtac conjI, rtac subset_UNIV], - CONJ_WRAP' (fn (i, thm) => EVERY' [rtac (thm RS ord_eq_le_trans), - rtac @{thm image_subsetI}, rtac CollectI, rtac exI, rtac exI, rtac conjI, rtac refl, - rtac conjI, - rtac conjI, etac @{thm empty_Shift}, dtac set_rev_mp, - etac equalityD1, etac CollectD, - rtac ballI, - CONJ_WRAP' (fn i => EVERY' [rtac ballI, etac CollectE, dtac @{thm ShiftD}, - dtac bspec, etac thin_rl, atac, dtac (mk_conjunctN n i), - dtac bspec, rtac CollectI, etac @{thm set_mp[OF equalityD1[OF Succ_Shift]]}, - REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], rtac exI, - rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans), - rtac (@{thm append_Cons} RS sym RS arg_cong RS trans), atac, - CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong}, - rtac @{thm eqset_imp_iff}, rtac sym, rtac trans, rtac @{thm Succ_Shift}, - rtac (@{thm append_Cons} RS sym RS arg_cong)])) ks]) ks, - dtac bspec, atac, dtac (mk_conjunctN n i), dtac bspec, - etac @{thm set_mp[OF equalityD1]}, atac, - REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], rtac exI, - rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans), - etac (@{thm append_Nil} RS sym RS arg_cong RS trans), - REPEAT_DETERM_N m o (rtac conjI THEN' atac), - CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong}, - rtac @{thm eqset_imp_iff}, rtac sym, rtac trans, rtac @{thm Succ_Shift}, - rtac (@{thm append_Nil} RS sym RS arg_cong)])) ks]) (ks ~~ active_sets)]; + EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], + hyp_subst_tac ctxt, rtac ctxt (def RS trans RS @{thm ssubst_mem}), etac ctxt (arg_cong RS trans), + rtac ctxt (mk_sum_caseN n i), rtac ctxt CollectI, + REPEAT_DETERM_N m o EVERY' [rtac ctxt conjI, rtac ctxt subset_UNIV], + CONJ_WRAP' (fn (i, thm) => EVERY' [rtac ctxt (thm RS ord_eq_le_trans), + rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, rtac ctxt exI, rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, + rtac ctxt conjI, + rtac ctxt conjI, etac ctxt @{thm empty_Shift}, dtac ctxt set_rev_mp, + etac ctxt equalityD1, etac ctxt CollectD, + rtac ctxt ballI, + CONJ_WRAP' (fn i => EVERY' [rtac ctxt ballI, etac ctxt CollectE, dtac ctxt @{thm ShiftD}, + dtac ctxt bspec, etac ctxt thin_rl, atac, dtac ctxt (mk_conjunctN n i), + dtac ctxt bspec, rtac ctxt CollectI, etac ctxt @{thm set_mp[OF equalityD1[OF Succ_Shift]]}, + REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], rtac ctxt exI, + rtac ctxt conjI, rtac ctxt (@{thm shift_def} RS fun_cong RS trans), + rtac ctxt (@{thm append_Cons} RS sym RS arg_cong RS trans), atac, + CONJ_WRAP' (K (EVERY' [etac ctxt trans, rtac ctxt @{thm Collect_cong}, + rtac ctxt @{thm eqset_imp_iff}, rtac ctxt sym, rtac ctxt trans, rtac ctxt @{thm Succ_Shift}, + rtac ctxt (@{thm append_Cons} RS sym RS arg_cong)])) ks]) ks, + dtac ctxt bspec, atac, dtac ctxt (mk_conjunctN n i), dtac ctxt bspec, + etac ctxt @{thm set_mp[OF equalityD1]}, atac, + REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], rtac ctxt exI, + rtac ctxt conjI, rtac ctxt (@{thm shift_def} RS fun_cong RS trans), + etac ctxt (@{thm append_Nil} RS sym RS arg_cong RS trans), + REPEAT_DETERM_N m o (rtac ctxt conjI THEN' atac), + CONJ_WRAP' (K (EVERY' [etac ctxt trans, rtac ctxt @{thm Collect_cong}, + rtac ctxt @{thm eqset_imp_iff}, rtac ctxt sym, rtac ctxt trans, rtac ctxt @{thm Succ_Shift}, + rtac ctxt (@{thm append_Nil} RS sym RS arg_cong)])) ks]) (ks ~~ active_sets)]; in unfold_thms_tac ctxt defs THEN CONJ_WRAP' coalg_tac (ks ~~ (map (drop m) set_map0ss ~~ strT_defs)) 1 @@ -395,45 +396,45 @@ val n = length Lev_0s; val ks = n downto 1; in - EVERY' [rtac (Drule.instantiate' [] cts nat_induct), - REPEAT_DETERM o rtac allI, + EVERY' [rtac ctxt (Drule.instantiate' [] cts nat_induct), + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn Lev_0 => - EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp), - etac @{thm singletonE}, etac ssubst, rtac @{thm list.size(3)}]) Lev_0s, - REPEAT_DETERM o rtac allI, + EVERY' [rtac ctxt impI, dtac ctxt (Lev_0 RS equalityD1 RS set_mp), + etac ctxt @{thm singletonE}, etac ctxt ssubst, rtac ctxt @{thm list.size(3)}]) Lev_0s, + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn Lev_Suc => - EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp), - CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) + EVERY' [rtac ctxt impI, dtac ctxt (Lev_Suc RS equalityD1 RS set_mp), + CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) (fn i => EVERY' [REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], hyp_subst_tac ctxt, - rtac trans, rtac @{thm length_Cons}, rtac @{thm arg_cong[of _ _ Suc]}, - REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i), etac mp, atac]) ks]) + rtac ctxt trans, rtac ctxt @{thm length_Cons}, rtac ctxt @{thm arg_cong[of _ _ Suc]}, + REPEAT_DETERM o etac ctxt allE, dtac ctxt (mk_conjunctN n i), etac ctxt mp, atac]) ks]) Lev_Sucs] 1 end; -fun mk_length_Lev'_tac length_Lev = - EVERY' [ftac length_Lev, etac ssubst, atac] 1; +fun mk_length_Lev'_tac ctxt length_Lev = + EVERY' [ftac ctxt length_Lev, etac ctxt ssubst, atac] 1; -fun mk_rv_last_tac cTs cts rv_Nils rv_Conss = +fun mk_rv_last_tac ctxt cTs cts rv_Nils rv_Conss = let val n = length rv_Nils; val ks = 1 upto n; in - EVERY' [rtac (Drule.instantiate' cTs cts @{thm list.induct}), - REPEAT_DETERM o rtac allI, + EVERY' [rtac ctxt (Drule.instantiate' cTs cts @{thm list.induct}), + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn rv_Cons => - CONJ_WRAP' (fn (i, rv_Nil) => (EVERY' [rtac exI, - rtac (@{thm append_Nil} RS arg_cong RS trans), - rtac (rv_Cons RS trans), rtac (mk_sum_caseN n i RS trans), rtac rv_Nil])) + CONJ_WRAP' (fn (i, rv_Nil) => (EVERY' [rtac ctxt exI, + rtac ctxt (@{thm append_Nil} RS arg_cong RS trans), + rtac ctxt (rv_Cons RS trans), rtac ctxt (mk_sum_caseN n i RS trans), rtac ctxt rv_Nil])) (ks ~~ rv_Nils)) rv_Conss, - REPEAT_DETERM o rtac allI, rtac (mk_sumEN n), + REPEAT_DETERM o rtac ctxt allI, rtac ctxt (mk_sumEN n), EVERY' (map (fn i => - CONJ_WRAP' (fn rv_Cons => EVERY' [REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i), - CONJ_WRAP' (fn i' => EVERY' [dtac (mk_conjunctN n i'), etac exE, rtac exI, - rtac (@{thm append_Cons} RS arg_cong RS trans), rtac (rv_Cons RS trans), - if n = 1 then K all_tac else etac (sum_case_cong_weak RS trans), - rtac (mk_sum_caseN n i RS trans), atac]) + CONJ_WRAP' (fn rv_Cons => EVERY' [REPEAT_DETERM o etac ctxt allE, dtac ctxt (mk_conjunctN n i), + CONJ_WRAP' (fn i' => EVERY' [dtac ctxt (mk_conjunctN n i'), etac ctxt exE, rtac ctxt exI, + rtac ctxt (@{thm append_Cons} RS arg_cong RS trans), rtac ctxt (rv_Cons RS trans), + if n = 1 then K all_tac else etac ctxt (sum_case_cong_weak RS trans), + rtac ctxt (mk_sum_caseN n i RS trans), atac]) ks]) rv_Conss) ks)] 1 @@ -444,41 +445,41 @@ val n = length Lev_0s; val ks = 1 upto n; in - EVERY' [rtac (Drule.instantiate' [] cts nat_induct), - REPEAT_DETERM o rtac allI, + EVERY' [rtac ctxt (Drule.instantiate' [] cts nat_induct), + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) => - EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp), - etac @{thm singletonE}, hyp_subst_tac ctxt, - CONJ_WRAP' (fn i' => rtac impI THEN' dtac (sym RS trans) THEN' rtac rv_Nil THEN' + EVERY' [rtac ctxt impI, dtac ctxt (Lev_0 RS equalityD1 RS set_mp), + etac ctxt @{thm singletonE}, hyp_subst_tac ctxt, + CONJ_WRAP' (fn i' => rtac ctxt impI THEN' dtac ctxt (sym RS trans) THEN' rtac ctxt rv_Nil THEN' (if i = i' - then EVERY' [dtac (mk_InN_inject n i), hyp_subst_tac ctxt, + then EVERY' [dtac ctxt (mk_InN_inject n i), hyp_subst_tac ctxt, CONJ_WRAP' (fn (i'', Lev_0'') => - EVERY' [rtac impI, rtac @{thm ssubst_mem[OF append_Nil]}, - rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i''), - rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, - etac conjI, rtac (Lev_0'' RS equalityD2 RS set_mp), - rtac @{thm singletonI}]) + EVERY' [rtac ctxt impI, rtac ctxt @{thm ssubst_mem[OF append_Nil]}, + rtac ctxt (Lev_Suc RS equalityD2 RS set_mp), rtac ctxt (mk_UnIN n i''), + rtac ctxt CollectI, REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, + etac ctxt conjI, rtac ctxt (Lev_0'' RS equalityD2 RS set_mp), + rtac ctxt @{thm singletonI}]) (ks ~~ Lev_0s)] - else etac (mk_InN_not_InM i' i RS notE))) + else etac ctxt (mk_InN_not_InM i' i RS notE))) ks]) ((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils), - REPEAT_DETERM o rtac allI, + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), from_to_sbds) => - EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp), - CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) + EVERY' [rtac ctxt impI, dtac ctxt (Lev_Suc RS equalityD1 RS set_mp), + CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) (fn (i, from_to_sbd) => EVERY' [REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], hyp_subst_tac ctxt, - CONJ_WRAP' (fn i' => rtac impI THEN' + CONJ_WRAP' (fn i' => rtac ctxt impI THEN' CONJ_WRAP' (fn i'' => - EVERY' [rtac impI, rtac (Lev_Suc RS equalityD2 RS set_mp), - rtac @{thm ssubst_mem[OF append_Cons]}, rtac (mk_UnIN n i), - rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, - rtac conjI, atac, dtac (sym RS trans RS sym), - rtac (rv_Cons RS trans), rtac (mk_sum_caseN n i RS trans), - etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE, - dtac (mk_conjunctN n i), dtac mp, atac, - dtac (mk_conjunctN n i'), dtac mp, atac, - dtac (mk_conjunctN n i''), etac mp, atac]) + EVERY' [rtac ctxt impI, rtac ctxt (Lev_Suc RS equalityD2 RS set_mp), + rtac ctxt @{thm ssubst_mem[OF append_Cons]}, rtac ctxt (mk_UnIN n i), + rtac ctxt CollectI, REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, + rtac ctxt conjI, atac, dtac ctxt (sym RS trans RS sym), + rtac ctxt (rv_Cons RS trans), rtac ctxt (mk_sum_caseN n i RS trans), + etac ctxt (from_to_sbd RS arg_cong), REPEAT_DETERM o etac ctxt allE, + dtac ctxt (mk_conjunctN n i), dtac ctxt mp, atac, + dtac ctxt (mk_conjunctN n i'), dtac ctxt mp, atac, + dtac ctxt (mk_conjunctN n i''), etac ctxt mp, atac]) ks) ks]) (rev (ks ~~ from_to_sbds))]) @@ -490,54 +491,54 @@ val n = length Lev_0s; val ks = 1 upto n; in - EVERY' [rtac (Drule.instantiate' [] cts nat_induct), - REPEAT_DETERM o rtac allI, + EVERY' [rtac ctxt (Drule.instantiate' [] cts nat_induct), + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) => - EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp), - etac @{thm singletonE}, hyp_subst_tac ctxt, - CONJ_WRAP' (fn i' => rtac impI THEN' - CONJ_WRAP' (fn i'' => rtac impI THEN' dtac (sym RS trans) THEN' rtac rv_Nil THEN' + EVERY' [rtac ctxt impI, dtac ctxt (Lev_0 RS equalityD1 RS set_mp), + etac ctxt @{thm singletonE}, hyp_subst_tac ctxt, + CONJ_WRAP' (fn i' => rtac ctxt impI THEN' + CONJ_WRAP' (fn i'' => rtac ctxt impI THEN' dtac ctxt (sym RS trans) THEN' rtac ctxt rv_Nil THEN' (if i = i'' - then EVERY' [dtac @{thm ssubst_mem[OF sym[OF append_Nil]]}, - dtac (Lev_Suc RS equalityD1 RS set_mp), dtac (mk_InN_inject n i), + then EVERY' [dtac ctxt @{thm ssubst_mem[OF sym[OF append_Nil]]}, + dtac ctxt (Lev_Suc RS equalityD1 RS set_mp), dtac ctxt (mk_InN_inject n i), hyp_subst_tac ctxt, - CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) + CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) (fn k => REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE] THEN' - dtac list_inject_iffD1 THEN' etac conjE THEN' + dtac ctxt list_inject_iffD1 THEN' etac ctxt conjE THEN' (if k = i' - then EVERY' [dtac (mk_InN_inject n k), hyp_subst_tac ctxt, etac imageI] - else etac (mk_InN_not_InM i' k RS notE))) + then EVERY' [dtac ctxt (mk_InN_inject n k), hyp_subst_tac ctxt, etac ctxt imageI] + else etac ctxt (mk_InN_not_InM i' k RS notE))) (rev ks)] - else etac (mk_InN_not_InM i'' i RS notE))) + else etac ctxt (mk_InN_not_InM i'' i RS notE))) ks) ks]) ((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils), - REPEAT_DETERM o rtac allI, + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), (from_to_sbds, to_sbd_injs)) => - EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp), - CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) + EVERY' [rtac ctxt impI, dtac ctxt (Lev_Suc RS equalityD1 RS set_mp), + CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) (fn (i, (from_to_sbd, to_sbd_inj)) => REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE] THEN' hyp_subst_tac ctxt THEN' - CONJ_WRAP' (fn i' => rtac impI THEN' - dtac @{thm ssubst_mem[OF sym[OF append_Cons]]} THEN' - dtac (Lev_Suc RS equalityD1 RS set_mp) THEN' - CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn k => + CONJ_WRAP' (fn i' => rtac ctxt impI THEN' + dtac ctxt @{thm ssubst_mem[OF sym[OF append_Cons]]} THEN' + dtac ctxt (Lev_Suc RS equalityD1 RS set_mp) THEN' + CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) (fn k => REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE] THEN' - dtac list_inject_iffD1 THEN' etac conjE THEN' + dtac ctxt list_inject_iffD1 THEN' etac ctxt conjE THEN' (if k = i - then EVERY' [dtac (mk_InN_inject n i), - dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)), + then EVERY' [dtac ctxt (mk_InN_inject n i), + dtac ctxt (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)), atac, atac, hyp_subst_tac ctxt] THEN' CONJ_WRAP' (fn i'' => - EVERY' [rtac impI, dtac (sym RS trans), - rtac (rv_Cons RS trans), rtac (mk_sum_caseN n i RS arg_cong RS trans), - etac (from_to_sbd RS arg_cong), - REPEAT_DETERM o etac allE, - dtac (mk_conjunctN n i), dtac mp, atac, - dtac (mk_conjunctN n i'), dtac mp, atac, - dtac (mk_conjunctN n i''), etac mp, etac sym]) + EVERY' [rtac ctxt impI, dtac ctxt (sym RS trans), + rtac ctxt (rv_Cons RS trans), rtac ctxt (mk_sum_caseN n i RS arg_cong RS trans), + etac ctxt (from_to_sbd RS arg_cong), + REPEAT_DETERM o etac ctxt allE, + dtac ctxt (mk_conjunctN n i), dtac ctxt mp, atac, + dtac ctxt (mk_conjunctN n i'), dtac ctxt mp, atac, + dtac ctxt (mk_conjunctN n i''), etac ctxt mp, etac ctxt sym]) ks - else etac (mk_InN_not_InM i k RS notE))) + else etac ctxt (mk_InN_not_InM i k RS notE))) (rev ks)) ks) (rev (ks ~~ (from_to_sbds ~~ to_sbd_injs)))]) @@ -553,90 +554,90 @@ fun fbetw_tac (i, (carT_def, (isNode_def, (Lev_0, (rv_Nil, ((length_Lev, length_Lev'), (rv_lasts, (set_map0s, (set_Levss, set_image_Levss))))))))) = - EVERY' [rtac ballI, rtac (carT_def RS equalityD2 RS set_mp), - rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, rtac conjI, - rtac conjI, - rtac @{thm UN_I}, rtac UNIV_I, rtac (Lev_0 RS equalityD2 RS set_mp), - rtac @{thm singletonI}, + EVERY' [rtac ctxt ballI, rtac ctxt (carT_def RS equalityD2 RS set_mp), + rtac ctxt CollectI, REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, rtac ctxt conjI, + rtac ctxt conjI, + rtac ctxt @{thm UN_I}, rtac ctxt UNIV_I, rtac ctxt (Lev_0 RS equalityD2 RS set_mp), + rtac ctxt @{thm singletonI}, (**) - rtac ballI, etac @{thm UN_E}, + rtac ctxt ballI, etac ctxt @{thm UN_E}, CONJ_WRAP' (fn (i, (rv_last, (isNode_def, (set_map0s, (set_Levs, set_image_Levs))))) => - EVERY' [rtac ballI, + EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}], - rtac (rev_mp OF [rv_last, impI]), etac exE, rtac (isNode_def RS iffD2), - rtac exI, rtac conjI, - if n = 1 then rtac refl - else etac (sum_case_cong_weak RS trans) THEN' rtac (mk_sum_caseN n i), + rtac ctxt (rev_mp OF [rv_last, impI]), etac ctxt exE, rtac ctxt (isNode_def RS iffD2), + rtac ctxt exI, rtac ctxt conjI, + if n = 1 then rtac ctxt refl + else etac ctxt (sum_case_cong_weak RS trans) THEN' rtac ctxt (mk_sum_caseN n i), CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) => - EVERY' [rtac (set_map0 RS trans), rtac equalityI, rtac @{thm image_subsetI}, - rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev, - if n = 1 then rtac refl else atac, atac, rtac subsetI, + EVERY' [rtac ctxt (set_map0 RS trans), rtac ctxt equalityI, rtac ctxt @{thm image_subsetI}, + rtac ctxt CollectI, rtac ctxt @{thm SuccI}, rtac ctxt @{thm UN_I}, rtac ctxt UNIV_I, etac ctxt set_Lev, + if n = 1 then rtac ctxt refl else atac, atac, rtac ctxt subsetI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}], - REPEAT_DETERM_N 4 o etac thin_rl, - rtac set_image_Lev, - atac, dtac length_Lev, hyp_subst_tac ctxt, dtac length_Lev', - etac @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]}, - if n = 1 then rtac refl else atac]) + REPEAT_DETERM_N 4 o etac ctxt thin_rl, + rtac ctxt set_image_Lev, + atac, dtac ctxt length_Lev, hyp_subst_tac ctxt, dtac ctxt length_Lev', + etac ctxt @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]}, + if n = 1 then rtac ctxt refl else atac]) (drop m set_map0s ~~ (set_Levs ~~ set_image_Levs))]) (ks ~~ (rv_lasts ~~ (isNode_defs ~~ (set_map0ss ~~ (set_Levss ~~ set_image_Levss))))), (*root isNode*) - rtac (isNode_def RS iffD2), rtac exI, rtac conjI, + rtac ctxt (isNode_def RS iffD2), rtac ctxt exI, rtac ctxt conjI, CONVERSION (Conv.top_conv (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt), - if n = 1 then rtac refl else rtac (mk_sum_caseN n i), + if n = 1 then rtac ctxt refl else rtac ctxt (mk_sum_caseN n i), CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) => - EVERY' [rtac (set_map0 RS trans), rtac equalityI, rtac @{thm image_subsetI}, - rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, rtac set_Lev, - rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI}, rtac rv_Nil, - atac, rtac subsetI, + EVERY' [rtac ctxt (set_map0 RS trans), rtac ctxt equalityI, rtac ctxt @{thm image_subsetI}, + rtac ctxt CollectI, rtac ctxt @{thm SuccI}, rtac ctxt @{thm UN_I}, rtac ctxt UNIV_I, rtac ctxt set_Lev, + rtac ctxt (Lev_0 RS equalityD2 RS set_mp), rtac ctxt @{thm singletonI}, rtac ctxt rv_Nil, + atac, rtac ctxt subsetI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}], - rtac set_image_Lev, rtac (Lev_0 RS equalityD2 RS set_mp), - rtac @{thm singletonI}, dtac length_Lev', - etac @{thm set_mp[OF equalityD1[OF arg_cong[OF + rtac ctxt set_image_Lev, rtac ctxt (Lev_0 RS equalityD2 RS set_mp), + rtac ctxt @{thm singletonI}, dtac ctxt length_Lev', + etac ctxt @{thm set_mp[OF equalityD1[OF arg_cong[OF trans[OF length_append_singleton arg_cong[of _ _ Suc, OF list.size(3)]]]]]}, - rtac rv_Nil]) + rtac ctxt rv_Nil]) (drop m set_map0s ~~ (nth set_Levss (i - 1) ~~ nth set_image_Levss (i - 1)))]; fun mor_tac (i, (strT_def, ((Lev_Suc, (rv_Nil, rv_Cons)), ((map_comp_id, map_cong0), (length_Lev', (from_to_sbds, to_sbd_injs)))))) = - EVERY' [rtac ballI, rtac sym, rtac trans, rtac strT_def, + EVERY' [rtac ctxt ballI, rtac ctxt sym, rtac ctxt trans, rtac ctxt strT_def, CONVERSION (Conv.top_conv (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt), if n = 1 then K all_tac - else (rtac (sum_case_cong_weak RS trans) THEN' - rtac (mk_sum_caseN n i) THEN' rtac (mk_sum_caseN n i RS trans)), - rtac (map_comp_id RS trans), rtac (map_cong0 OF replicate m refl), + else (rtac ctxt (sum_case_cong_weak RS trans) THEN' + rtac ctxt (mk_sum_caseN n i) THEN' rtac ctxt (mk_sum_caseN n i RS trans)), + rtac ctxt (map_comp_id RS trans), rtac ctxt (map_cong0 OF replicate m refl), EVERY' (@{map 3} (fn i' => fn to_sbd_inj => fn from_to_sbd => - DETERM o EVERY' [rtac trans, rtac o_apply, rtac prod_injectI, rtac conjI, - rtac trans, rtac @{thm Shift_def}, - rtac equalityI, rtac subsetI, etac thin_rl, - REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm UN_E}], dtac length_Lev', dtac asm_rl, - etac thin_rl, dtac @{thm set_rev_mp[OF _ equalityD1]}, - rtac (@{thm length_Cons} RS arg_cong RS trans), rtac Lev_Suc, - CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' => + DETERM o EVERY' [rtac ctxt trans, rtac ctxt o_apply, rtac ctxt prod_injectI, rtac ctxt conjI, + rtac ctxt trans, rtac ctxt @{thm Shift_def}, + rtac ctxt equalityI, rtac ctxt subsetI, etac ctxt thin_rl, + REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm UN_E}], dtac ctxt length_Lev', dtac ctxt asm_rl, + etac ctxt thin_rl, dtac ctxt @{thm set_rev_mp[OF _ equalityD1]}, + rtac ctxt (@{thm length_Cons} RS arg_cong RS trans), rtac ctxt Lev_Suc, + CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) (fn i'' => EVERY' [REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], - dtac list_inject_iffD1, etac conjE, - if i' = i'' then EVERY' [dtac (mk_InN_inject n i'), - dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)), - atac, atac, hyp_subst_tac ctxt, etac @{thm UN_I[OF UNIV_I]}] - else etac (mk_InN_not_InM i' i'' RS notE)]) + dtac ctxt list_inject_iffD1, etac ctxt conjE, + if i' = i'' then EVERY' [dtac ctxt (mk_InN_inject n i'), + dtac ctxt (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)), + atac, atac, hyp_subst_tac ctxt, etac ctxt @{thm UN_I[OF UNIV_I]}] + else etac ctxt (mk_InN_not_InM i' i'' RS notE)]) (rev ks), - rtac @{thm UN_least}, rtac subsetI, rtac CollectI, rtac @{thm UN_I[OF UNIV_I]}, - rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i'), rtac CollectI, - REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, etac conjI, atac, - rtac trans, rtac @{thm shift_def}, rtac iffD2, rtac @{thm fun_eq_iff}, rtac allI, + rtac ctxt @{thm UN_least}, rtac ctxt subsetI, rtac ctxt CollectI, rtac ctxt @{thm UN_I[OF UNIV_I]}, + rtac ctxt (Lev_Suc RS equalityD2 RS set_mp), rtac ctxt (mk_UnIN n i'), rtac ctxt CollectI, + REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, etac ctxt conjI, atac, + rtac ctxt trans, rtac ctxt @{thm shift_def}, rtac ctxt iffD2, rtac ctxt @{thm fun_eq_iff}, rtac ctxt allI, CONVERSION (Conv.top_conv (K (Conv.try_conv (Conv.rewr_conv (rv_Cons RS eq_reflection)))) ctxt), if n = 1 then K all_tac - else rtac sum_case_cong_weak THEN' rtac (mk_sum_caseN n i' RS trans), - SELECT_GOAL (unfold_thms_tac ctxt [from_to_sbd]), rtac refl]) + else rtac ctxt sum_case_cong_weak THEN' rtac ctxt (mk_sum_caseN n i' RS trans), + SELECT_GOAL (unfold_thms_tac ctxt [from_to_sbd]), rtac ctxt refl]) ks to_sbd_injs from_to_sbds)]; in - (rtac mor_cong THEN' - EVERY' (map (fn thm => rtac (thm RS @{thm ext})) beh_defs) THEN' - rtac (mor_def RS iffD2) THEN' rtac conjI THEN' + (rtac ctxt mor_cong THEN' + EVERY' (map (fn thm => rtac ctxt (thm RS @{thm ext})) beh_defs) THEN' + rtac ctxt (mor_def RS iffD2) THEN' rtac ctxt conjI THEN' CONJ_WRAP' fbetw_tac (ks ~~ (carT_defs ~~ (isNode_defs ~~ (Lev_0s ~~ (rv_Nils ~~ ((length_Levs ~~ length_Lev's) ~~ (rv_lastss ~~ (set_map0ss ~~ (set_Levsss ~~ set_image_Levsss))))))))) THEN' @@ -646,162 +647,162 @@ (length_Lev's ~~ (from_to_sbdss ~~ to_sbd_injss))))))) 1 end; -fun mk_congruent_str_final_tac m lsbisE map_comp_id map_cong0 equiv_LSBISs = - EVERY' [rtac @{thm congruentI}, dtac lsbisE, - REPEAT_DETERM o eresolve0_tac [CollectE, conjE, bexE], rtac (o_apply RS trans), - etac (sym RS arg_cong RS trans), rtac (map_comp_id RS trans), - rtac (map_cong0 RS trans), REPEAT_DETERM_N m o rtac refl, +fun mk_congruent_str_final_tac ctxt m lsbisE map_comp_id map_cong0 equiv_LSBISs = + EVERY' [rtac ctxt @{thm congruentI}, dtac ctxt lsbisE, + REPEAT_DETERM o eresolve0_tac [CollectE, conjE, bexE], rtac ctxt (o_apply RS trans), + etac ctxt (sym RS arg_cong RS trans), rtac ctxt (map_comp_id RS trans), + rtac ctxt (map_cong0 RS trans), REPEAT_DETERM_N m o rtac ctxt refl, EVERY' (map (fn equiv_LSBIS => - EVERY' [rtac @{thm equiv_proj}, rtac equiv_LSBIS, etac set_mp, atac]) - equiv_LSBISs), rtac sym, rtac (o_apply RS trans), - etac (sym RS arg_cong RS trans), rtac map_comp_id] 1; + EVERY' [rtac ctxt @{thm equiv_proj}, rtac ctxt equiv_LSBIS, etac ctxt set_mp, atac]) + equiv_LSBISs), rtac ctxt sym, rtac ctxt (o_apply RS trans), + etac ctxt (sym RS arg_cong RS trans), rtac ctxt map_comp_id] 1; -fun mk_coalg_final_tac m coalg_def congruent_str_finals equiv_LSBISs set_map0ss coalgT_setss = - EVERY' [rtac (coalg_def RS iffD2), +fun mk_coalg_final_tac ctxt m coalg_def congruent_str_finals equiv_LSBISs set_map0ss coalgT_setss = + EVERY' [rtac ctxt (coalg_def RS iffD2), CONJ_WRAP' (fn ((set_map0s, coalgT_sets), (equiv_LSBIS, congruent_str_final)) => - EVERY' [rtac @{thm univ_preserves}, rtac equiv_LSBIS, rtac congruent_str_final, - rtac ballI, rtac @{thm ssubst_mem}, rtac o_apply, rtac CollectI, - REPEAT_DETERM_N m o EVERY' [rtac conjI, rtac subset_UNIV], + EVERY' [rtac ctxt @{thm univ_preserves}, rtac ctxt equiv_LSBIS, rtac ctxt congruent_str_final, + rtac ctxt ballI, rtac ctxt @{thm ssubst_mem}, rtac ctxt o_apply, rtac ctxt CollectI, + REPEAT_DETERM_N m o EVERY' [rtac ctxt conjI, rtac ctxt subset_UNIV], CONJ_WRAP' (fn (equiv_LSBIS, (set_map0, coalgT_set)) => - EVERY' [rtac (set_map0 RS ord_eq_le_trans), - rtac @{thm image_subsetI}, rtac iffD2, rtac @{thm proj_in_iff}, - rtac equiv_LSBIS, etac set_rev_mp, etac coalgT_set]) + EVERY' [rtac ctxt (set_map0 RS ord_eq_le_trans), + rtac ctxt @{thm image_subsetI}, rtac ctxt iffD2, rtac ctxt @{thm proj_in_iff}, + rtac ctxt equiv_LSBIS, etac ctxt set_rev_mp, etac ctxt coalgT_set]) (equiv_LSBISs ~~ (drop m set_map0s ~~ coalgT_sets))]) ((set_map0ss ~~ coalgT_setss) ~~ (equiv_LSBISs ~~ congruent_str_finals))] 1; -fun mk_mor_T_final_tac mor_def congruent_str_finals equiv_LSBISs = - EVERY' [rtac (mor_def RS iffD2), rtac conjI, +fun mk_mor_T_final_tac ctxt mor_def congruent_str_finals equiv_LSBISs = + EVERY' [rtac ctxt (mor_def RS iffD2), rtac ctxt conjI, CONJ_WRAP' (fn equiv_LSBIS => - EVERY' [rtac ballI, rtac iffD2, rtac @{thm proj_in_iff}, rtac equiv_LSBIS, atac]) + EVERY' [rtac ctxt ballI, rtac ctxt iffD2, rtac ctxt @{thm proj_in_iff}, rtac ctxt equiv_LSBIS, atac]) equiv_LSBISs, CONJ_WRAP' (fn (equiv_LSBIS, congruent_str_final) => - EVERY' [rtac ballI, rtac sym, rtac trans, rtac @{thm univ_commute}, rtac equiv_LSBIS, - rtac congruent_str_final, atac, rtac o_apply]) + EVERY' [rtac ctxt ballI, rtac ctxt sym, rtac ctxt trans, rtac ctxt @{thm univ_commute}, rtac ctxt equiv_LSBIS, + rtac ctxt congruent_str_final, atac, rtac ctxt o_apply]) (equiv_LSBISs ~~ congruent_str_finals)] 1; fun mk_mor_Rep_tac ctxt defs Reps Abs_inverses coalg_final_setss map_comp_ids map_cong0Ls = unfold_thms_tac ctxt defs THEN - EVERY' [rtac conjI, - CONJ_WRAP' (fn thm => rtac ballI THEN' rtac thm) Reps, + EVERY' [rtac ctxt conjI, + CONJ_WRAP' (fn thm => rtac ctxt ballI THEN' rtac ctxt thm) Reps, CONJ_WRAP' (fn (Rep, ((map_comp_id, map_cong0L), coalg_final_sets)) => - EVERY' [rtac ballI, rtac (map_comp_id RS trans), rtac map_cong0L, + EVERY' [rtac ctxt ballI, rtac ctxt (map_comp_id RS trans), rtac ctxt map_cong0L, EVERY' (map2 (fn Abs_inverse => fn coalg_final_set => - EVERY' [rtac ballI, rtac (o_apply RS trans), rtac Abs_inverse, - etac set_rev_mp, rtac coalg_final_set, rtac Rep]) + EVERY' [rtac ctxt ballI, rtac ctxt (o_apply RS trans), rtac ctxt Abs_inverse, + etac ctxt set_rev_mp, rtac ctxt coalg_final_set, rtac ctxt Rep]) Abs_inverses coalg_final_sets)]) (Reps ~~ ((map_comp_ids ~~ map_cong0Ls) ~~ coalg_final_setss))] 1; fun mk_mor_Abs_tac ctxt defs Abs_inverses = unfold_thms_tac ctxt defs THEN - EVERY' [rtac conjI, - CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) Abs_inverses, - CONJ_WRAP' (fn thm => rtac ballI THEN' etac (thm RS arg_cong RS sym)) Abs_inverses] 1; + EVERY' [rtac ctxt conjI, + CONJ_WRAP' (K (rtac ctxt ballI THEN' rtac ctxt UNIV_I)) Abs_inverses, + CONJ_WRAP' (fn thm => rtac ctxt ballI THEN' etac ctxt (thm RS arg_cong RS sym)) Abs_inverses] 1; -fun mk_mor_unfold_tac m mor_UNIV dtor_defs unfold_defs Abs_inverses morEs map_comp_ids map_cong0s = - EVERY' [rtac iffD2, rtac mor_UNIV, +fun mk_mor_unfold_tac ctxt m mor_UNIV dtor_defs unfold_defs Abs_inverses morEs map_comp_ids map_cong0s = + EVERY' [rtac ctxt iffD2, rtac ctxt mor_UNIV, CONJ_WRAP' (fn ((Abs_inverse, morE), ((dtor_def, unfold_def), (map_comp_id, map_cong0))) => - EVERY' [rtac @{thm ext}, rtac (o_apply RS trans RS sym), rtac (dtor_def RS trans), - rtac (unfold_def RS arg_cong RS trans), rtac (Abs_inverse RS arg_cong RS trans), - rtac (morE RS arg_cong RS trans), rtac (map_comp_id RS trans), - rtac (o_apply RS trans RS sym), rtac map_cong0, - REPEAT_DETERM_N m o rtac refl, - EVERY' (map (fn thm => rtac (thm RS trans) THEN' rtac (o_apply RS sym)) unfold_defs)]) + EVERY' [rtac ctxt @{thm ext}, rtac ctxt (o_apply RS trans RS sym), rtac ctxt (dtor_def RS trans), + rtac ctxt (unfold_def RS arg_cong RS trans), rtac ctxt (Abs_inverse RS arg_cong RS trans), + rtac ctxt (morE RS arg_cong RS trans), rtac ctxt (map_comp_id RS trans), + rtac ctxt (o_apply RS trans RS sym), rtac ctxt map_cong0, + REPEAT_DETERM_N m o rtac ctxt refl, + EVERY' (map (fn thm => rtac ctxt (thm RS trans) THEN' rtac ctxt (o_apply RS sym)) unfold_defs)]) ((Abs_inverses ~~ morEs) ~~ ((dtor_defs ~~ unfold_defs) ~~ (map_comp_ids ~~ map_cong0s)))] 1; -fun mk_raw_coind_tac bis_def bis_cong bis_O bis_converse bis_Gr tcoalg coalgT mor_T_final +fun mk_raw_coind_tac ctxt bis_def bis_cong bis_O bis_converse bis_Gr tcoalg coalgT mor_T_final sbis_lsbis lsbis_incls incl_lsbiss equiv_LSBISs mor_Rep Rep_injects = let val n = length Rep_injects; in - EVERY' [rtac rev_mp, ftac (bis_def RS iffD1), - REPEAT_DETERM o etac conjE, rtac bis_cong, rtac bis_O, rtac bis_converse, - rtac bis_Gr, rtac tcoalg, rtac mor_Rep, rtac bis_O, atac, rtac bis_Gr, rtac tcoalg, - rtac mor_Rep, REPEAT_DETERM_N n o etac @{thm relImage_Gr}, - rtac impI, rtac rev_mp, rtac bis_cong, rtac bis_O, rtac bis_Gr, rtac coalgT, - rtac mor_T_final, rtac bis_O, rtac sbis_lsbis, rtac bis_converse, rtac bis_Gr, rtac coalgT, - rtac mor_T_final, EVERY' (map (fn thm => rtac (thm RS @{thm relInvImage_Gr})) lsbis_incls), - rtac impI, + EVERY' [rtac ctxt rev_mp, ftac ctxt (bis_def RS iffD1), + REPEAT_DETERM o etac ctxt conjE, rtac ctxt bis_cong, rtac ctxt bis_O, rtac ctxt bis_converse, + rtac ctxt bis_Gr, rtac ctxt tcoalg, rtac ctxt mor_Rep, rtac ctxt bis_O, atac, rtac ctxt bis_Gr, rtac ctxt tcoalg, + rtac ctxt mor_Rep, REPEAT_DETERM_N n o etac ctxt @{thm relImage_Gr}, + rtac ctxt impI, rtac ctxt rev_mp, rtac ctxt bis_cong, rtac ctxt bis_O, rtac ctxt bis_Gr, rtac ctxt coalgT, + rtac ctxt mor_T_final, rtac ctxt bis_O, rtac ctxt sbis_lsbis, rtac ctxt bis_converse, rtac ctxt bis_Gr, rtac ctxt coalgT, + rtac ctxt mor_T_final, EVERY' (map (fn thm => rtac ctxt (thm RS @{thm relInvImage_Gr})) lsbis_incls), + rtac ctxt impI, CONJ_WRAP' (fn (Rep_inject, (equiv_LSBIS , (incl_lsbis, lsbis_incl))) => - EVERY' [rtac subset_trans, rtac @{thm relInvImage_UNIV_relImage}, rtac subset_trans, - rtac @{thm relInvImage_mono}, rtac subset_trans, etac incl_lsbis, - rtac ord_eq_le_trans, rtac @{thm sym[OF relImage_relInvImage]}, - rtac @{thm xt1(3)}, rtac @{thm Sigma_cong}, - rtac @{thm proj_image}, rtac @{thm proj_image}, rtac lsbis_incl, - rtac subset_trans, rtac @{thm relImage_mono}, rtac incl_lsbis, atac, - rtac @{thm relImage_proj}, rtac equiv_LSBIS, rtac @{thm relInvImage_Id_on}, - rtac Rep_inject]) + EVERY' [rtac ctxt subset_trans, rtac ctxt @{thm relInvImage_UNIV_relImage}, rtac ctxt subset_trans, + rtac ctxt @{thm relInvImage_mono}, rtac ctxt subset_trans, etac ctxt incl_lsbis, + rtac ctxt ord_eq_le_trans, rtac ctxt @{thm sym[OF relImage_relInvImage]}, + rtac ctxt @{thm xt1(3)}, rtac ctxt @{thm Sigma_cong}, + rtac ctxt @{thm proj_image}, rtac ctxt @{thm proj_image}, rtac ctxt lsbis_incl, + rtac ctxt subset_trans, rtac ctxt @{thm relImage_mono}, rtac ctxt incl_lsbis, atac, + rtac ctxt @{thm relImage_proj}, rtac ctxt equiv_LSBIS, rtac ctxt @{thm relInvImage_Id_on}, + rtac ctxt Rep_inject]) (Rep_injects ~~ (equiv_LSBISs ~~ (incl_lsbiss ~~ lsbis_incls)))] 1 end; -fun mk_unfold_unique_mor_tac raw_coinds bis mor unfold_defs = +fun mk_unfold_unique_mor_tac ctxt raw_coinds bis mor unfold_defs = CONJ_WRAP' (fn (raw_coind, unfold_def) => - EVERY' [rtac @{thm ext}, etac (bis RS raw_coind RS set_mp RS @{thm IdD}), rtac mor, - rtac @{thm image2_eqI}, rtac refl, rtac (unfold_def RS arg_cong RS trans), - rtac (o_apply RS sym), rtac UNIV_I]) (raw_coinds ~~ unfold_defs) 1; + EVERY' [rtac ctxt @{thm ext}, etac ctxt (bis RS raw_coind RS set_mp RS @{thm IdD}), rtac ctxt mor, + rtac ctxt @{thm image2_eqI}, rtac ctxt refl, rtac ctxt (unfold_def RS arg_cong RS trans), + rtac ctxt (o_apply RS sym), rtac ctxt UNIV_I]) (raw_coinds ~~ unfold_defs) 1; fun mk_dtor_o_ctor_tac ctxt ctor_def unfold map_comp_id map_cong0L unfold_o_dtors = - unfold_thms_tac ctxt [ctor_def] THEN EVERY' [rtac @{thm ext}, rtac trans, rtac o_apply, - rtac trans, rtac unfold, rtac trans, rtac map_comp_id, rtac trans, rtac map_cong0L, + unfold_thms_tac ctxt [ctor_def] THEN EVERY' [rtac ctxt @{thm ext}, rtac ctxt trans, rtac ctxt o_apply, + rtac ctxt trans, rtac ctxt unfold, rtac ctxt trans, rtac ctxt map_comp_id, rtac ctxt trans, rtac ctxt map_cong0L, EVERY' (map (fn thm => - rtac ballI THEN' rtac (trans OF [thm RS fun_cong, id_apply])) unfold_o_dtors), - rtac sym, rtac id_apply] 1; + rtac ctxt ballI THEN' rtac ctxt (trans OF [thm RS fun_cong, id_apply])) unfold_o_dtors), + rtac ctxt sym, rtac ctxt id_apply] 1; fun mk_corec_tac ctxt m corec_defs unfold map_cong0 corec_Inls = - unfold_thms_tac ctxt corec_defs THEN EVERY' [rtac trans, rtac (o_apply RS arg_cong), - rtac trans, rtac unfold, fo_rtac (@{thm sum.case(2)} RS arg_cong RS trans) ctxt, rtac map_cong0, - REPEAT_DETERM_N m o rtac refl, - EVERY' (map (fn thm => rtac @{thm case_sum_expand_Inr} THEN' rtac thm) corec_Inls)] 1; + unfold_thms_tac ctxt corec_defs THEN EVERY' [rtac ctxt trans, rtac ctxt (o_apply RS arg_cong), + rtac ctxt trans, rtac ctxt unfold, fo_rtac ctxt (@{thm sum.case(2)} RS arg_cong RS trans), + rtac ctxt map_cong0, REPEAT_DETERM_N m o rtac ctxt refl, + EVERY' (map (fn thm => rtac ctxt @{thm case_sum_expand_Inr} THEN' rtac ctxt thm) corec_Inls)] 1; fun mk_corec_unique_mor_tac ctxt corec_defs corec_Inls unfold_unique_mor = unfold_thms_tac ctxt (corec_defs @ map (fn thm => thm RS @{thm case_sum_expand_Inr'}) corec_Inls) THEN - etac unfold_unique_mor 1; + etac ctxt unfold_unique_mor 1; -fun mk_dtor_coinduct_tac m raw_coind bis_rel rel_congs = - EVERY' [rtac rev_mp, rtac raw_coind, rtac iffD2, rtac bis_rel, rtac conjI, - CONJ_WRAP' (K (rtac @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]})) +fun mk_dtor_coinduct_tac ctxt m raw_coind bis_rel rel_congs = + EVERY' [rtac ctxt rev_mp, rtac ctxt raw_coind, rtac ctxt iffD2, rtac ctxt bis_rel, rtac ctxt conjI, + CONJ_WRAP' (K (rtac ctxt @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]})) rel_congs, - CONJ_WRAP' (fn rel_cong => EVERY' [rtac allI, rtac allI, rtac impI, - REPEAT_DETERM o etac allE, rtac (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}), - REPEAT_DETERM_N m o rtac refl, - REPEAT_DETERM_N (length rel_congs) o rtac @{thm in_rel_Collect_split_eq[symmetric]}, - etac mp, etac CollectE, etac @{thm splitD}]) + CONJ_WRAP' (fn rel_cong => EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, + REPEAT_DETERM o etac ctxt allE, rtac ctxt (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}), + REPEAT_DETERM_N m o rtac ctxt refl, + REPEAT_DETERM_N (length rel_congs) o rtac ctxt @{thm in_rel_Collect_split_eq[symmetric]}, + etac ctxt mp, etac ctxt CollectE, etac ctxt @{thm splitD}]) rel_congs, - rtac impI, REPEAT_DETERM o etac conjE, - CONJ_WRAP' (K (EVERY' [rtac impI, rtac @{thm IdD}, etac set_mp, - rtac CollectI, etac @{thm case_prodI}])) rel_congs] 1; + rtac ctxt impI, REPEAT_DETERM o etac ctxt conjE, + CONJ_WRAP' (K (EVERY' [rtac ctxt impI, rtac ctxt @{thm IdD}, etac ctxt set_mp, + rtac ctxt CollectI, etac ctxt @{thm case_prodI}])) rel_congs] 1; -fun mk_map_tac m n map_arg_cong unfold map_comp map_cong0 = - EVERY' [rtac @{thm ext}, rtac (o_apply RS trans RS sym), rtac (o_apply RS trans RS sym), - rtac (unfold RS trans), rtac (Thm.permute_prems 0 1 (map_comp RS @{thm box_equals})), - rtac map_cong0, - REPEAT_DETERM_N m o rtac (@{thm id_comp} RS fun_cong), - REPEAT_DETERM_N n o rtac (@{thm comp_id} RS fun_cong), - rtac map_arg_cong, rtac (o_apply RS sym)] 1; +fun mk_map_tac ctxt m n map_arg_cong unfold map_comp map_cong0 = + EVERY' [rtac ctxt @{thm ext}, rtac ctxt (o_apply RS trans RS sym), rtac ctxt (o_apply RS trans RS sym), + rtac ctxt (unfold RS trans), rtac ctxt (Thm.permute_prems 0 1 (map_comp RS @{thm box_equals})), + rtac ctxt map_cong0, + REPEAT_DETERM_N m o rtac ctxt (@{thm id_comp} RS fun_cong), + REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong), + rtac ctxt map_arg_cong, rtac ctxt (o_apply RS sym)] 1; -fun mk_set_le_tac n Jset_minimal set_Jsets set_Jset_Jsetss = - EVERY' [rtac Jset_minimal, - REPEAT_DETERM_N n o rtac @{thm Un_upper1}, +fun mk_set_le_tac ctxt n Jset_minimal set_Jsets set_Jset_Jsetss = + EVERY' [rtac ctxt Jset_minimal, + REPEAT_DETERM_N n o rtac ctxt @{thm Un_upper1}, REPEAT_DETERM_N n o EVERY' (@{map 3} (fn i => fn set_Jset => fn set_Jset_Jsets => - EVERY' [rtac subsetI, rtac @{thm UnI2}, rtac (mk_UnIN n i), etac @{thm UN_I}, - etac UnE, etac set_Jset, REPEAT_DETERM_N (n - 1) o etac UnE, - EVERY' (map (fn thm => EVERY' [etac @{thm UN_E}, etac thm, atac]) set_Jset_Jsets)]) + EVERY' [rtac ctxt subsetI, rtac ctxt @{thm UnI2}, rtac ctxt (mk_UnIN n i), etac ctxt @{thm UN_I}, + etac ctxt UnE, etac ctxt set_Jset, REPEAT_DETERM_N (n - 1) o etac ctxt UnE, + EVERY' (map (fn thm => EVERY' [etac ctxt @{thm UN_E}, etac ctxt thm, atac]) set_Jset_Jsets)]) (1 upto n) set_Jsets set_Jset_Jsetss)] 1; -fun mk_set_ge_tac n set_incl_Jset set_Jset_incl_Jsets = - EVERY' [rtac @{thm Un_least}, rtac set_incl_Jset, - REPEAT_DETERM_N (n - 1) o rtac @{thm Un_least}, - EVERY' (map (fn thm => rtac @{thm UN_least} THEN' etac thm) set_Jset_incl_Jsets)] 1; +fun mk_set_ge_tac ctxt n set_incl_Jset set_Jset_incl_Jsets = + EVERY' [rtac ctxt @{thm Un_least}, rtac ctxt set_incl_Jset, + REPEAT_DETERM_N (n - 1) o rtac ctxt @{thm Un_least}, + EVERY' (map (fn thm => rtac ctxt @{thm UN_least} THEN' etac ctxt thm) set_Jset_incl_Jsets)] 1; -fun mk_map_id0_tac maps unfold_unique unfold_dtor = - EVERY' [rtac (unfold_unique RS trans), EVERY' (map (rtac o mk_sym) maps), - rtac unfold_dtor] 1; +fun mk_map_id0_tac ctxt maps unfold_unique unfold_dtor = + EVERY' [rtac ctxt (unfold_unique RS trans), EVERY' (map (rtac ctxt o mk_sym) maps), + rtac ctxt unfold_dtor] 1; -fun mk_map_comp0_tac maps map_comp0s map_unique = - EVERY' [rtac map_unique, +fun mk_map_comp0_tac ctxt maps map_comp0s map_unique = + EVERY' [rtac ctxt map_unique, EVERY' (map2 (fn map_thm => fn map_comp0 => - EVERY' (map rtac + EVERY' (map (rtac ctxt) [@{thm comp_assoc[symmetric]} RS trans, @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_thm, refl] RS trans, @{thm comp_assoc[symmetric]} RS sym RS trans, map_thm RS arg_cong RS trans, @@ -815,136 +816,136 @@ val n = length map_comps; val ks = 1 upto n; in - EVERY' ([rtac rev_mp, coinduct_tac] @ + EVERY' ([rtac ctxt rev_mp, coinduct_tac] @ maps (fn ((((((map_comp_trans, dtor_maps_trans), map_cong0), set_map0s), set_Jsets), set_Jset_Jsetss), in_rel) => [REPEAT_DETERM o resolve_tac ctxt [allI, impI, in_rel RS iffD2], - REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], hyp_subst_tac ctxt, rtac exI, - rtac (Drule.rotate_prems 1 conjI), - rtac conjI, rtac map_comp_trans, rtac sym, rtac dtor_maps_trans, rtac map_cong0, - REPEAT_DETERM_N m o (rtac o_apply_trans_sym THEN' rtac @{thm fst_conv}), - REPEAT_DETERM_N n o rtac fst_convol_fun_cong_sym, - rtac map_comp_trans, rtac sym, rtac dtor_maps_trans, rtac map_cong0, + REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], hyp_subst_tac ctxt, rtac ctxt exI, + rtac ctxt (Drule.rotate_prems 1 conjI), + rtac ctxt conjI, rtac ctxt map_comp_trans, rtac ctxt sym, rtac ctxt dtor_maps_trans, rtac ctxt map_cong0, + REPEAT_DETERM_N m o (rtac ctxt o_apply_trans_sym THEN' rtac ctxt @{thm fst_conv}), + REPEAT_DETERM_N n o rtac ctxt fst_convol_fun_cong_sym, + rtac ctxt map_comp_trans, rtac ctxt sym, rtac ctxt dtor_maps_trans, rtac ctxt map_cong0, EVERY' (maps (fn set_Jset => - [rtac o_apply_trans_sym, rtac (@{thm snd_conv} RS trans), etac CollectE, - REPEAT_DETERM o etac conjE, etac bspec, etac set_Jset]) set_Jsets), - REPEAT_DETERM_N n o rtac snd_convol_fun_cong_sym, - rtac CollectI, - EVERY' (map (fn set_map0 => EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_map0, - rtac @{thm image_subsetI}, rtac CollectI, rtac @{thm case_prodI}, rtac refl]) + [rtac ctxt o_apply_trans_sym, rtac ctxt (@{thm snd_conv} RS trans), etac ctxt CollectE, + REPEAT_DETERM o etac ctxt conjE, etac ctxt bspec, etac ctxt set_Jset]) set_Jsets), + REPEAT_DETERM_N n o rtac ctxt snd_convol_fun_cong_sym, + rtac ctxt CollectI, + EVERY' (map (fn set_map0 => EVERY' [rtac ctxt conjI, rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, + rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, rtac ctxt refl]) (take m set_map0s)), CONJ_WRAP' (fn (set_map0, set_Jset_Jsets) => - EVERY' [rtac ord_eq_le_trans, rtac set_map0, - rtac @{thm image_subsetI}, rtac CollectI, rtac @{thm case_prodI}, rtac exI, rtac conjI, - rtac CollectI, etac CollectE, - REPEAT_DETERM o etac conjE, + EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, + rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, rtac ctxt exI, rtac ctxt conjI, + rtac ctxt CollectI, etac ctxt CollectE, + REPEAT_DETERM o etac ctxt conjE, CONJ_WRAP' (fn set_Jset_Jset => - EVERY' [rtac ballI, etac bspec, etac set_Jset_Jset, atac]) set_Jset_Jsets, - rtac (conjI OF [refl, refl])]) + EVERY' [rtac ctxt ballI, etac ctxt bspec, etac ctxt set_Jset_Jset, atac]) set_Jset_Jsets, + rtac ctxt (conjI OF [refl, refl])]) (drop m set_map0s ~~ set_Jset_Jsetss)]) (map (fn th => th RS trans) map_comps ~~ map (fn th => th RS trans) dtor_maps ~~ map_cong0s ~~ set_map0ss ~~ set_Jsetss ~~ set_Jset_Jsetsss ~~ in_rels) @ - [rtac impI, + [rtac ctxt impI, CONJ_WRAP' (fn k => - EVERY' [rtac impI, dtac (mk_conjunctN n k), etac mp, rtac exI, rtac conjI, etac CollectI, - rtac conjI, rtac refl, rtac refl]) ks]) 1 + EVERY' [rtac ctxt impI, dtac ctxt (mk_conjunctN n k), etac ctxt mp, rtac ctxt exI, rtac ctxt conjI, etac ctxt CollectI, + rtac ctxt conjI, rtac ctxt refl, rtac ctxt refl]) ks]) 1 end fun mk_dtor_map_unique_tac ctxt unfold_unique sym_map_comps = - rtac unfold_unique 1 THEN + rtac ctxt unfold_unique 1 THEN unfold_thms_tac ctxt (sym_map_comps @ @{thms comp_assoc[symmetric] id_comp comp_id}) THEN - ALLGOALS (etac sym); + ALLGOALS (etac ctxt sym); fun mk_col_natural_tac ctxt cts rec_0s rec_Sucs dtor_maps set_map0ss = let val n = length dtor_maps; in - EVERY' [rtac (Drule.instantiate' [] cts nat_induct), - REPEAT_DETERM o rtac allI, SELECT_GOAL (unfold_thms_tac ctxt rec_0s), - CONJ_WRAP' (K (rtac @{thm image_empty})) rec_0s, - REPEAT_DETERM o rtac allI, + EVERY' [rtac ctxt (Drule.instantiate' [] cts nat_induct), + REPEAT_DETERM o rtac ctxt allI, SELECT_GOAL (unfold_thms_tac ctxt rec_0s), + CONJ_WRAP' (K (rtac ctxt @{thm image_empty})) rec_0s, + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn (rec_Suc, (dtor_map, set_nats)) => EVERY' [SELECT_GOAL (unfold_thms_tac ctxt (rec_Suc :: dtor_map :: set_nats @ @{thms image_Un image_UN UN_simps(10)})), - rtac Un_cong, rtac refl, - CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 Un_cong)) - (fn i => EVERY' [rtac @{thm SUP_cong[OF refl]}, - REPEAT_DETERM o etac allE, etac (mk_conjunctN n i)]) (n downto 1)]) + rtac ctxt Un_cong, rtac ctxt refl, + CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 Un_cong)) + (fn i => EVERY' [rtac ctxt @{thm SUP_cong[OF refl]}, + REPEAT_DETERM o etac ctxt allE, etac ctxt (mk_conjunctN n i)]) (n downto 1)]) (rec_Sucs ~~ (dtor_maps ~~ set_map0ss))] 1 end; -fun mk_set_map0_tac col_natural = - EVERY' (map rtac [@{thm ext}, o_apply RS trans, sym, o_apply RS trans, @{thm image_UN} RS trans, - refl RS @{thm SUP_cong}, col_natural]) 1; +fun mk_set_map0_tac ctxt col_natural = + EVERY' (map (rtac ctxt) [@{thm ext}, o_apply RS trans, sym, o_apply RS trans, + @{thm image_UN} RS trans, refl RS @{thm SUP_cong}, col_natural]) 1; -fun mk_col_bd_tac m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss = +fun mk_col_bd_tac ctxt m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss = let val n = length rec_0s; in - EVERY' [rtac (Drule.instantiate' [] cts nat_induct), - REPEAT_DETERM o rtac allI, - CONJ_WRAP' (fn rec_0 => EVERY' (map rtac [ordIso_ordLeq_trans, + EVERY' [rtac ctxt (Drule.instantiate' [] cts nat_induct), + REPEAT_DETERM o rtac ctxt allI, + CONJ_WRAP' (fn rec_0 => EVERY' (map (rtac ctxt) [ordIso_ordLeq_trans, @{thm card_of_ordIso_subst}, rec_0, @{thm Card_order_empty}, sbd_Card_order])) rec_0s, - REPEAT_DETERM o rtac allI, + REPEAT_DETERM o rtac ctxt allI, CONJ_WRAP' (fn (rec_Suc, set_sbds) => EVERY' - [rtac ordIso_ordLeq_trans, rtac @{thm card_of_ordIso_subst}, rtac rec_Suc, - rtac (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_sbds (j - 1)), - REPEAT_DETERM_N (n - 1) o rtac (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), - EVERY' (map2 (fn i => fn set_sbd => EVERY' [rtac @{thm UNION_Cinfinite_bound}, - rtac set_sbd, rtac ballI, REPEAT_DETERM o etac allE, - etac (mk_conjunctN n i), rtac sbd_Cinfinite]) (1 upto n) (drop m set_sbds))]) + [rtac ctxt ordIso_ordLeq_trans, rtac ctxt @{thm card_of_ordIso_subst}, rtac ctxt rec_Suc, + rtac ctxt (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac ctxt (nth set_sbds (j - 1)), + REPEAT_DETERM_N (n - 1) o rtac ctxt (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), + EVERY' (map2 (fn i => fn set_sbd => EVERY' [rtac ctxt @{thm UNION_Cinfinite_bound}, + rtac ctxt set_sbd, rtac ctxt ballI, REPEAT_DETERM o etac ctxt allE, + etac ctxt (mk_conjunctN n i), rtac ctxt sbd_Cinfinite]) (1 upto n) (drop m set_sbds))]) (rec_Sucs ~~ set_sbdss)] 1 end; -fun mk_set_bd_tac sbd_Cinfinite col_bd = - EVERY' (map rtac [@{thm UNION_Cinfinite_bound}, ordIso_ordLeq_trans, @{thm card_of_nat}, +fun mk_set_bd_tac ctxt sbd_Cinfinite col_bd = + EVERY' (map (rtac ctxt) [@{thm UNION_Cinfinite_bound}, ordIso_ordLeq_trans, @{thm card_of_nat}, @{thm natLeq_ordLeq_cinfinite}, sbd_Cinfinite, ballI, col_bd, sbd_Cinfinite]) 1; -fun mk_le_rel_OO_tac coinduct rel_Jrels le_rel_OOs = - EVERY' (rtac coinduct :: map2 (fn rel_Jrel => fn le_rel_OO => +fun mk_le_rel_OO_tac ctxt coinduct rel_Jrels le_rel_OOs = + EVERY' (rtac ctxt coinduct :: map2 (fn rel_Jrel => fn le_rel_OO => let val Jrel_imp_rel = rel_Jrel RS iffD1; in - EVERY' [rtac (le_rel_OO RS @{thm predicate2D}), etac @{thm relcomppE}, - rtac @{thm relcomppI}, etac Jrel_imp_rel, etac Jrel_imp_rel] + EVERY' [rtac ctxt (le_rel_OO RS @{thm predicate2D}), etac ctxt @{thm relcomppE}, + rtac ctxt @{thm relcomppI}, etac ctxt Jrel_imp_rel, etac ctxt Jrel_imp_rel] end) rel_Jrels le_rel_OOs) 1; fun mk_wit_tac ctxt n dtor_ctors dtor_set wit coind_wits = - ALLGOALS (TRY o (eresolve_tac ctxt coind_wits THEN' rtac refl)) THEN + ALLGOALS (TRY o (eresolve_tac ctxt coind_wits THEN' rtac ctxt refl)) THEN REPEAT_DETERM (atac 1 ORELSE - EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac ctxt dtor_set, + EVERY' [dtac ctxt set_rev_mp, rtac ctxt equalityD1, resolve_tac ctxt dtor_set, K (unfold_thms_tac ctxt dtor_ctors), - REPEAT_DETERM_N n o etac UnE, + REPEAT_DETERM_N n o etac ctxt UnE, REPEAT_DETERM o - (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN' + (TRY o REPEAT_DETERM o etac ctxt UnE THEN' TRY o etac ctxt @{thm UN_E} THEN' (eresolve_tac ctxt wit ORELSE' (dresolve_tac ctxt wit THEN' - (etac FalseE ORELSE' - EVERY' [hyp_subst_tac ctxt, dtac set_rev_mp, rtac equalityD1, resolve_tac ctxt dtor_set, - K (unfold_thms_tac ctxt dtor_ctors), REPEAT_DETERM_N n o etac UnE]))))] 1); + (etac ctxt FalseE ORELSE' + EVERY' [hyp_subst_tac ctxt, dtac ctxt set_rev_mp, rtac ctxt equalityD1, resolve_tac ctxt dtor_set, + K (unfold_thms_tac ctxt dtor_ctors), REPEAT_DETERM_N n o etac ctxt UnE]))))] 1); fun mk_coind_wit_tac ctxt induct unfolds set_nats wits = - rtac induct 1 THEN ALLGOALS (TRY o rtac impI THEN' TRY o hyp_subst_tac ctxt) THEN + rtac ctxt induct 1 THEN ALLGOALS (TRY o rtac ctxt impI THEN' TRY o hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt (unfolds @ set_nats @ @{thms image_id id_apply}) THEN - ALLGOALS (REPEAT_DETERM o etac imageE THEN' TRY o hyp_subst_tac ctxt) THEN + ALLGOALS (REPEAT_DETERM o etac ctxt imageE THEN' TRY o hyp_subst_tac ctxt) THEN ALLGOALS (TRY o - FIRST' [rtac TrueI, rtac refl, etac (refl RSN (2, mp)), dresolve_tac ctxt wits THEN' etac FalseE]); + FIRST' [rtac ctxt TrueI, rtac ctxt refl, etac ctxt (refl RSN (2, mp)), dresolve_tac ctxt wits THEN' etac ctxt FalseE]); fun mk_dtor_corec_transfer_tac ctxt n m dtor_corec_defs dtor_unfold_transfer pre_T_map_transfers dtor_rels = CONJ_WRAP (fn (dtor_corec_def, dtor_unfold_transfer) => - REPEAT_DETERM (HEADGOAL (rtac rel_funI)) THEN + REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN unfold_thms_tac ctxt [dtor_corec_def, o_apply] THEN - HEADGOAL (rtac (mk_rel_funDN (n + 1) dtor_unfold_transfer) THEN' + HEADGOAL (rtac ctxt (mk_rel_funDN (n + 1) dtor_unfold_transfer) THEN' EVERY' (map2 (fn pre_T_map_transfer => fn dtor_rel => - etac (mk_rel_funDN_rotated 2 @{thm case_sum_transfer}) THEN' - rtac (mk_rel_funDN 2 @{thm comp_transfer}) THEN' - rtac (mk_rel_funDN (m + n) pre_T_map_transfer) THEN' - REPEAT_DETERM_N m o rtac @{thm id_transfer} THEN' - REPEAT_DETERM_N n o rtac @{thm Inl_transfer} THEN' - rtac rel_funI THEN' - etac (dtor_rel RS iffD1)) pre_T_map_transfers dtor_rels) THEN' - etac (mk_rel_funDN 1 @{thm Inr_transfer}))) + etac ctxt (mk_rel_funDN_rotated 2 @{thm case_sum_transfer}) THEN' + rtac ctxt (mk_rel_funDN 2 @{thm comp_transfer}) THEN' + rtac ctxt (mk_rel_funDN (m + n) pre_T_map_transfer) THEN' + REPEAT_DETERM_N m o rtac ctxt @{thm id_transfer} THEN' + REPEAT_DETERM_N n o rtac ctxt @{thm Inl_transfer} THEN' + rtac ctxt rel_funI THEN' + etac ctxt (dtor_rel RS iffD1)) pre_T_map_transfers dtor_rels) THEN' + etac ctxt (mk_rel_funDN 1 @{thm Inr_transfer}))) (dtor_corec_defs ~~ dtor_unfold_transfer); fun mk_dtor_rel_tac ctxt in_Jrels i in_rel map_comp0 map_cong0 dtor_map dtor_sets dtor_inject @@ -955,59 +956,59 @@ val (passive_set_map0s, active_set_map0s) = chop m set_map0s; val in_Jrel = nth in_Jrels (i - 1); val if_tac = - EVERY' [dtac (in_Jrel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], - rtac (in_rel RS iffD2), rtac exI, rtac conjI, rtac CollectI, + EVERY' [dtac ctxt (in_Jrel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], + rtac ctxt (in_rel RS iffD2), rtac ctxt exI, rtac ctxt conjI, rtac ctxt CollectI, EVERY' (map2 (fn set_map0 => fn set_incl => - EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_map0, - rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply, - etac (set_incl RS @{thm subset_trans})]) + EVERY' [rtac ctxt conjI, rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, + rtac ctxt ord_eq_le_trans, rtac ctxt trans_fun_cong_image_id_id_apply, + etac ctxt (set_incl RS @{thm subset_trans})]) passive_set_map0s dtor_set_incls), CONJ_WRAP' (fn (in_Jrel, (set_map0, dtor_set_set_incls)) => - EVERY' [rtac ord_eq_le_trans, rtac set_map0, rtac @{thm image_subsetI}, rtac CollectI, - rtac @{thm case_prodI}, rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, rtac CollectI, - CONJ_WRAP' (fn thm => etac (thm RS @{thm subset_trans}) THEN' atac) dtor_set_set_incls, - rtac conjI, rtac refl, rtac refl]) + EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, + rtac ctxt @{thm case_prodI}, rtac ctxt (in_Jrel RS iffD2), rtac ctxt exI, rtac ctxt conjI, rtac ctxt CollectI, + CONJ_WRAP' (fn thm => etac ctxt (thm RS @{thm subset_trans}) THEN' atac) dtor_set_set_incls, + rtac ctxt conjI, rtac ctxt refl, rtac ctxt refl]) (in_Jrels ~~ (active_set_map0s ~~ dtor_set_set_inclss)), CONJ_WRAP' (fn conv => - EVERY' [rtac trans, rtac map_comp0, rtac trans, rtac map_cong0, - REPEAT_DETERM_N m o rtac @{thm fun_cong[OF comp_id]}, - REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]), - rtac trans, rtac sym, rtac dtor_map, rtac (dtor_inject RS iffD2), atac]) + EVERY' [rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0, + REPEAT_DETERM_N m o rtac ctxt @{thm fun_cong[OF comp_id]}, + REPEAT_DETERM_N n o EVERY' (map (rtac ctxt) [trans, o_apply, conv]), + rtac ctxt trans, rtac ctxt sym, rtac ctxt dtor_map, rtac ctxt (dtor_inject RS iffD2), atac]) @{thms fst_conv snd_conv}]; val only_if_tac = - EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], - rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, rtac CollectI, + EVERY' [dtac ctxt (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], + rtac ctxt (in_Jrel RS iffD2), rtac ctxt exI, rtac ctxt conjI, rtac ctxt CollectI, CONJ_WRAP' (fn (dtor_set, passive_set_map0) => - EVERY' [rtac ord_eq_le_trans, rtac dtor_set, rtac @{thm Un_least}, - rtac ord_eq_le_trans, rtac @{thm box_equals}, rtac passive_set_map0, - rtac (dtor_ctor RS sym RS arg_cong), rtac trans_fun_cong_image_id_id_apply, atac, - CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least})) - (fn (active_set_map0, in_Jrel) => EVERY' [rtac ord_eq_le_trans, - rtac @{thm SUP_cong[OF _ refl]}, rtac @{thm box_equals[OF _ _ refl]}, - rtac active_set_map0, rtac (dtor_ctor RS sym RS arg_cong), rtac @{thm UN_least}, - dtac set_rev_mp, etac @{thm image_mono}, etac imageE, - dtac @{thm ssubst_mem[OF pair_collapse]}, + EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt dtor_set, rtac ctxt @{thm Un_least}, + rtac ctxt ord_eq_le_trans, rtac ctxt @{thm box_equals}, rtac ctxt passive_set_map0, + rtac ctxt (dtor_ctor RS sym RS arg_cong), rtac ctxt trans_fun_cong_image_id_id_apply, atac, + CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 @{thm Un_least})) + (fn (active_set_map0, in_Jrel) => EVERY' [rtac ctxt ord_eq_le_trans, + rtac ctxt @{thm SUP_cong[OF _ refl]}, rtac ctxt @{thm box_equals[OF _ _ refl]}, + rtac ctxt active_set_map0, rtac ctxt (dtor_ctor RS sym RS arg_cong), rtac ctxt @{thm UN_least}, + dtac ctxt set_rev_mp, etac ctxt @{thm image_mono}, etac ctxt imageE, + dtac ctxt @{thm ssubst_mem[OF pair_collapse]}, REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE :: @{thms case_prodE iffD1[OF prod.inject, elim_format]}), hyp_subst_tac ctxt, - dtac (in_Jrel RS iffD1), - dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, + dtac ctxt (in_Jrel RS iffD1), + dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], atac]) (rev (active_set_map0s ~~ in_Jrels))]) (dtor_sets ~~ passive_set_map0s), - rtac conjI, - REPEAT_DETERM_N 2 o EVERY'[rtac (dtor_inject RS iffD1), rtac trans, rtac dtor_map, - rtac @{thm box_equals}, rtac map_comp0, rtac (dtor_ctor RS sym RS arg_cong), rtac trans, - rtac map_cong0, REPEAT_DETERM_N m o rtac @{thm fun_cong[OF comp_id]}, - EVERY' (map (fn in_Jrel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac, - dtac @{thm ssubst_mem[OF pair_collapse]}, + rtac ctxt conjI, + REPEAT_DETERM_N 2 o EVERY'[rtac ctxt (dtor_inject RS iffD1), rtac ctxt trans, rtac ctxt dtor_map, + rtac ctxt @{thm box_equals}, rtac ctxt map_comp0, rtac ctxt (dtor_ctor RS sym RS arg_cong), rtac ctxt trans, + rtac ctxt map_cong0, REPEAT_DETERM_N m o rtac ctxt @{thm fun_cong[OF comp_id]}, + EVERY' (map (fn in_Jrel => EVERY' [rtac ctxt trans, rtac ctxt o_apply, dtac ctxt set_rev_mp, atac, + dtac ctxt @{thm ssubst_mem[OF pair_collapse]}, REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE :: @{thms case_prodE iffD1[OF prod.inject, elim_format]}), - hyp_subst_tac ctxt, dtac (in_Jrel RS iffD1), - dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) in_Jrels), + hyp_subst_tac ctxt, dtac ctxt (in_Jrel RS iffD1), + dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE, atac]) in_Jrels), atac]] in - EVERY' [rtac iffI, if_tac, only_if_tac] 1 + EVERY' [rtac ctxt iffI, if_tac, only_if_tac] 1 end; fun mk_rel_coinduct_coind_tac ctxt fst m coinduct ks map_comp0s map_congs map_arg_congs set_map0ss @@ -1017,29 +1018,29 @@ val fst_diag_nth = if fst then @{thm fst_diag_fst} else @{thm fst_diag_snd}; val snd_diag_nth = if fst then @{thm snd_diag_fst} else @{thm snd_diag_snd}; in - EVERY' [rtac coinduct, + EVERY' [rtac ctxt coinduct, EVERY' (@{map 8} (fn i => fn map_comp0 => fn map_cong => fn map_arg_cong => fn set_map0s => fn dtor_unfold => fn dtor_map => fn in_rel => EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI, in_rel RS iffD2], REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], - select_prem_tac (length ks) (dtac @{thm spec2}) i, dtac mp, atac, + select_prem_tac ctxt (length ks) (dtac ctxt @{thm spec2}) i, dtac ctxt mp, atac, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], hyp_subst_tac ctxt, - rtac exI, rtac (Drule.rotate_prems 1 conjI), rtac conjI, - rtac (map_comp0 RS trans), rtac (dtor_map RS trans RS sym), - rtac (dtor_unfold RS map_arg_cong RS trans), rtac (trans OF [map_comp0, map_cong]), - REPEAT_DETERM_N m o rtac (fst_diag_nth RS @{thm fun_cong[OF trans[OF o_id sym]]}), - REPEAT_DETERM_N n o (rtac @{thm sym[OF trans[OF o_apply]]} THEN' rtac @{thm fst_conv}), - rtac (map_comp0 RS trans), rtac (map_cong RS trans), - REPEAT_DETERM_N m o rtac (snd_diag_nth RS fun_cong), - REPEAT_DETERM_N n o (rtac @{thm trans[OF o_apply]} THEN' rtac @{thm snd_conv}), - etac (@{thm prod.case} RS map_arg_cong RS trans), + rtac ctxt exI, rtac ctxt (Drule.rotate_prems 1 conjI), rtac ctxt conjI, + rtac ctxt (map_comp0 RS trans), rtac ctxt (dtor_map RS trans RS sym), + rtac ctxt (dtor_unfold RS map_arg_cong RS trans), rtac ctxt (trans OF [map_comp0, map_cong]), + REPEAT_DETERM_N m o rtac ctxt (fst_diag_nth RS @{thm fun_cong[OF trans[OF o_id sym]]}), + REPEAT_DETERM_N n o (rtac ctxt @{thm sym[OF trans[OF o_apply]]} THEN' rtac ctxt @{thm fst_conv}), + rtac ctxt (map_comp0 RS trans), rtac ctxt (map_cong RS trans), + REPEAT_DETERM_N m o rtac ctxt (snd_diag_nth RS fun_cong), + REPEAT_DETERM_N n o (rtac ctxt @{thm trans[OF o_apply]} THEN' rtac ctxt @{thm snd_conv}), + etac ctxt (@{thm prod.case} RS map_arg_cong RS trans), SELECT_GOAL (unfold_thms_tac ctxt @{thms prod.case o_def fst_conv snd_conv}), - rtac CollectI, + rtac ctxt CollectI, CONJ_WRAP' (fn set_map0 => - EVERY' [rtac (set_map0 RS ord_eq_le_trans), - rtac @{thm image_subsetI}, rtac CollectI, rtac @{thm case_prodI}, - FIRST' [rtac refl, EVERY'[rtac exI, rtac conjI, etac Collect_splitD_set_mp, atac, - rtac (@{thm surjective_pairing} RS arg_cong)]]]) + EVERY' [rtac ctxt (set_map0 RS ord_eq_le_trans), + rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, + FIRST' [rtac ctxt refl, EVERY'[rtac ctxt exI, rtac ctxt conjI, etac ctxt Collect_splitD_set_mp, atac, + rtac ctxt (@{thm surjective_pairing} RS arg_cong)]]]) set_map0s]) ks map_comp0s map_congs map_arg_congs set_map0ss dtor_unfolds dtor_maps in_rels)] 1 end; @@ -1049,23 +1050,23 @@ fun mk_rel_coinduct_ind_tac ctxt m ks unfolds set_map0ss j set_induct = let val n = length ks; in - rtac set_induct 1 THEN + rtac ctxt set_induct 1 THEN EVERY' (@{map 3} (fn unfold => fn set_map0s => fn i => - EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac conjE, - select_prem_tac n (dtac @{thm spec2}) i, dtac mp, atac, + EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac ctxt conjE, + select_prem_tac ctxt n (dtac ctxt @{thm spec2}) i, dtac ctxt mp, atac, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE, Collect_splitD_set_mp, set_rev_mp], hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt ([unfold, nth set_map0s (j - 1)] @ split_id_unfolds)), - rtac subset_refl]) + rtac ctxt subset_refl]) unfolds set_map0ss ks) 1 THEN EVERY' (@{map 3} (fn unfold => fn set_map0s => fn i => EVERY' (map (fn set_map0 => - EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac conjE, - select_prem_tac n (dtac @{thm spec2}) i, dtac mp, atac, + EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac ctxt conjE, + select_prem_tac ctxt n (dtac ctxt @{thm spec2}) i, dtac ctxt mp, atac, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt ([unfold, set_map0] @ split_id_unfolds)), - etac imageE, hyp_subst_tac ctxt, REPEAT_DETERM o eresolve_tac ctxt [allE, mp], - rtac conjI, etac Collect_splitD_set_mp, atac, rtac (@{thm surjective_pairing} RS arg_cong)]) + etac ctxt imageE, hyp_subst_tac ctxt, REPEAT_DETERM o eresolve_tac ctxt [allE, mp], + rtac ctxt conjI, etac ctxt Collect_splitD_set_mp, atac, rtac ctxt (@{thm surjective_pairing} RS arg_cong)]) (drop m set_map0s))) unfolds set_map0ss ks) 1 end; @@ -1075,23 +1076,23 @@ in Method.insert_tac CIHs 1 THEN unfold_thms_tac ctxt (@{thm choice_iff} :: @{thm ex_simps(6)[symmetric]} :: in_rels) THEN - REPEAT_DETERM (etac exE 1) THEN + REPEAT_DETERM (etac ctxt exE 1) THEN CONJ_WRAP' (fn (in_Jrel, (helper_inds, (helper_coind1, helper_coind2))) => - EVERY' [rtac @{thm predicate2I}, rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, - if null helper_inds then rtac UNIV_I - else rtac CollectI THEN' + EVERY' [rtac ctxt @{thm predicate2I}, rtac ctxt (in_Jrel RS iffD2), rtac ctxt exI, rtac ctxt conjI, + if null helper_inds then rtac ctxt UNIV_I + else rtac ctxt CollectI THEN' CONJ_WRAP' (fn helper_ind => - EVERY' [rtac (helper_ind RS rev_mp), REPEAT_DETERM_N n o atac, - REPEAT_DETERM_N n o etac thin_rl, rtac impI, + EVERY' [rtac ctxt (helper_ind RS rev_mp), REPEAT_DETERM_N n o atac, + REPEAT_DETERM_N n o etac ctxt thin_rl, rtac ctxt impI, REPEAT_DETERM o resolve_tac ctxt [subsetI, CollectI, @{thm iffD2[OF split_beta]}], - dtac bspec, atac, REPEAT_DETERM o eresolve_tac ctxt [allE, mp, conjE], - etac (refl RSN (2, conjI))]) + dtac ctxt bspec, atac, REPEAT_DETERM o eresolve_tac ctxt [allE, mp, conjE], + etac ctxt (refl RSN (2, conjI))]) helper_inds, - rtac conjI, - rtac (helper_coind1 RS rev_mp), REPEAT_DETERM_N n o atac, REPEAT_DETERM_N n o etac thin_rl, - rtac impI, etac mp, rtac exI, etac (refl RSN (2, conjI)), - rtac (helper_coind2 RS rev_mp), REPEAT_DETERM_N n o atac, REPEAT_DETERM_N n o etac thin_rl, - rtac impI, etac mp, rtac exI, etac (refl RSN (2, conjI))]) + rtac ctxt conjI, + rtac ctxt (helper_coind1 RS rev_mp), REPEAT_DETERM_N n o atac, REPEAT_DETERM_N n o etac ctxt thin_rl, + rtac ctxt impI, etac ctxt mp, rtac ctxt exI, etac ctxt (refl RSN (2, conjI)), + rtac ctxt (helper_coind2 RS rev_mp), REPEAT_DETERM_N n o atac, REPEAT_DETERM_N n o etac ctxt thin_rl, + rtac ctxt impI, etac ctxt mp, rtac ctxt exI, etac ctxt (refl RSN (2, conjI))]) (in_Jrels ~~ (helper_indss ~~ (helper_coind1s ~~ helper_coind2s))) 1 end; @@ -1103,14 +1104,14 @@ @{thms rel_fun_def_butlast all_conj_distrib[symmetric] imp_conjR[symmetric]} THEN unfold_thms_tac ctxt @{thms rel_fun_iff_geq_image2p} THEN HEADGOAL (EVERY' - [REPEAT_DETERM o resolve_tac ctxt [allI, impI], rtac ctor_rel_coinduct, + [REPEAT_DETERM o resolve_tac ctxt [allI, impI], rtac ctxt ctor_rel_coinduct, EVERY' (map (fn map_transfer => EVERY' - [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac @{thm image2pE}, hyp_subst_tac ctxt, + [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac ctxt @{thm image2pE}, hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt unfolds), - rtac (funpow (m + n + 1) (fn thm => thm RS rel_funD) map_transfer), - REPEAT_DETERM_N m o rtac @{thm id_transfer}, - REPEAT_DETERM_N n o rtac @{thm rel_fun_image2p}, - etac @{thm predicate2D}, etac @{thm image2pI}]) + rtac ctxt (funpow (m + n + 1) (fn thm => thm RS rel_funD) map_transfer), + REPEAT_DETERM_N m o rtac ctxt @{thm id_transfer}, + REPEAT_DETERM_N n o rtac ctxt @{thm rel_fun_image2p}, + etac ctxt @{thm predicate2D}, etac ctxt @{thm image2pI}]) map_transfers)]) end; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_lfp.ML --- a/src/HOL/Tools/BNF/bnf_lfp.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_lfp.ML Thu Jul 16 12:23:22 2015 +0200 @@ -221,7 +221,7 @@ val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x); in Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, goal)) - (K (mk_map_cong0L_tac m map_cong0 map_id)) + (fn {context = ctxt, prems = _} => mk_map_cong0L_tac ctxt m map_cong0 map_id) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -278,7 +278,8 @@ Logic.list_implies (alg_prem :: prems, concl)) premss concls; in map (fn goal => - Goal.prove_sorry lthy [] [] goal (K (mk_alg_set_tac alg_def)) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_alg_set_tac ctxt alg_def) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) goals @@ -355,7 +356,8 @@ Logic.list_implies ([prem, mk_elim_prem sets x T], mk_Trueprop_eq (f $ (s $ x), s' $ Term.list_comb (mapAsBs, passive_ids @ fs @ [x]))); val elim_goals = @{map 7} mk_elim_goal setssAs mapsAsBs fs ss s's xFs FTsAs; - fun prove goal = Goal.prove_sorry lthy [] [] goal (K (mk_mor_elim_tac mor_def)) + fun prove goal = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_mor_elim_tac ctxt mor_def) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); in @@ -368,7 +370,7 @@ val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids); in Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl)) - (fn {context = ctxt, ...} => mk_mor_incl_tac ctxt mor_def map_ids) + (fn {context = ctxt, prems = _} => mk_mor_incl_tac ctxt mor_def map_ids) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -382,7 +384,7 @@ HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs)); in Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl)) - (K (mk_mor_comp_tac mor_def set_mapss map_comp_id_thms)) + (fn {context = ctxt, prems = _} => mk_mor_comp_tac ctxt mor_def set_mapss map_comp_id_thms) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -394,7 +396,7 @@ val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy); in Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl)) - (K ((hyp_subst_tac lthy THEN' assume_tac lthy) 1)) + (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' assume_tac ctxt) 1) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -407,7 +409,7 @@ Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor (map HOLogic.mk_UNIV FTsAs) maps active_UNIVs ss ss)) - (K (mk_mor_str_tac ks mor_def)) + (fn {context = ctxt, prems = _} => mk_mor_str_tac ctxt ks mor_def) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -421,7 +423,7 @@ Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor prod_UNIVs maps (map HOLogic.mk_UNIV activeBs) s's fsts)) - (K (mk_mor_convol_tac ks mor_def)) + (fn {context = ctxt, prems = _} => mk_mor_convol_tac ctxt ks mor_def) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -435,7 +437,7 @@ val rhs = Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct mapsAsBs fs ss s's); in Goal.prove_sorry lthy [] [] (mk_Trueprop_eq (lhs, rhs)) - (K (mk_mor_UNIV_tac m morE_thms mor_def)) + (fn {context = ctxt, prems = _} => mk_mor_UNIV_tac ctxt m morE_thms mor_def) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -457,7 +459,8 @@ val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) = typedef (sbdT_bind, sum_bdT_params', NoSyn) - (HOLogic.mk_UNIV sum_bdT) NONE (fn _ => EVERY' [rtac exI, rtac UNIV_I] 1) lthy; + (HOLogic.mk_UNIV sum_bdT) NONE (fn ctxt => + EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val sbdT = Type (sbdT_name, sum_bdT_params); val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT); @@ -548,7 +551,7 @@ (Term.absfree jdx' (Library.foldr1 HOLogic.mk_conj (map mk_conjunct idxs)))); in Goal.prove_sorry lthy [] [] (Logic.list_implies ([prem], concl)) - (K (mk_bd_limit_tac n suc_bd_Cinfinite)) + (fn {context = ctxt, prems = _} => mk_bd_limit_tac ctxt n suc_bd_Cinfinite) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -589,7 +592,7 @@ val goal = Logic.mk_implies (HOLogic.mk_Trueprop i_field, concl); val min_algs_thm = Goal.prove_sorry lthy [] [] goal - (K (mk_min_algs_tac suc_bd_worel in_cong'_thms)) + (fn {context = ctxt, prems = _} => mk_min_algs_tac ctxt suc_bd_worel in_cong'_thms) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); @@ -613,10 +616,10 @@ val card_of = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, card_conjunction))) - (K (mk_min_algs_card_of_tac card_cT card_ct + (fn {context = ctxt, prems = _} => mk_min_algs_card_of_tac ctxt card_cT card_ct m suc_bd_worel min_algs_thms in_sbds sbd_Card_order sbd_Cnotzero suc_bd_Card_order suc_bd_Cinfinite suc_bd_Cnotzero - suc_bd_Asuc_bd Asuc_bd_Cinfinite)) + suc_bd_Asuc_bd Asuc_bd_Cinfinite) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); @@ -629,8 +632,8 @@ (Goal.prove_sorry lthy [] [] (Logic.mk_implies (least_prem, HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, least_conjunction)))) - (K (mk_min_algs_least_tac least_cT least_ct - suc_bd_worel min_algs_thms alg_set_thms))) + (fn {context = ctxt, prems = _} => mk_min_algs_least_tac ctxt least_cT least_ct + suc_bd_worel min_algs_thms alg_set_thms)) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); in @@ -678,21 +681,21 @@ let val goal = HOLogic.mk_Trueprop (mk_alg min_algs ss); val alg_min_alg = Goal.prove_sorry lthy [] [] goal - (K (mk_alg_min_alg_tac m alg_def min_alg_defs suc_bd_limit_thm sbd_Cinfinite - set_sbdss min_algs_thms min_algs_mono_thms)) + (fn {context = ctxt, prems = _} => mk_alg_min_alg_tac ctxt m alg_def min_alg_defs + suc_bd_limit_thm sbd_Cinfinite set_sbdss min_algs_thms min_algs_mono_thms) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); fun mk_card_of_thm min_alg def = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of min_alg) Asuc_bd)) - (K (mk_card_of_min_alg_tac def card_of_min_algs_thm - suc_bd_Card_order suc_bd_Asuc_bd Asuc_bd_Cinfinite)) + (fn {context = ctxt, prems = _} => mk_card_of_min_alg_tac ctxt def card_of_min_algs_thm + suc_bd_Card_order suc_bd_Asuc_bd Asuc_bd_Cinfinite) |> Thm.close_derivation; val least_prem = HOLogic.mk_Trueprop (mk_alg Bs ss); fun mk_least_thm min_alg B def = Goal.prove_sorry lthy [] [] (Logic.mk_implies (least_prem, HOLogic.mk_Trueprop (mk_leq min_alg B))) - (K (mk_least_min_alg_tac def least_min_algs_thm)) + (fn {context = ctxt, prems = _} => mk_least_min_alg_tac ctxt def least_min_algs_thm) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); @@ -702,7 +705,8 @@ val incl = Goal.prove_sorry lthy [] [] (Logic.mk_implies (incl_prem, HOLogic.mk_Trueprop (mk_mor min_algs ss Bs ss active_ids))) - (K (EVERY' (rtac mor_incl_thm :: map etac leasts) 1)) + (fn {context = ctxt, prems = _} => + EVERY' (rtac ctxt mor_incl_thm :: map (etac ctxt) leasts) 1) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); in @@ -716,7 +720,7 @@ val ((IIT_name, (IIT_glob_info, IIT_loc_info)), lthy) = typedef (IIT_bind, params, NoSyn) - (HOLogic.mk_UNIV II_repT) NONE (fn _ => EVERY' [rtac exI, rtac UNIV_I] 1) lthy; + (HOLogic.mk_UNIV II_repT) NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val IIT = Type (IIT_name, params'); val Abs_IIT = Const (#Abs_name IIT_glob_info, II_repT --> IIT); @@ -797,8 +801,9 @@ (map (fn f => HOLogic.mk_comp (f, mk_rapp iidx Asuc_bdT)) Asuc_fs)); in Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl)) - (K (mk_mor_select_tac mor_def mor_cong_thm mor_comp_thm mor_incl_min_alg_thm alg_def - alg_select_thm alg_set_thms set_mapss str_init_defs)) + (fn {context = ctxt, prems = _} => mk_mor_select_tac ctxt mor_def mor_cong_thm + mor_comp_thm mor_incl_min_alg_thm alg_def alg_select_thm alg_set_thms set_mapss + str_init_defs) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -815,8 +820,8 @@ val cts = map (Thm.cterm_of lthy) ss; val unique_mor = Goal.prove_sorry lthy [] [] (Logic.list_implies (prems @ mor_prems, unique)) - (K (mk_init_unique_mor_tac cts m alg_def alg_init_thm least_min_alg_thms - in_mono'_thms alg_set_thms morE_thms map_cong0s)) + (fn {context = ctxt, prems = _} => mk_init_unique_mor_tac ctxt cts m alg_def + alg_init_thm least_min_alg_thms in_mono'_thms alg_set_thms morE_thms map_cong0s) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); in @@ -849,7 +854,8 @@ (map2 mk_Ball car_inits init_phis)); in Goal.prove_sorry lthy [] [] (Logic.mk_implies (prem, concl)) - (K (mk_init_induct_tac m alg_def alg_init_thm least_min_alg_thms alg_set_thms)) + (fn {context = ctxt, prems = _} => mk_init_induct_tac ctxt m alg_def alg_init_thm + least_min_alg_thms alg_set_thms) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -861,8 +867,8 @@ |> @{fold_map 3} (fn b => fn mx => fn car_init => typedef (b, params, mx) car_init NONE (fn ctxt => - EVERY' [rtac iffD2, rtac @{thm ex_in_conv}, resolve_tac ctxt alg_not_empty_thms, - rtac alg_init_thm] 1)) bs mixfixes car_inits + EVERY' [rtac ctxt iffD2, rtac ctxt @{thm ex_in_conv}, resolve_tac ctxt alg_not_empty_thms, + rtac ctxt alg_init_thm] 1)) bs mixfixes car_inits |>> apsnd split_list o split_list; val Ts = map (fn name => Type (name, params')) T_names; @@ -999,7 +1005,8 @@ (HOLogic.mk_conj (mk_alg B's s's, mk_mor B's s's Bs ss inv_fs))); in Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl)) - (K (mk_copy_tac m alg_def mor_def alg_set_thms set_mapss)) + (fn {context = ctxt, prems = _} => mk_copy_tac ctxt m alg_def mor_def alg_set_thms + set_mapss) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy) end; @@ -1032,7 +1039,7 @@ end; val ctor_fold_thms = map (fn morE => rule_by_tactic lthy - ((rtac CollectI THEN' CONJ_WRAP' (K (rtac @{thm subset_UNIV})) (1 upto m + n)) 1) + ((rtac lthy CollectI THEN' CONJ_WRAP' (K (rtac lthy @{thm subset_UNIV})) (1 upto m + n)) 1) (mor_fold_thm RS morE)) morE_thms; val (fold_unique_mor_thms, fold_unique_mor_thm) = @@ -1041,8 +1048,8 @@ fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_fold Ts ss i); val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq fs ks)); val unique_mor = Goal.prove_sorry lthy [] [] (Logic.mk_implies (prem, unique)) - (K (mk_fold_unique_mor_tac type_defs init_unique_mor_thms Reps - mor_comp_thm mor_Abs_thm mor_fold_thm)) + (fn {context = ctxt, prems = _} => mk_fold_unique_mor_tac ctxt type_defs + init_unique_mor_thms Reps mor_comp_thm mor_Abs_thm mor_fold_thm) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); in @@ -1100,7 +1107,8 @@ in @{map 5} (fn goal => fn dtor_def => fn foldx => fn map_comp_id => fn map_cong0L => Goal.prove_sorry lthy [] [] goal - (K (mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_cong0L ctor_o_fold_thms)) + (fn {context = ctxt, prems = _} => mk_dtor_o_ctor_tac ctxt dtor_def foldx map_comp_id + map_cong0L ctor_o_fold_thms) |> Thm.close_derivation) goals dtor_defs ctor_fold_thms map_comp_id_thms map_cong0L_thms end; @@ -1395,7 +1403,8 @@ val ((sbd0T_name, (sbd0T_glob_info, sbd0T_loc_info)), lthy) = typedef (sbd0T_bind, sum_bd0T_params', NoSyn) - (HOLogic.mk_UNIV sum_bd0T) NONE (fn _ => EVERY' [rtac exI, rtac UNIV_I] 1) lthy; + (HOLogic.mk_UNIV sum_bd0T) NONE (fn ctxt => + EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val sbd0T = Type (sbd0T_name, sum_bd0T_params); val Abs_sbd0T = Const (#Abs_name sbd0T_glob_info, sum_bd0T --> sbd0T); @@ -1481,7 +1490,7 @@ @{map 4} (fn goal => fn foldx => fn map_comp_id => fn map_cong0 => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN - mk_map_tac m n foldx map_comp_id map_cong0) + mk_map_tac ctxt m n foldx map_comp_id map_cong0) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) goals ctor_fold_thms map_comp_id_thms map_cong0s; @@ -1519,7 +1528,7 @@ Isetss_by_range colss map_setss; val setss = map (map2 (fn foldx => fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => - unfold_thms_tac ctxt Iset_defs THEN mk_set_tac foldx) + unfold_thms_tac ctxt Iset_defs THEN mk_set_tac ctxt foldx) |> Thm.close_derivation) ctor_fold_thms) goalss; @@ -1535,7 +1544,8 @@ val ctor_setss = @{map 3} (fn i => @{map 3} (fn set_nats => fn goal => fn set => Goal.prove_sorry lthy [] [] goal - (K (mk_ctor_set_tac set (nth set_nats (i - 1)) (drop m set_nats))) + (fn {context = ctxt, prems = _} => + mk_ctor_set_tac ctxt set (nth set_nats (i - 1)) (drop m set_nats)) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy)) set_mapss) ls simp_goalss setss; @@ -1578,7 +1588,7 @@ (@{map 4} (mk_set_map0 f) fs_Imaps Izs sets sets'))) fs Isetss_by_range Isetss_by_range'; - fun mk_tac ctxt induct = mk_set_nat_tac ctxt m (rtac induct) set_mapss ctor_Imap_thms; + fun mk_tac ctxt induct = mk_set_nat_tac ctxt m (rtac ctxt induct) set_mapss ctor_Imap_thms; val thms = @{map 5} (fn goal => fn csets => fn ctor_sets => fn induct => fn i => Goal.prove_sorry lthy [] [] goal @@ -1606,7 +1616,7 @@ HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_set_bd Izs Ibds sets))) Isetss_by_range; - fun mk_tac ctxt induct = mk_set_bd_tac ctxt m (rtac induct) sbd0_Cinfinite set_sbd0ss; + fun mk_tac ctxt induct = mk_set_bd_tac ctxt m (rtac ctxt induct) sbd0_Cinfinite set_sbd0ss; val thms = @{map 4} (fn goal => fn ctor_sets => fn induct => fn i => Goal.prove_sorry lthy [] [] goal @@ -1641,7 +1651,7 @@ (@{map 4} mk_map_cong0 Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps)); val thm = Goal.prove_sorry lthy [] [] goal - (fn {context = ctxt, prems = _} => mk_mcong_tac ctxt (rtac induct) set_Iset_thmsss + (fn {context = ctxt, prems = _} => mk_mcong_tac ctxt (rtac ctxt induct) set_Iset_thmsss map_cong0s ctor_Imap_thms) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); @@ -1708,20 +1718,24 @@ val timer = time (timer "helpers for BNF properties"); - val map_id0_tacs = map (K o mk_map_id0_tac map_id0s) ctor_Imap_unique_thms; + val map_id0_tacs = map (fn thm => fn ctxt => mk_map_id0_tac ctxt map_id0s thm) + ctor_Imap_unique_thms; val map_comp0_tacs = - map2 (K oo mk_map_comp0_tac map_comps ctor_Imap_thms) ctor_Imap_unique_thms ks; + map2 (fn thm => fn i => fn ctxt => + mk_map_comp0_tac ctxt map_comps ctor_Imap_thms thm i) + ctor_Imap_unique_thms ks; val map_cong0_tacs = map (fn thm => fn ctxt => mk_map_cong0_tac ctxt m thm) Imap_cong0_thms; - val set_map0_tacss = map (map (K o mk_set_map0_tac)) (transpose Iset_Imap0_thmss); + val set_map0_tacss = map (map (fn thm => fn ctxt => mk_set_map0_tac ctxt thm)) + (transpose Iset_Imap0_thmss); val bd_co_tacs = replicate n (fn ctxt => - unfold_thms_tac ctxt Ibd_defs THEN rtac sbd0_card_order 1); + unfold_thms_tac ctxt Ibd_defs THEN rtac ctxt sbd0_card_order 1); val bd_cinf_tacs = replicate n (fn ctxt => - unfold_thms_tac ctxt Ibd_defs THEN rtac (sbd0_Cinfinite RS conjunct1) 1); - val set_bd_tacss = map (map (fn thm => K (rtac thm 1))) (transpose Iset_bd_thmss); - val le_rel_OO_tacs = map (fn i => - K ((rtac @{thm predicate2I} THEN' etac (le_Irel_OO_thm RS mk_conjunctN n i RS mp)) 1)) ks; + unfold_thms_tac ctxt Ibd_defs THEN rtac ctxt (sbd0_Cinfinite RS conjunct1) 1); + val set_bd_tacss = map (map (fn thm => fn ctxt => rtac ctxt thm 1)) (transpose Iset_bd_thmss); + val le_rel_OO_tacs = map (fn i => fn ctxt => + (rtac ctxt @{thm predicate2I} THEN' etac ctxt (le_Irel_OO_thm RS mk_conjunctN n i RS mp)) 1) ks; - val rel_OO_Grp_tacs = map (fn def => K (rtac def 1)) Irel_unabs_defs; + val rel_OO_Grp_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Irel_unabs_defs; val tacss = @{map 9} zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_lfp_compat.ML --- a/src/HOL/Tools/BNF/bnf_lfp_compat.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_lfp_compat.ML Thu Jul 16 12:23:22 2015 +0200 @@ -153,7 +153,7 @@ fun tac ctxt = unfold_thms_tac ctxt (@{thms o_apply fst_conv snd_conv} @ rec_defs @ flat rec0_thmss) THEN - HEADGOAL (rtac refl); + HEADGOAL (rtac ctxt refl); fun prove goal = Goal.prove_sorry ctxt [] [] goal (tac o #context) diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML --- a/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML Thu Jul 16 12:23:22 2015 +0200 @@ -493,9 +493,9 @@ fun mk_primrec_tac ctxt num_extra_args map_ident0s map_comps fun_defs recx = unfold_thms_tac ctxt fun_defs THEN - HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN + HEADGOAL (rtac ctxt (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN unfold_thms_tac ctxt (nested_simps ctxt @ map_ident0s @ map_comps) THEN - HEADGOAL (rtac refl); + HEADGOAL (rtac ctxt refl); fun prepare_primrec plugins nonexhaustives transfers fixes specs lthy0 = let diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_lfp_size.ML --- a/src/HOL/Tools/BNF/bnf_lfp_size.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML Thu Jul 16 12:23:22 2015 +0200 @@ -63,15 +63,15 @@ fun mk_size_gen_o_map_tac ctxt size_def rec_o_map inj_maps size_maps = unfold_thms_tac ctxt [size_def] THEN - HEADGOAL (rtac (rec_o_map RS trans) THEN' + HEADGOAL (rtac ctxt (rec_o_map RS trans) THEN' asm_simp_tac (ss_only (inj_maps @ size_maps @ size_gen_o_map_simps) ctxt)) THEN - IF_UNSOLVED (unfold_thms_tac ctxt @{thms id_def o_def} THEN HEADGOAL (rtac refl)); + IF_UNSOLVED (unfold_thms_tac ctxt @{thms id_def o_def} THEN HEADGOAL (rtac ctxt refl)); fun mk_size_neq ctxt cts exhaust sizes = - HEADGOAL (rtac (Ctr_Sugar_Util.cterm_instantiate_pos (map SOME cts) exhaust)) THEN + HEADGOAL (rtac ctxt (Ctr_Sugar_Util.cterm_instantiate_pos (map SOME cts) exhaust)) THEN ALLGOALS (hyp_subst_tac ctxt) THEN Ctr_Sugar_Tactics.unfold_thms_tac ctxt (@{thm neq0_conv} :: sizes) THEN - ALLGOALS (REPEAT_DETERM o (rtac @{thm zero_less_Suc} ORELSE' rtac @{thm trans_less_add2})); + ALLGOALS (REPEAT_DETERM o (rtac ctxt @{thm zero_less_Suc} ORELSE' rtac ctxt @{thm trans_less_add2})); fun generate_datatype_size (fp_sugars as ({T = Type (_, As), BT = Type (_, Bs), fp = Least_FP, fp_res = {bnfs = fp_bnfs, ...}, fp_nesting_bnfs, live_nesting_bnfs, ...} : fp_sugar) :: _) diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_lfp_tactics.ML --- a/src/HOL/Tools/BNF/bnf_lfp_tactics.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_lfp_tactics.ML Thu Jul 16 12:23:22 2015 +0200 @@ -8,59 +8,60 @@ signature BNF_LFP_TACTICS = sig - val mk_alg_min_alg_tac: int -> thm -> thm list -> thm -> thm -> thm list list -> thm list -> - thm list -> tactic + val mk_alg_min_alg_tac: Proof.context -> int -> thm -> thm list -> thm -> thm -> thm list list -> + thm list -> thm list -> tactic val mk_alg_not_empty_tac: Proof.context -> thm -> thm list -> thm list -> tactic val mk_alg_select_tac: Proof.context -> thm -> tactic - val mk_alg_set_tac: thm -> tactic - val mk_bd_card_order_tac: thm list -> tactic - val mk_bd_limit_tac: int -> thm -> tactic - val mk_card_of_min_alg_tac: thm -> thm -> thm -> thm -> thm -> tactic - val mk_copy_tac: int -> thm -> thm -> thm list -> thm list list -> tactic + val mk_alg_set_tac: Proof.context -> thm -> tactic + val mk_bd_card_order_tac: Proof.context -> thm list -> tactic + val mk_bd_limit_tac: Proof.context -> int -> thm -> tactic + val mk_card_of_min_alg_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> tactic + val mk_copy_tac: Proof.context -> int -> thm -> thm -> thm list -> thm list list -> tactic val mk_ctor_induct_tac: Proof.context -> int -> thm list list -> thm -> thm list -> thm -> thm list -> thm list -> thm list -> tactic val mk_ctor_induct2_tac: Proof.context -> ctyp option list -> cterm option list -> thm -> thm list -> tactic - val mk_ctor_set_tac: thm -> thm -> thm list -> tactic + val mk_ctor_set_tac: Proof.context -> thm -> thm -> thm list -> tactic val mk_ctor_rec_transfer_tac: Proof.context -> int -> int -> thm list -> thm list -> thm list -> thm list -> tactic val mk_ctor_rel_tac: Proof.context -> thm list -> int -> thm -> thm -> thm -> thm -> thm list -> thm -> thm -> thm list -> thm list -> thm list list -> tactic - val mk_dtor_o_ctor_tac: thm -> thm -> thm -> thm -> thm list -> tactic + val mk_dtor_o_ctor_tac: Proof.context -> thm -> thm -> thm -> thm -> thm list -> tactic val mk_init_ex_mor_tac: Proof.context -> thm -> thm -> thm list -> thm -> thm -> thm -> thm -> tactic - val mk_init_induct_tac: int -> thm -> thm -> thm list -> thm list -> tactic - val mk_init_unique_mor_tac: cterm list -> int -> thm -> thm -> thm list -> thm list -> thm list -> - thm list -> thm list -> tactic - val mk_fold_unique_mor_tac: thm list -> thm list -> thm list -> thm -> thm -> thm -> tactic + val mk_init_induct_tac: Proof.context -> int -> thm -> thm -> thm list -> thm list -> tactic + val mk_init_unique_mor_tac: Proof.context -> cterm list -> int -> thm -> thm -> thm list -> + thm list -> thm list -> thm list -> thm list -> tactic + val mk_fold_unique_mor_tac: Proof.context -> thm list -> thm list -> thm list -> thm -> thm -> + thm -> tactic val mk_fold_transfer_tac: Proof.context -> int -> thm -> thm list -> thm list -> tactic - val mk_least_min_alg_tac: thm -> thm -> tactic + val mk_least_min_alg_tac: Proof.context -> thm -> thm -> tactic val mk_le_rel_OO_tac: Proof.context -> int -> thm -> thm list -> thm list -> thm list -> thm list -> tactic - val mk_map_comp0_tac: thm list -> thm list -> thm -> int -> tactic - val mk_map_id0_tac: thm list -> thm -> tactic - val mk_map_tac: int -> int -> thm -> thm -> thm -> tactic + val mk_map_comp0_tac: Proof.context -> thm list -> thm list -> thm -> int -> tactic + val mk_map_id0_tac: Proof.context -> thm list -> thm -> tactic + val mk_map_tac: Proof.context -> int -> int -> thm -> thm -> thm -> tactic val mk_ctor_map_unique_tac: Proof.context -> thm -> thm list -> tactic val mk_mcong_tac: Proof.context -> (int -> tactic) -> thm list list list -> thm list -> thm list -> tactic - val mk_min_algs_card_of_tac: ctyp -> cterm -> int -> thm -> thm list -> thm list -> thm -> thm -> - thm -> thm -> thm -> thm -> thm -> tactic - val mk_min_algs_least_tac: ctyp -> cterm -> thm -> thm list -> thm list -> tactic + val mk_min_algs_card_of_tac: Proof.context -> ctyp -> cterm -> int -> thm -> thm list -> + thm list -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> tactic + val mk_min_algs_least_tac: Proof.context -> ctyp -> cterm -> thm -> thm list -> thm list -> tactic val mk_min_algs_mono_tac: Proof.context -> thm -> tactic - val mk_min_algs_tac: thm -> thm list -> tactic + val mk_min_algs_tac: Proof.context -> thm -> thm list -> tactic val mk_mor_Abs_tac: Proof.context -> cterm list -> thm list -> thm list -> thm list -> thm list -> tactic val mk_mor_Rep_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm -> thm list -> thm list list -> tactic - val mk_mor_UNIV_tac: int -> thm list -> thm -> tactic - val mk_mor_comp_tac: thm -> thm list list -> thm list -> tactic - val mk_mor_convol_tac: 'a list -> thm -> tactic - val mk_mor_elim_tac: thm -> tactic + val mk_mor_UNIV_tac: Proof.context -> int -> thm list -> thm -> tactic + val mk_mor_comp_tac: Proof.context -> thm -> thm list list -> thm list -> tactic + val mk_mor_convol_tac: Proof.context -> 'a list -> thm -> tactic + val mk_mor_elim_tac: Proof.context -> thm -> tactic val mk_mor_incl_tac: Proof.context -> thm -> thm list -> tactic val mk_mor_fold_tac: Proof.context -> ctyp -> cterm -> thm list -> thm -> thm -> tactic - val mk_mor_select_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm list -> thm list list -> - thm list -> tactic - val mk_mor_str_tac: 'a list -> thm -> tactic + val mk_mor_select_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm list -> + thm list list -> thm list -> tactic + val mk_mor_str_tac: Proof.context -> 'a list -> thm -> tactic val mk_rel_induct_tac: Proof.context -> thm list -> int -> thm -> int list -> thm list -> thm list -> tactic val mk_rec_tac: Proof.context -> thm list -> thm -> thm list -> tactic @@ -69,8 +70,8 @@ int -> tactic val mk_set_nat_tac: Proof.context -> int -> (int -> tactic) -> thm list list -> thm list -> cterm list -> thm list -> int -> tactic - val mk_set_map0_tac: thm -> tactic - val mk_set_tac: thm -> tactic + val mk_set_map0_tac: Proof.context -> thm -> tactic + val mk_set_tac: Proof.context -> thm -> tactic val mk_wit_tac: Proof.context -> int -> thm list -> thm list -> tactic end; @@ -89,132 +90,132 @@ val Un_cong = @{thm arg_cong2[of _ _ _ _ "op \"]}; val relChainD = @{thm iffD2[OF meta_eq_to_obj_eq[OF relChain_def]]}; -fun mk_alg_set_tac alg_def = - EVERY' [dtac (alg_def RS iffD1), REPEAT_DETERM o etac conjE, etac bspec, rtac CollectI, - REPEAT_DETERM o (rtac (subset_UNIV RS conjI) ORELSE' etac conjI), atac] 1; +fun mk_alg_set_tac ctxt alg_def = + EVERY' [dtac ctxt (alg_def RS iffD1), REPEAT_DETERM o etac ctxt conjE, etac ctxt bspec, rtac ctxt CollectI, + REPEAT_DETERM o (rtac ctxt (subset_UNIV RS conjI) ORELSE' etac ctxt conjI), atac] 1; fun mk_alg_not_empty_tac ctxt alg_set alg_sets wits = - (EVERY' [rtac notI, hyp_subst_tac ctxt, forward_tac ctxt [alg_set]] THEN' + (EVERY' [rtac ctxt notI, hyp_subst_tac ctxt, forward_tac ctxt [alg_set]] THEN' REPEAT_DETERM o FIRST' - [EVERY' [rtac @{thm subset_emptyI}, eresolve_tac ctxt wits], - EVERY' [rtac subsetI, rtac FalseE, eresolve_tac ctxt wits], - EVERY' [rtac subsetI, dresolve_tac ctxt wits, hyp_subst_tac ctxt, - FIRST' (map (fn thm => rtac thm THEN' atac) alg_sets)]] THEN' - etac @{thm emptyE}) 1; + [EVERY' [rtac ctxt @{thm subset_emptyI}, eresolve_tac ctxt wits], + EVERY' [rtac ctxt subsetI, rtac ctxt FalseE, eresolve_tac ctxt wits], + EVERY' [rtac ctxt subsetI, dresolve_tac ctxt wits, hyp_subst_tac ctxt, + FIRST' (map (fn thm => rtac ctxt thm THEN' atac) alg_sets)]] THEN' + etac ctxt @{thm emptyE}) 1; -fun mk_mor_elim_tac mor_def = - (dtac (mor_def RS iffD1) THEN' - REPEAT o etac conjE THEN' - TRY o rtac @{thm image_subsetI} THEN' - etac bspec THEN' +fun mk_mor_elim_tac ctxt mor_def = + (dtac ctxt (mor_def RS iffD1) THEN' + REPEAT o etac ctxt conjE THEN' + TRY o rtac ctxt @{thm image_subsetI} THEN' + etac ctxt bspec THEN' atac) 1; fun mk_mor_incl_tac ctxt mor_def map_ids = - (rtac (mor_def RS iffD2) THEN' - rtac conjI THEN' - CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, etac (id_apply RS @{thm ssubst_mem})])) + (rtac ctxt (mor_def RS iffD2) THEN' + rtac ctxt conjI THEN' + CONJ_WRAP' (K (EVERY' [rtac ctxt ballI, etac ctxt set_mp, etac ctxt (id_apply RS @{thm ssubst_mem})])) map_ids THEN' CONJ_WRAP' (fn thm => - (EVERY' [rtac ballI, rtac trans, rtac id_apply, stac ctxt thm, rtac refl])) map_ids) 1; + (EVERY' [rtac ctxt ballI, rtac ctxt trans, rtac ctxt id_apply, stac ctxt thm, rtac ctxt refl])) map_ids) 1; -fun mk_mor_comp_tac mor_def set_maps map_comp_ids = +fun mk_mor_comp_tac ctxt mor_def set_maps map_comp_ids = let val fbetw_tac = - EVERY' [rtac ballI, rtac (o_apply RS @{thm ssubst_mem}), etac bspec, etac bspec, atac]; + EVERY' [rtac ctxt ballI, rtac ctxt (o_apply RS @{thm ssubst_mem}), etac ctxt bspec, etac ctxt bspec, atac]; fun mor_tac (set_map, map_comp_id) = - EVERY' [rtac ballI, rtac (o_apply RS trans), rtac trans, - rtac trans, dtac rev_bspec, atac, etac arg_cong, - REPEAT o eresolve0_tac [CollectE, conjE], etac bspec, rtac CollectI] THEN' + EVERY' [rtac ctxt ballI, rtac ctxt (o_apply RS trans), rtac ctxt trans, + rtac ctxt trans, dtac ctxt rev_bspec, atac, etac ctxt arg_cong, + REPEAT o eresolve0_tac [CollectE, conjE], etac ctxt bspec, rtac ctxt CollectI] THEN' CONJ_WRAP' (fn thm => - FIRST' [rtac subset_UNIV, - (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI}, - etac bspec, etac set_mp, atac])]) set_map THEN' - rtac (map_comp_id RS arg_cong); + FIRST' [rtac ctxt subset_UNIV, + (EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt thm, rtac ctxt @{thm image_subsetI}, + etac ctxt bspec, etac ctxt set_mp, atac])]) set_map THEN' + rtac ctxt (map_comp_id RS arg_cong); in - (dtac (mor_def RS iffD1) THEN' dtac (mor_def RS iffD1) THEN' rtac (mor_def RS iffD2) THEN' - REPEAT o etac conjE THEN' - rtac conjI THEN' + (dtac ctxt (mor_def RS iffD1) THEN' dtac ctxt (mor_def RS iffD1) THEN' rtac ctxt (mor_def RS iffD2) THEN' + REPEAT o etac ctxt conjE THEN' + rtac ctxt conjI THEN' CONJ_WRAP' (K fbetw_tac) set_maps THEN' CONJ_WRAP' mor_tac (set_maps ~~ map_comp_ids)) 1 end; -fun mk_mor_str_tac ks mor_def = - (rtac (mor_def RS iffD2) THEN' rtac conjI THEN' - CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN' - CONJ_WRAP' (K (EVERY' [rtac ballI, rtac refl])) ks) 1; +fun mk_mor_str_tac ctxt ks mor_def = + (rtac ctxt (mor_def RS iffD2) THEN' rtac ctxt conjI THEN' + CONJ_WRAP' (K (EVERY' [rtac ctxt ballI, rtac ctxt UNIV_I])) ks THEN' + CONJ_WRAP' (K (EVERY' [rtac ctxt ballI, rtac ctxt refl])) ks) 1; -fun mk_mor_convol_tac ks mor_def = - (rtac (mor_def RS iffD2) THEN' rtac conjI THEN' - CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN' - CONJ_WRAP' (K (EVERY' [rtac ballI, rtac trans, rtac @{thm fst_convol'}, rtac o_apply])) ks) 1; +fun mk_mor_convol_tac ctxt ks mor_def = + (rtac ctxt (mor_def RS iffD2) THEN' rtac ctxt conjI THEN' + CONJ_WRAP' (K (EVERY' [rtac ctxt ballI, rtac ctxt UNIV_I])) ks THEN' + CONJ_WRAP' (K (EVERY' [rtac ctxt ballI, rtac ctxt trans, rtac ctxt @{thm fst_convol'}, rtac ctxt o_apply])) ks) 1; -fun mk_mor_UNIV_tac m morEs mor_def = +fun mk_mor_UNIV_tac ctxt m morEs mor_def = let val n = length morEs; - fun mor_tac morE = EVERY' [rtac @{thm ext}, rtac trans, rtac o_apply, rtac trans, etac morE, - rtac CollectI, CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto m + n), - rtac sym, rtac o_apply]; + fun mor_tac morE = EVERY' [rtac ctxt @{thm ext}, rtac ctxt trans, rtac ctxt o_apply, rtac ctxt trans, etac ctxt morE, + rtac ctxt CollectI, CONJ_WRAP' (K (rtac ctxt subset_UNIV)) (1 upto m + n), + rtac ctxt sym, rtac ctxt o_apply]; in - EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs, - rtac (mor_def RS iffD2), rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs, - REPEAT_DETERM o etac conjE, REPEAT_DETERM_N n o dtac (@{thm fun_eq_iff} RS iffD1), - CONJ_WRAP' (K (EVERY' [rtac ballI, REPEAT_DETERM o etac allE, rtac trans, - etac (o_apply RS sym RS trans), rtac o_apply])) morEs] 1 + EVERY' [rtac ctxt iffI, CONJ_WRAP' mor_tac morEs, + rtac ctxt (mor_def RS iffD2), rtac ctxt conjI, CONJ_WRAP' (K (rtac ctxt ballI THEN' rtac ctxt UNIV_I)) morEs, + REPEAT_DETERM o etac ctxt conjE, REPEAT_DETERM_N n o dtac ctxt (@{thm fun_eq_iff} RS iffD1), + CONJ_WRAP' (K (EVERY' [rtac ctxt ballI, REPEAT_DETERM o etac ctxt allE, rtac ctxt trans, + etac ctxt (o_apply RS sym RS trans), rtac ctxt o_apply])) morEs] 1 end; -fun mk_copy_tac m alg_def mor_def alg_sets set_mapss = +fun mk_copy_tac ctxt m alg_def mor_def alg_sets set_mapss = let val n = length alg_sets; fun set_tac thm = - EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans, etac @{thm image_mono}, - rtac equalityD1, etac @{thm bij_betw_imageE}]; + EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt thm, rtac ctxt subset_trans, etac ctxt @{thm image_mono}, + rtac ctxt equalityD1, etac ctxt @{thm bij_betw_imageE}]; val alg_tac = CONJ_WRAP' (fn (set_maps, alg_set) => - EVERY' [rtac ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac set_mp, - rtac equalityD1, etac @{thm bij_betw_imageE[OF bij_betw_the_inv_into]}, - rtac imageI, etac alg_set, EVERY' (map set_tac (drop m set_maps))]) + EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac ctxt set_mp, + rtac ctxt equalityD1, etac ctxt @{thm bij_betw_imageE[OF bij_betw_the_inv_into]}, + rtac ctxt imageI, etac ctxt alg_set, EVERY' (map set_tac (drop m set_maps))]) (set_mapss ~~ alg_sets); - val mor_tac = rtac conjI THEN' CONJ_WRAP' (K (etac @{thm bij_betwE})) alg_sets THEN' + val mor_tac = rtac ctxt conjI THEN' CONJ_WRAP' (K (etac ctxt @{thm bij_betwE})) alg_sets THEN' CONJ_WRAP' (fn (set_maps, alg_set) => - EVERY' [rtac ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], - etac @{thm f_the_inv_into_f_bij_betw}, etac alg_set, + EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], + etac ctxt @{thm f_the_inv_into_f_bij_betw}, etac ctxt alg_set, EVERY' (map set_tac (drop m set_maps))]) (set_mapss ~~ alg_sets); in - (REPEAT_DETERM_N n o rtac exI THEN' rtac conjI THEN' - rtac (alg_def RS iffD2) THEN' alg_tac THEN' rtac (mor_def RS iffD2) THEN' mor_tac) 1 + (REPEAT_DETERM_N n o rtac ctxt exI THEN' rtac ctxt conjI THEN' + rtac ctxt (alg_def RS iffD2) THEN' alg_tac THEN' rtac ctxt (mor_def RS iffD2) THEN' mor_tac) 1 end; -fun mk_bd_limit_tac n bd_Cinfinite = - EVERY' [REPEAT_DETERM o etac conjE, rtac rev_mp, rtac @{thm Cinfinite_limit_finite}, - REPEAT_DETERM_N n o rtac @{thm finite.insertI}, rtac @{thm finite.emptyI}, - REPEAT_DETERM_N n o etac @{thm insert_subsetI}, rtac @{thm empty_subsetI}, - rtac bd_Cinfinite, rtac impI, etac bexE, rtac bexI, +fun mk_bd_limit_tac ctxt n bd_Cinfinite = + EVERY' [REPEAT_DETERM o etac ctxt conjE, rtac ctxt rev_mp, rtac ctxt @{thm Cinfinite_limit_finite}, + REPEAT_DETERM_N n o rtac ctxt @{thm finite.insertI}, rtac ctxt @{thm finite.emptyI}, + REPEAT_DETERM_N n o etac ctxt @{thm insert_subsetI}, rtac ctxt @{thm empty_subsetI}, + rtac ctxt bd_Cinfinite, rtac ctxt impI, etac ctxt bexE, rtac ctxt bexI, CONJ_WRAP' (fn i => - EVERY' [etac bspec, REPEAT_DETERM_N i o rtac @{thm insertI2}, rtac @{thm insertI1}]) + EVERY' [etac ctxt bspec, REPEAT_DETERM_N i o rtac ctxt @{thm insertI2}, rtac ctxt @{thm insertI1}]) (0 upto n - 1), atac] 1; -fun mk_min_algs_tac worel in_congs = +fun mk_min_algs_tac ctxt worel in_congs = let - val minG_tac = EVERY' [rtac @{thm SUP_cong}, rtac refl, dtac bspec, atac, etac arg_cong]; + val minG_tac = EVERY' [rtac ctxt @{thm SUP_cong}, rtac ctxt refl, dtac ctxt bspec, atac, etac ctxt arg_cong]; fun minH_tac thm = - EVERY' [rtac Un_cong, minG_tac, rtac @{thm image_cong}, rtac thm, - REPEAT_DETERM_N (length in_congs) o minG_tac, rtac refl]; + EVERY' [rtac ctxt Un_cong, minG_tac, rtac ctxt @{thm image_cong}, rtac ctxt thm, + REPEAT_DETERM_N (length in_congs) o minG_tac, rtac ctxt refl]; in - (rtac (worel RS (@{thm wo_rel.worec_fixpoint} RS fun_cong)) THEN' rtac iffD2 THEN' - rtac meta_eq_to_obj_eq THEN' rtac (worel RS @{thm wo_rel.adm_wo_def}) THEN' - REPEAT_DETERM_N 3 o rtac allI THEN' rtac impI THEN' - CONJ_WRAP_GEN' (EVERY' [rtac prod_injectI, rtac conjI]) minH_tac in_congs) 1 + (rtac ctxt (worel RS (@{thm wo_rel.worec_fixpoint} RS fun_cong)) THEN' rtac ctxt iffD2 THEN' + rtac ctxt meta_eq_to_obj_eq THEN' rtac ctxt (worel RS @{thm wo_rel.adm_wo_def}) THEN' + REPEAT_DETERM_N 3 o rtac ctxt allI THEN' rtac ctxt impI THEN' + CONJ_WRAP_GEN' (EVERY' [rtac ctxt prod_injectI, rtac ctxt conjI]) minH_tac in_congs) 1 end; -fun mk_min_algs_mono_tac ctxt min_algs = EVERY' [rtac relChainD, rtac allI, rtac allI, rtac impI, - rtac @{thm case_split}, rtac @{thm xt1(3)}, rtac min_algs, etac @{thm FieldI2}, rtac subsetI, - rtac UnI1, rtac @{thm UN_I}, etac @{thm underS_I}, atac, atac, rtac equalityD1, - dtac @{thm notnotD}, hyp_subst_tac ctxt, rtac refl] 1; +fun mk_min_algs_mono_tac ctxt min_algs = EVERY' [rtac ctxt relChainD, rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, + rtac ctxt @{thm case_split}, rtac ctxt @{thm xt1(3)}, rtac ctxt min_algs, etac ctxt @{thm FieldI2}, rtac ctxt subsetI, + rtac ctxt UnI1, rtac ctxt @{thm UN_I}, etac ctxt @{thm underS_I}, atac, atac, rtac ctxt equalityD1, + dtac ctxt @{thm notnotD}, hyp_subst_tac ctxt, rtac ctxt refl] 1; -fun mk_min_algs_card_of_tac cT ct m worel min_algs in_bds bd_Card_order bd_Cnotzero +fun mk_min_algs_card_of_tac ctxt cT ct m worel min_algs in_bds bd_Card_order bd_Cnotzero suc_Card_order suc_Cinfinite suc_Cnotzero suc_Asuc Asuc_Cinfinite = let val induct = worel RS @@ -222,121 +223,121 @@ val src = 1 upto m + 1; val dest = (m + 1) :: (1 upto m); val absorbAs_tac = if m = 0 then K (all_tac) - else EVERY' [rtac @{thm ordIso_transitive}, rtac @{thm csum_cong1}, - rtac @{thm ordIso_transitive}, - BNF_Tactics.mk_rotate_eq_tac (rtac @{thm ordIso_refl} THEN' - FIRST' [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum}, - rtac @{thm Card_order_cexp}]) + else EVERY' [rtac ctxt @{thm ordIso_transitive}, rtac ctxt @{thm csum_cong1}, + rtac ctxt @{thm ordIso_transitive}, + BNF_Tactics.mk_rotate_eq_tac ctxt (rtac ctxt @{thm ordIso_refl} THEN' + FIRST' [rtac ctxt @{thm card_of_Card_order}, rtac ctxt @{thm Card_order_csum}, + rtac ctxt @{thm Card_order_cexp}]) @{thm ordIso_transitive} @{thm csum_assoc} @{thm csum_com} @{thm csum_cong} src dest, - rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac ctrans, rtac @{thm ordLeq_csum1}, - FIRST' [rtac @{thm Card_order_csum}, rtac @{thm card_of_Card_order}], - rtac @{thm ordLeq_cexp1}, rtac suc_Cnotzero, rtac @{thm Card_order_csum}]; + rtac ctxt @{thm csum_absorb1}, rtac ctxt Asuc_Cinfinite, rtac ctxt ctrans, rtac ctxt @{thm ordLeq_csum1}, + FIRST' [rtac ctxt @{thm Card_order_csum}, rtac ctxt @{thm card_of_Card_order}], + rtac ctxt @{thm ordLeq_cexp1}, rtac ctxt suc_Cnotzero, rtac ctxt @{thm Card_order_csum}]; - val minG_tac = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac @{thm ordLess_imp_ordLeq}, - rtac @{thm ordLess_transitive}, rtac @{thm card_of_underS}, rtac suc_Card_order, - atac, rtac suc_Asuc, rtac ballI, etac allE, dtac mp, etac @{thm underS_E}, - dtac mp, etac @{thm underS_Field}, REPEAT o etac conjE, atac, rtac Asuc_Cinfinite] + val minG_tac = EVERY' [rtac ctxt @{thm UNION_Cinfinite_bound}, rtac ctxt @{thm ordLess_imp_ordLeq}, + rtac ctxt @{thm ordLess_transitive}, rtac ctxt @{thm card_of_underS}, rtac ctxt suc_Card_order, + atac, rtac ctxt suc_Asuc, rtac ctxt ballI, etac ctxt allE, dtac ctxt mp, etac ctxt @{thm underS_E}, + dtac ctxt mp, etac ctxt @{thm underS_Field}, REPEAT o etac ctxt conjE, atac, rtac ctxt Asuc_Cinfinite] - fun mk_minH_tac (min_alg, in_bd) = EVERY' [rtac @{thm ordIso_ordLeq_trans}, - rtac @{thm card_of_ordIso_subst}, etac min_alg, rtac @{thm Un_Cinfinite_bound}, - minG_tac, rtac ctrans, rtac @{thm card_of_image}, rtac ctrans, rtac in_bd, rtac ctrans, - rtac @{thm cexp_mono1}, rtac @{thm csum_mono1}, - REPEAT_DETERM_N m o rtac @{thm csum_mono2}, - CONJ_WRAP_GEN' (rtac @{thm csum_cinfinite_bound}) (K minG_tac) min_algs, + fun mk_minH_tac (min_alg, in_bd) = EVERY' [rtac ctxt @{thm ordIso_ordLeq_trans}, + rtac ctxt @{thm card_of_ordIso_subst}, etac ctxt min_alg, rtac ctxt @{thm Un_Cinfinite_bound}, + minG_tac, rtac ctxt ctrans, rtac ctxt @{thm card_of_image}, rtac ctxt ctrans, rtac ctxt in_bd, rtac ctxt ctrans, + rtac ctxt @{thm cexp_mono1}, rtac ctxt @{thm csum_mono1}, + REPEAT_DETERM_N m o rtac ctxt @{thm csum_mono2}, + CONJ_WRAP_GEN' (rtac ctxt @{thm csum_cinfinite_bound}) (K minG_tac) min_algs, REPEAT_DETERM o FIRST' - [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum}, - rtac Asuc_Cinfinite, rtac bd_Card_order], - rtac @{thm ordIso_ordLeq_trans}, rtac @{thm cexp_cong1}, absorbAs_tac, - rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac @{thm ctwo_ordLeq_Cinfinite}, - rtac Asuc_Cinfinite, rtac bd_Card_order, - rtac @{thm ordIso_imp_ordLeq}, rtac @{thm cexp_cprod_ordLeq}, - resolve0_tac @{thms Card_order_csum Card_order_ctwo}, rtac suc_Cinfinite, - rtac bd_Cnotzero, rtac @{thm cardSuc_ordLeq}, rtac bd_Card_order, rtac Asuc_Cinfinite]; + [rtac ctxt @{thm card_of_Card_order}, rtac ctxt @{thm Card_order_csum}, + rtac ctxt Asuc_Cinfinite, rtac ctxt bd_Card_order], + rtac ctxt @{thm ordIso_ordLeq_trans}, rtac ctxt @{thm cexp_cong1}, absorbAs_tac, + rtac ctxt @{thm csum_absorb1}, rtac ctxt Asuc_Cinfinite, rtac ctxt @{thm ctwo_ordLeq_Cinfinite}, + rtac ctxt Asuc_Cinfinite, rtac ctxt bd_Card_order, + rtac ctxt @{thm ordIso_imp_ordLeq}, rtac ctxt @{thm cexp_cprod_ordLeq}, + resolve0_tac @{thms Card_order_csum Card_order_ctwo}, rtac ctxt suc_Cinfinite, + rtac ctxt bd_Cnotzero, rtac ctxt @{thm cardSuc_ordLeq}, rtac ctxt bd_Card_order, rtac ctxt Asuc_Cinfinite]; in - (rtac induct THEN' - rtac impI THEN' + (rtac ctxt induct THEN' + rtac ctxt impI THEN' CONJ_WRAP' mk_minH_tac (min_algs ~~ in_bds)) 1 end; -fun mk_min_algs_least_tac cT ct worel min_algs alg_sets = +fun mk_min_algs_least_tac ctxt cT ct worel min_algs alg_sets = let val induct = worel RS Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp}; - val minG_tac = EVERY' [rtac @{thm UN_least}, etac allE, dtac mp, etac @{thm underS_E}, - dtac mp, etac @{thm underS_Field}, REPEAT_DETERM o etac conjE, atac]; + val minG_tac = EVERY' [rtac ctxt @{thm UN_least}, etac ctxt allE, dtac ctxt mp, etac ctxt @{thm underS_E}, + dtac ctxt mp, etac ctxt @{thm underS_Field}, REPEAT_DETERM o etac ctxt conjE, atac]; - fun mk_minH_tac (min_alg, alg_set) = EVERY' [rtac ord_eq_le_trans, etac min_alg, - rtac @{thm Un_least}, minG_tac, rtac @{thm image_subsetI}, - REPEAT_DETERM o eresolve0_tac [CollectE, conjE], etac alg_set, - REPEAT_DETERM o (etac subset_trans THEN' minG_tac)]; + fun mk_minH_tac (min_alg, alg_set) = EVERY' [rtac ctxt ord_eq_le_trans, etac ctxt min_alg, + rtac ctxt @{thm Un_least}, minG_tac, rtac ctxt @{thm image_subsetI}, + REPEAT_DETERM o eresolve0_tac [CollectE, conjE], etac ctxt alg_set, + REPEAT_DETERM o (etac ctxt subset_trans THEN' minG_tac)]; in - (rtac induct THEN' - rtac impI THEN' + (rtac ctxt induct THEN' + rtac ctxt impI THEN' CONJ_WRAP' mk_minH_tac (min_algs ~~ alg_sets)) 1 end; -fun mk_alg_min_alg_tac m alg_def min_alg_defs bd_limit bd_Cinfinite +fun mk_alg_min_alg_tac ctxt m alg_def min_alg_defs bd_limit bd_Cinfinite set_bdss min_algs min_alg_monos = let val n = length min_algs; fun mk_cardSuc_UNION_tac set_bds (mono, def) = EVERY' - [rtac bexE, rtac @{thm cardSuc_UNION_Cinfinite}, rtac bd_Cinfinite, rtac mono, - etac (def RSN (2, @{thm subset_trans[OF _ equalityD1]})), resolve0_tac set_bds]; + [rtac ctxt bexE, rtac ctxt @{thm cardSuc_UNION_Cinfinite}, rtac ctxt bd_Cinfinite, rtac ctxt mono, + etac ctxt (def RSN (2, @{thm subset_trans[OF _ equalityD1]})), resolve0_tac set_bds]; fun mk_conjunct_tac (set_bds, (min_alg, min_alg_def)) = - EVERY' [rtac ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], - EVERY' (map (mk_cardSuc_UNION_tac set_bds) (min_alg_monos ~~ min_alg_defs)), rtac bexE, - rtac bd_limit, REPEAT_DETERM_N (n - 1) o etac conjI, atac, - rtac (min_alg_def RS @{thm set_mp[OF equalityD2]}), - rtac @{thm UN_I}, REPEAT_DETERM_N (m + 3 * n) o etac thin_rl, atac, rtac set_mp, - rtac equalityD2, rtac min_alg, atac, rtac UnI2, rtac @{thm image_eqI}, rtac refl, - rtac CollectI, REPEAT_DETERM_N m o dtac asm_rl, REPEAT_DETERM_N n o etac thin_rl, - REPEAT_DETERM o etac conjE, + EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], + EVERY' (map (mk_cardSuc_UNION_tac set_bds) (min_alg_monos ~~ min_alg_defs)), rtac ctxt bexE, + rtac ctxt bd_limit, REPEAT_DETERM_N (n - 1) o etac ctxt conjI, atac, + rtac ctxt (min_alg_def RS @{thm set_mp[OF equalityD2]}), + rtac ctxt @{thm UN_I}, REPEAT_DETERM_N (m + 3 * n) o etac ctxt thin_rl, atac, rtac ctxt set_mp, + rtac ctxt equalityD2, rtac ctxt min_alg, atac, rtac ctxt UnI2, rtac ctxt @{thm image_eqI}, rtac ctxt refl, + rtac ctxt CollectI, REPEAT_DETERM_N m o dtac ctxt asm_rl, REPEAT_DETERM_N n o etac ctxt thin_rl, + REPEAT_DETERM o etac ctxt conjE, CONJ_WRAP' (K (FIRST' [atac, - EVERY' [etac subset_trans, rtac subsetI, rtac @{thm UN_I}, - etac @{thm underS_I}, atac, atac]])) + EVERY' [etac ctxt subset_trans, rtac ctxt subsetI, rtac ctxt @{thm UN_I}, + etac ctxt @{thm underS_I}, atac, atac]])) set_bds]; in - (rtac (alg_def RS iffD2) THEN' + (rtac ctxt (alg_def RS iffD2) THEN' CONJ_WRAP' mk_conjunct_tac (set_bdss ~~ (min_algs ~~ min_alg_defs))) 1 end; -fun mk_card_of_min_alg_tac min_alg_def card_of suc_Card_order suc_Asuc Asuc_Cinfinite = - EVERY' [rtac @{thm ordIso_ordLeq_trans}, rtac (min_alg_def RS @{thm card_of_ordIso_subst}), - rtac @{thm UNION_Cinfinite_bound}, rtac @{thm ordIso_ordLeq_trans}, - rtac @{thm card_of_Field_ordIso}, rtac suc_Card_order, rtac @{thm ordLess_imp_ordLeq}, - rtac suc_Asuc, rtac ballI, dtac rev_mp, rtac card_of, REPEAT_DETERM o etac conjE, atac, - rtac Asuc_Cinfinite] 1; +fun mk_card_of_min_alg_tac ctxt min_alg_def card_of suc_Card_order suc_Asuc Asuc_Cinfinite = + EVERY' [rtac ctxt @{thm ordIso_ordLeq_trans}, rtac ctxt (min_alg_def RS @{thm card_of_ordIso_subst}), + rtac ctxt @{thm UNION_Cinfinite_bound}, rtac ctxt @{thm ordIso_ordLeq_trans}, + rtac ctxt @{thm card_of_Field_ordIso}, rtac ctxt suc_Card_order, rtac ctxt @{thm ordLess_imp_ordLeq}, + rtac ctxt suc_Asuc, rtac ctxt ballI, dtac ctxt rev_mp, rtac ctxt card_of, REPEAT_DETERM o etac ctxt conjE, atac, + rtac ctxt Asuc_Cinfinite] 1; -fun mk_least_min_alg_tac min_alg_def least = - EVERY' [rtac (min_alg_def RS ord_eq_le_trans), rtac @{thm UN_least}, dtac least, dtac mp, atac, - REPEAT_DETERM o etac conjE, atac] 1; +fun mk_least_min_alg_tac ctxt min_alg_def least = + EVERY' [rtac ctxt (min_alg_def RS ord_eq_le_trans), rtac ctxt @{thm UN_least}, dtac ctxt least, dtac ctxt mp, atac, + REPEAT_DETERM o etac ctxt conjE, atac] 1; fun mk_alg_select_tac ctxt Abs_inverse = - EVERY' [rtac ballI, + EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], hyp_subst_tac ctxt] 1 THEN unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs) THEN atac 1; -fun mk_mor_select_tac mor_def mor_cong mor_comp mor_incl_min_alg alg_def alg_select alg_sets +fun mk_mor_select_tac ctxt mor_def mor_cong mor_comp mor_incl_min_alg alg_def alg_select alg_sets set_maps str_init_defs = let val n = length alg_sets; val fbetw_tac = - CONJ_WRAP' (K (EVERY' [rtac ballI, etac rev_bspec, etac CollectE, atac])) alg_sets; + CONJ_WRAP' (K (EVERY' [rtac ctxt ballI, etac ctxt rev_bspec, etac ctxt CollectE, atac])) alg_sets; val mor_tac = - CONJ_WRAP' (fn thm => EVERY' [rtac ballI, rtac thm]) str_init_defs; + CONJ_WRAP' (fn thm => EVERY' [rtac ctxt ballI, rtac ctxt thm]) str_init_defs; fun alg_epi_tac ((alg_set, str_init_def), set_map) = - EVERY' [rtac ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac CollectI, - rtac ballI, ftac (alg_select RS bspec), rtac (str_init_def RS @{thm ssubst_mem}), - etac alg_set, REPEAT_DETERM o EVERY' [rtac ord_eq_le_trans, resolve0_tac set_map, - rtac subset_trans, etac @{thm image_mono}, rtac @{thm image_Collect_subsetI}, etac bspec, + EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac ctxt CollectI, + rtac ctxt ballI, ftac ctxt (alg_select RS bspec), rtac ctxt (str_init_def RS @{thm ssubst_mem}), + etac ctxt alg_set, REPEAT_DETERM o EVERY' [rtac ctxt ord_eq_le_trans, resolve0_tac set_map, + rtac ctxt subset_trans, etac ctxt @{thm image_mono}, rtac ctxt @{thm image_Collect_subsetI}, etac ctxt bspec, atac]]; in - EVERY' [rtac mor_cong, REPEAT_DETERM_N n o (rtac sym THEN' rtac @{thm comp_id}), - rtac (Thm.permute_prems 0 1 mor_comp), etac (Thm.permute_prems 0 1 mor_comp), - rtac (mor_def RS iffD2), rtac conjI, fbetw_tac, mor_tac, rtac mor_incl_min_alg, - rtac (alg_def RS iffD2), CONJ_WRAP' alg_epi_tac ((alg_sets ~~ str_init_defs) ~~ set_maps)] 1 + EVERY' [rtac ctxt mor_cong, REPEAT_DETERM_N n o (rtac ctxt sym THEN' rtac ctxt @{thm comp_id}), + rtac ctxt (Thm.permute_prems 0 1 mor_comp), etac ctxt (Thm.permute_prems 0 1 mor_comp), + rtac ctxt (mor_def RS iffD2), rtac ctxt conjI, fbetw_tac, mor_tac, rtac ctxt mor_incl_min_alg, + rtac ctxt (alg_def RS iffD2), CONJ_WRAP' alg_epi_tac ((alg_sets ~~ str_init_defs) ~~ set_maps)] 1 end; fun mk_init_ex_mor_tac ctxt Abs_inverse copy card_of_min_algs mor_Rep mor_comp mor_select mor_incl = @@ -344,59 +345,59 @@ val n = length card_of_min_algs; in EVERY' [Method.insert_tac (map (fn thm => thm RS @{thm ex_bij_betw}) card_of_min_algs), - REPEAT_DETERM o etac exE, rtac rev_mp, rtac copy, REPEAT_DETERM_N n o atac, - rtac impI, REPEAT_DETERM o eresolve0_tac [exE, conjE], REPEAT_DETERM_N n o rtac exI, - rtac mor_comp, rtac mor_Rep, rtac mor_select, rtac CollectI, REPEAT_DETERM o rtac exI, - rtac conjI, rtac refl, atac, + REPEAT_DETERM o etac ctxt exE, rtac ctxt rev_mp, rtac ctxt copy, REPEAT_DETERM_N n o atac, + rtac ctxt impI, REPEAT_DETERM o eresolve0_tac [exE, conjE], REPEAT_DETERM_N n o rtac ctxt exI, + rtac ctxt mor_comp, rtac ctxt mor_Rep, rtac ctxt mor_select, rtac ctxt CollectI, REPEAT_DETERM o rtac ctxt exI, + rtac ctxt conjI, rtac ctxt refl, atac, SELECT_GOAL (unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs)), - etac mor_comp, rtac mor_incl, REPEAT_DETERM_N n o rtac subset_UNIV] 1 + etac ctxt mor_comp, rtac ctxt mor_incl, REPEAT_DETERM_N n o rtac ctxt subset_UNIV] 1 end; -fun mk_init_unique_mor_tac cts m +fun mk_init_unique_mor_tac ctxt cts m alg_def alg_min_alg least_min_algs in_monos alg_sets morEs map_cong0s = let val n = length least_min_algs; val ks = (1 upto n); - fun mor_tac morE in_mono = EVERY' [etac morE, rtac set_mp, rtac in_mono, - REPEAT_DETERM_N n o rtac @{thm Collect_restrict}, rtac CollectI, - REPEAT_DETERM_N (m + n) o (TRY o rtac conjI THEN' atac)]; + fun mor_tac morE in_mono = EVERY' [etac ctxt morE, rtac ctxt set_mp, rtac ctxt in_mono, + REPEAT_DETERM_N n o rtac ctxt @{thm Collect_restrict}, rtac ctxt CollectI, + REPEAT_DETERM_N (m + n) o (TRY o rtac ctxt conjI THEN' atac)]; fun cong_tac ct map_cong0 = EVERY' - [rtac (map_cong0 RS cterm_instantiate_pos [NONE, NONE, SOME ct] arg_cong), - REPEAT_DETERM_N m o rtac refl, - REPEAT_DETERM_N n o (etac @{thm prop_restrict} THEN' atac)]; + [rtac ctxt (map_cong0 RS cterm_instantiate_pos [NONE, NONE, SOME ct] arg_cong), + REPEAT_DETERM_N m o rtac ctxt refl, + REPEAT_DETERM_N n o (etac ctxt @{thm prop_restrict} THEN' atac)]; fun mk_alg_tac (ct, (alg_set, (in_mono, (morE, map_cong0)))) = - EVERY' [rtac ballI, rtac CollectI, - REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set), - REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}), - rtac trans, mor_tac morE in_mono, - rtac trans, cong_tac ct map_cong0, - rtac sym, mor_tac morE in_mono]; + EVERY' [rtac ctxt ballI, rtac ctxt CollectI, + REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac ctxt conjI, rtac ctxt (alg_min_alg RS alg_set), + REPEAT_DETERM_N n o (etac ctxt subset_trans THEN' rtac ctxt @{thm Collect_restrict}), + rtac ctxt trans, mor_tac morE in_mono, + rtac ctxt trans, cong_tac ct map_cong0, + rtac ctxt sym, mor_tac morE in_mono]; fun mk_unique_tac (k, least_min_alg) = - select_prem_tac n (etac @{thm prop_restrict}) k THEN' rtac least_min_alg THEN' - rtac (alg_def RS iffD2) THEN' + select_prem_tac ctxt n (etac ctxt @{thm prop_restrict}) k THEN' rtac ctxt least_min_alg THEN' + rtac ctxt (alg_def RS iffD2) THEN' CONJ_WRAP' mk_alg_tac (cts ~~ (alg_sets ~~ (in_monos ~~ (morEs ~~ map_cong0s)))); in CONJ_WRAP' mk_unique_tac (ks ~~ least_min_algs) 1 end; -fun mk_init_induct_tac m alg_def alg_min_alg least_min_algs alg_sets = +fun mk_init_induct_tac ctxt m alg_def alg_min_alg least_min_algs alg_sets = let val n = length least_min_algs; - fun mk_alg_tac alg_set = EVERY' [rtac ballI, rtac CollectI, - REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set), - REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}), - rtac mp, etac bspec, rtac CollectI, - REPEAT_DETERM_N m o (rtac conjI THEN' atac), - CONJ_WRAP' (K (etac subset_trans THEN' rtac @{thm Collect_restrict})) alg_sets, - CONJ_WRAP' (K (rtac ballI THEN' etac @{thm prop_restrict} THEN' atac)) alg_sets]; + fun mk_alg_tac alg_set = EVERY' [rtac ctxt ballI, rtac ctxt CollectI, + REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac ctxt conjI, rtac ctxt (alg_min_alg RS alg_set), + REPEAT_DETERM_N n o (etac ctxt subset_trans THEN' rtac ctxt @{thm Collect_restrict}), + rtac ctxt mp, etac ctxt bspec, rtac ctxt CollectI, + REPEAT_DETERM_N m o (rtac ctxt conjI THEN' atac), + CONJ_WRAP' (K (etac ctxt subset_trans THEN' rtac ctxt @{thm Collect_restrict})) alg_sets, + CONJ_WRAP' (K (rtac ctxt ballI THEN' etac ctxt @{thm prop_restrict} THEN' atac)) alg_sets]; fun mk_induct_tac least_min_alg = - rtac ballI THEN' etac @{thm prop_restrict} THEN' rtac least_min_alg THEN' - rtac (alg_def RS iffD2) THEN' + rtac ctxt ballI THEN' etac ctxt @{thm prop_restrict} THEN' rtac ctxt least_min_alg THEN' + rtac ctxt (alg_def RS iffD2) THEN' CONJ_WRAP' mk_alg_tac alg_sets; in CONJ_WRAP' mk_induct_tac least_min_algs 1 @@ -404,61 +405,61 @@ fun mk_mor_Rep_tac ctxt m defs Reps Abs_inverses alg_min_alg alg_sets set_mapss = unfold_thms_tac ctxt (@{thm o_apply} :: defs) THEN - EVERY' [rtac conjI, - CONJ_WRAP' (fn thm => rtac ballI THEN' rtac thm) Reps, + EVERY' [rtac ctxt conjI, + CONJ_WRAP' (fn thm => rtac ctxt ballI THEN' rtac ctxt thm) Reps, CONJ_WRAP' (fn (Abs_inverse, (set_maps, alg_set)) => - EVERY' [rtac ballI, rtac Abs_inverse, rtac (alg_min_alg RS alg_set), + EVERY' [rtac ctxt ballI, rtac ctxt Abs_inverse, rtac ctxt (alg_min_alg RS alg_set), EVERY' (map2 (fn Rep => fn set_map => - EVERY' [rtac (set_map RS ord_eq_le_trans), rtac @{thm image_subsetI}, rtac Rep]) + EVERY' [rtac ctxt (set_map RS ord_eq_le_trans), rtac ctxt @{thm image_subsetI}, rtac ctxt Rep]) Reps (drop m set_maps))]) (Abs_inverses ~~ (set_mapss ~~ alg_sets))] 1; fun mk_mor_Abs_tac ctxt cts defs Abs_inverses map_comp_ids map_congLs = unfold_thms_tac ctxt (@{thm o_apply} :: defs) THEN - EVERY' [rtac conjI, - CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) Abs_inverses, + EVERY' [rtac ctxt conjI, + CONJ_WRAP' (K (rtac ctxt ballI THEN' rtac ctxt UNIV_I)) Abs_inverses, CONJ_WRAP' (fn (ct, thm) => - EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], - rtac (thm RS (cterm_instantiate_pos [NONE, NONE, SOME ct] arg_cong) RS sym), + EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], + rtac ctxt (thm RS (cterm_instantiate_pos [NONE, NONE, SOME ct] arg_cong) RS sym), EVERY' (map (fn Abs_inverse => - EVERY' [rtac (o_apply RS trans RS ballI), etac (set_mp RS Abs_inverse), atac]) + EVERY' [rtac ctxt (o_apply RS trans RS ballI), etac ctxt (set_mp RS Abs_inverse), atac]) Abs_inverses)]) (cts ~~ map2 mk_trans map_comp_ids map_congLs)] 1;; fun mk_mor_fold_tac ctxt cT ct fold_defs ex_mor mor = - (EVERY' (map (stac ctxt) fold_defs) THEN' EVERY' [rtac rev_mp, rtac ex_mor, rtac impI] THEN' - REPEAT_DETERM_N (length fold_defs) o etac exE THEN' - rtac (Drule.instantiate' [SOME cT] [SOME ct] @{thm someI}) THEN' etac mor) 1; + (EVERY' (map (stac ctxt) fold_defs) THEN' EVERY' [rtac ctxt rev_mp, rtac ctxt ex_mor, rtac ctxt impI] THEN' + REPEAT_DETERM_N (length fold_defs) o etac ctxt exE THEN' + rtac ctxt (Drule.instantiate' [SOME cT] [SOME ct] @{thm someI}) THEN' etac ctxt mor) 1; -fun mk_fold_unique_mor_tac type_defs init_unique_mors Reps mor_comp mor_Abs mor_fold = +fun mk_fold_unique_mor_tac ctxt type_defs init_unique_mors Reps mor_comp mor_Abs mor_fold = let fun mk_unique type_def = - EVERY' [rtac @{thm surj_fun_eq}, rtac (type_def RS @{thm type_definition.Abs_image}), - rtac ballI, resolve0_tac init_unique_mors, - EVERY' (map (fn thm => atac ORELSE' rtac thm) Reps), - rtac mor_comp, rtac mor_Abs, atac, - rtac mor_comp, rtac mor_Abs, rtac mor_fold]; + EVERY' [rtac ctxt @{thm surj_fun_eq}, rtac ctxt (type_def RS @{thm type_definition.Abs_image}), + rtac ctxt ballI, resolve0_tac init_unique_mors, + EVERY' (map (fn thm => atac ORELSE' rtac ctxt thm) Reps), + rtac ctxt mor_comp, rtac ctxt mor_Abs, atac, + rtac ctxt mor_comp, rtac ctxt mor_Abs, rtac ctxt mor_fold]; in CONJ_WRAP' mk_unique type_defs 1 end; -fun mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_cong0L ctor_o_folds = - EVERY' [rtac @{thm ext}, rtac trans, rtac o_apply, rtac (dtor_def RS fun_cong RS trans), - rtac trans, rtac foldx, rtac trans, rtac map_comp_id, rtac trans, rtac map_cong0L, - EVERY' (map (fn thm => rtac ballI THEN' rtac (trans OF [thm RS fun_cong, id_apply])) +fun mk_dtor_o_ctor_tac ctxt dtor_def foldx map_comp_id map_cong0L ctor_o_folds = + EVERY' [rtac ctxt @{thm ext}, rtac ctxt trans, rtac ctxt o_apply, rtac ctxt (dtor_def RS fun_cong RS trans), + rtac ctxt trans, rtac ctxt foldx, rtac ctxt trans, rtac ctxt map_comp_id, rtac ctxt trans, rtac ctxt map_cong0L, + EVERY' (map (fn thm => rtac ctxt ballI THEN' rtac ctxt (trans OF [thm RS fun_cong, id_apply])) ctor_o_folds), - rtac sym, rtac id_apply] 1; + rtac ctxt sym, rtac ctxt id_apply] 1; fun mk_rec_tac ctxt rec_defs foldx fst_recs = unfold_thms_tac ctxt (rec_defs @ map (fn thm => thm RS @{thm convol_expand_snd}) fst_recs) THEN - EVERY' [rtac trans, rtac o_apply, rtac trans, rtac (foldx RS @{thm arg_cong[of _ _ snd]}), - rtac @{thm snd_convol'}] 1; + EVERY' [rtac ctxt trans, rtac ctxt o_apply, rtac ctxt trans, rtac ctxt (foldx RS @{thm arg_cong[of _ _ snd]}), + rtac ctxt @{thm snd_convol'}] 1; fun mk_rec_unique_mor_tac ctxt rec_defs fst_recs fold_unique_mor = unfold_thms_tac ctxt (rec_defs @ map (fn thm => thm RS @{thm convol_expand_snd'}) fst_recs) THEN - etac fold_unique_mor 1; + etac ctxt fold_unique_mor 1; fun mk_ctor_induct_tac ctxt m set_mapss init_induct morEs mor_Abs Rep_invs Abs_invs Reps = let @@ -466,22 +467,22 @@ val ks = 1 upto n; fun mk_IH_tac Rep_inv Abs_inv set_map = - DETERM o EVERY' [dtac meta_mp, rtac (Rep_inv RS arg_cong RS iffD1), etac bspec, - dtac set_rev_mp, rtac equalityD1, rtac set_map, etac imageE, - hyp_subst_tac ctxt, rtac (Abs_inv RS @{thm ssubst_mem}), etac set_mp, atac, atac]; + DETERM o EVERY' [dtac ctxt meta_mp, rtac ctxt (Rep_inv RS arg_cong RS iffD1), etac ctxt bspec, + dtac ctxt set_rev_mp, rtac ctxt equalityD1, rtac ctxt set_map, etac ctxt imageE, + hyp_subst_tac ctxt, rtac ctxt (Abs_inv RS @{thm ssubst_mem}), etac ctxt set_mp, atac, atac]; fun mk_closed_tac (k, (morE, set_maps)) = - EVERY' [select_prem_tac n (dtac asm_rl) k, rtac ballI, rtac impI, - rtac (mor_Abs RS morE RS arg_cong RS iffD2), atac, - REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], dtac @{thm meta_spec}, + EVERY' [select_prem_tac ctxt n (dtac ctxt asm_rl) k, rtac ctxt ballI, rtac ctxt impI, + rtac ctxt (mor_Abs RS morE RS arg_cong RS iffD2), atac, + REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], dtac ctxt @{thm meta_spec}, EVERY' (@{map 3} mk_IH_tac Rep_invs Abs_invs (drop m set_maps)), atac]; fun mk_induct_tac (Rep, Rep_inv) = - EVERY' [rtac (Rep_inv RS arg_cong RS iffD1), etac (Rep RSN (2, bspec))]; + EVERY' [rtac ctxt (Rep_inv RS arg_cong RS iffD1), etac ctxt (Rep RSN (2, bspec))]; in - (rtac mp THEN' rtac impI THEN' - DETERM o CONJ_WRAP_GEN' (etac conjE THEN' rtac conjI) mk_induct_tac (Reps ~~ Rep_invs) THEN' - rtac init_induct THEN' + (rtac ctxt mp THEN' rtac ctxt impI THEN' + DETERM o CONJ_WRAP_GEN' (etac ctxt conjE THEN' rtac ctxt conjI) mk_induct_tac (Reps ~~ Rep_invs) THEN' + rtac ctxt init_induct THEN' DETERM o CONJ_WRAP' mk_closed_tac (ks ~~ (morEs ~~ set_mapss))) 1 end; @@ -490,44 +491,44 @@ val n = length weak_ctor_inducts; val ks = 1 upto n; fun mk_inner_induct_tac induct i = - EVERY' [rtac allI, fo_rtac induct ctxt, - select_prem_tac n (dtac @{thm meta_spec2}) i, + EVERY' [rtac ctxt allI, fo_rtac ctxt induct, + select_prem_tac ctxt n (dtac ctxt @{thm meta_spec2}) i, REPEAT_DETERM_N n o - EVERY' [dtac meta_mp THEN_ALL_NEW Goal.norm_hhf_tac ctxt, - REPEAT_DETERM o dtac @{thm meta_spec}, etac (spec RS meta_mp), atac], + EVERY' [dtac ctxt meta_mp THEN_ALL_NEW Goal.norm_hhf_tac ctxt, + REPEAT_DETERM o dtac ctxt @{thm meta_spec}, etac ctxt (spec RS meta_mp), atac], atac]; in - EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts ctor_induct), - EVERY' (map2 mk_inner_induct_tac weak_ctor_inducts ks), rtac impI, + EVERY' [rtac ctxt rev_mp, rtac ctxt (Drule.instantiate' cTs cts ctor_induct), + EVERY' (map2 mk_inner_induct_tac weak_ctor_inducts ks), rtac ctxt impI, REPEAT_DETERM o eresolve_tac ctxt [conjE, allE], CONJ_WRAP' (K atac) ks] 1 end; -fun mk_map_tac m n foldx map_comp_id map_cong0 = - EVERY' [rtac @{thm ext}, rtac trans, rtac o_apply, rtac trans, rtac foldx, rtac trans, - rtac o_apply, - rtac trans, rtac (map_comp_id RS arg_cong), rtac trans, rtac (map_cong0 RS arg_cong), - REPEAT_DETERM_N m o rtac refl, - REPEAT_DETERM_N n o (EVERY' (map rtac [trans, o_apply, id_apply])), - rtac sym, rtac o_apply] 1; +fun mk_map_tac ctxt m n foldx map_comp_id map_cong0 = + EVERY' [rtac ctxt @{thm ext}, rtac ctxt trans, rtac ctxt o_apply, rtac ctxt trans, rtac ctxt foldx, rtac ctxt trans, + rtac ctxt o_apply, + rtac ctxt trans, rtac ctxt (map_comp_id RS arg_cong), rtac ctxt trans, rtac ctxt (map_cong0 RS arg_cong), + REPEAT_DETERM_N m o rtac ctxt refl, + REPEAT_DETERM_N n o (EVERY' (map (rtac ctxt) [trans, o_apply, id_apply])), + rtac ctxt sym, rtac ctxt o_apply] 1; fun mk_ctor_map_unique_tac ctxt fold_unique sym_map_comps = - rtac fold_unique 1 THEN + rtac ctxt fold_unique 1 THEN unfold_thms_tac ctxt (sym_map_comps @ @{thms comp_assoc id_comp comp_id}) THEN ALLGOALS atac; -fun mk_set_tac foldx = EVERY' [rtac @{thm ext}, rtac trans, rtac o_apply, - rtac trans, rtac foldx, rtac sym, rtac o_apply] 1; +fun mk_set_tac ctxt foldx = EVERY' [rtac ctxt @{thm ext}, rtac ctxt trans, rtac ctxt o_apply, + rtac ctxt trans, rtac ctxt foldx, rtac ctxt sym, rtac ctxt o_apply] 1; -fun mk_ctor_set_tac set set_map set_maps = +fun mk_ctor_set_tac ctxt set set_map set_maps = let val n = length set_maps; - fun mk_UN thm = rtac (thm RS @{thm arg_cong[of _ _ Union]} RS trans) THEN' - rtac @{thm Union_image_eq}; + fun mk_UN thm = rtac ctxt (thm RS @{thm arg_cong[of _ _ Union]} RS trans) THEN' + rtac ctxt @{thm Union_image_eq}; in - EVERY' [rtac (set RS @{thm comp_eq_dest} RS trans), rtac Un_cong, - rtac (trans OF [set_map, trans_fun_cong_image_id_id_apply]), - REPEAT_DETERM_N (n - 1) o rtac Un_cong, + EVERY' [rtac ctxt (set RS @{thm comp_eq_dest} RS trans), rtac ctxt Un_cong, + rtac ctxt (trans OF [set_map, trans_fun_cong_image_id_id_apply]), + REPEAT_DETERM_N (n - 1) o rtac ctxt Un_cong, EVERY' (map mk_UN set_maps)] 1 end; @@ -535,17 +536,17 @@ let val n = length ctor_maps; - fun useIH set_nat = EVERY' [rtac trans, rtac @{thm image_UN}, rtac trans, rtac @{thm SUP_cong}, - rtac refl, Goal.assume_rule_tac ctxt, rtac sym, rtac trans, rtac @{thm SUP_cong}, - rtac set_nat, rtac refl, rtac @{thm UN_simps(10)}]; + fun useIH set_nat = EVERY' [rtac ctxt trans, rtac ctxt @{thm image_UN}, rtac ctxt trans, rtac ctxt @{thm SUP_cong}, + rtac ctxt refl, Goal.assume_rule_tac ctxt, rtac ctxt sym, rtac ctxt trans, rtac ctxt @{thm SUP_cong}, + rtac ctxt set_nat, rtac ctxt refl, rtac ctxt @{thm UN_simps(10)}]; fun mk_set_nat cset ctor_map ctor_set set_nats = - EVERY' [rtac trans, rtac @{thm image_cong}, rtac ctor_set, rtac refl, rtac sym, - rtac (trans OF [ctor_map RS cterm_instantiate_pos [NONE, NONE, SOME cset] arg_cong, + EVERY' [rtac ctxt trans, rtac ctxt @{thm image_cong}, rtac ctxt ctor_set, rtac ctxt refl, rtac ctxt sym, + rtac ctxt (trans OF [ctor_map RS cterm_instantiate_pos [NONE, NONE, SOME cset] arg_cong, ctor_set RS trans]), - rtac sym, EVERY' (map rtac [trans, @{thm image_Un}, Un_cong]), - rtac sym, rtac (nth set_nats (i - 1)), - REPEAT_DETERM_N (n - 1) o EVERY' (map rtac [trans, @{thm image_Un}, Un_cong]), + rtac ctxt sym, EVERY' (map (rtac ctxt) [trans, @{thm image_Un}, Un_cong]), + rtac ctxt sym, rtac ctxt (nth set_nats (i - 1)), + REPEAT_DETERM_N (n - 1) o EVERY' (map (rtac ctxt) [trans, @{thm image_Un}, Un_cong]), EVERY' (map useIH (drop m set_nats))]; in (induct_tac THEN' EVERY' (@{map 4} mk_set_nat csets ctor_maps ctor_sets set_mapss)) 1 @@ -555,13 +556,13 @@ let val n = length ctor_sets; - fun useIH set_bd = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac set_bd, rtac ballI, - Goal.assume_rule_tac ctxt, rtac bd_Cinfinite]; + fun useIH set_bd = EVERY' [rtac ctxt @{thm UNION_Cinfinite_bound}, rtac ctxt set_bd, rtac ctxt ballI, + Goal.assume_rule_tac ctxt, rtac ctxt bd_Cinfinite]; fun mk_set_nat ctor_set set_bds = - EVERY' [rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_ordIso_subst}, rtac ctor_set, - rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_bds (i - 1)), - REPEAT_DETERM_N (n - 1) o rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), + EVERY' [rtac ctxt @{thm ordIso_ordLeq_trans}, rtac ctxt @{thm card_of_ordIso_subst}, rtac ctxt ctor_set, + rtac ctxt (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac ctxt (nth set_bds (i - 1)), + REPEAT_DETERM_N (n - 1) o rtac ctxt (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), EVERY' (map useIH (drop m set_bds))]; in (induct_tac THEN' EVERY' (map2 mk_set_nat ctor_sets set_bdss)) 1 @@ -569,72 +570,72 @@ fun mk_mcong_tac ctxt induct_tac set_setsss map_cong0s ctor_maps = let - fun use_asm thm = EVERY' [etac bspec, etac set_rev_mp, rtac thm]; + fun use_asm thm = EVERY' [etac ctxt bspec, etac ctxt set_rev_mp, rtac ctxt thm]; - fun useIH set_sets = EVERY' [rtac mp, Goal.assume_rule_tac ctxt, + fun useIH set_sets = EVERY' [rtac ctxt mp, Goal.assume_rule_tac ctxt, CONJ_WRAP' (fn thm => - EVERY' [rtac ballI, etac bspec, etac set_rev_mp, etac thm]) set_sets]; + EVERY' [rtac ctxt ballI, etac ctxt bspec, etac ctxt set_rev_mp, etac ctxt thm]) set_sets]; fun mk_map_cong0 ctor_map map_cong0 set_setss = - EVERY' [rtac impI, REPEAT_DETERM o etac conjE, - rtac trans, rtac ctor_map, rtac trans, rtac (map_cong0 RS arg_cong), + EVERY' [rtac ctxt impI, REPEAT_DETERM o etac ctxt conjE, + rtac ctxt trans, rtac ctxt ctor_map, rtac ctxt trans, rtac ctxt (map_cong0 RS arg_cong), EVERY' (map use_asm (map hd set_setss)), EVERY' (map useIH (transpose (map tl set_setss))), - rtac sym, rtac ctor_map]; + rtac ctxt sym, rtac ctxt ctor_map]; in (induct_tac THEN' EVERY' (@{map 3} mk_map_cong0 ctor_maps map_cong0s set_setsss)) 1 end; fun mk_le_rel_OO_tac ctxt m induct ctor_nchotomys ctor_Irels rel_mono_strong0s le_rel_OOs = - EVERY' (rtac induct :: + EVERY' (rtac ctxt induct :: @{map 4} (fn nchotomy => fn Irel => fn rel_mono => fn le_rel_OO => - EVERY' [rtac impI, etac (nchotomy RS @{thm nchotomy_relcomppE}), - REPEAT_DETERM_N 2 o dtac (Irel RS iffD1), rtac (Irel RS iffD2), - rtac rel_mono, rtac (le_rel_OO RS @{thm predicate2D}), - rtac @{thm relcomppI}, atac, atac, - REPEAT_DETERM_N m o EVERY' [rtac ballI, rtac ballI, rtac impI, atac], + EVERY' [rtac ctxt impI, etac ctxt (nchotomy RS @{thm nchotomy_relcomppE}), + REPEAT_DETERM_N 2 o dtac ctxt (Irel RS iffD1), rtac ctxt (Irel RS iffD2), + rtac ctxt rel_mono, rtac ctxt (le_rel_OO RS @{thm predicate2D}), + rtac ctxt @{thm relcomppI}, atac, atac, + REPEAT_DETERM_N m o EVERY' [rtac ctxt ballI, rtac ctxt ballI, rtac ctxt impI, atac], REPEAT_DETERM_N (length le_rel_OOs) o - EVERY' [rtac ballI, rtac ballI, Goal.assume_rule_tac ctxt]]) + EVERY' [rtac ctxt ballI, rtac ctxt ballI, Goal.assume_rule_tac ctxt]]) ctor_nchotomys ctor_Irels rel_mono_strong0s le_rel_OOs) 1; (* BNF tactics *) -fun mk_map_id0_tac map_id0s unique = - (rtac sym THEN' rtac unique THEN' +fun mk_map_id0_tac ctxt map_id0s unique = + (rtac ctxt sym THEN' rtac ctxt unique THEN' EVERY' (map (fn thm => - EVERY' [rtac trans, rtac @{thm id_comp}, rtac trans, rtac sym, rtac @{thm comp_id}, - rtac (thm RS sym RS arg_cong)]) map_id0s)) 1; + EVERY' [rtac ctxt trans, rtac ctxt @{thm id_comp}, rtac ctxt trans, rtac ctxt sym, rtac ctxt @{thm comp_id}, + rtac ctxt (thm RS sym RS arg_cong)]) map_id0s)) 1; -fun mk_map_comp0_tac map_comp0s ctor_maps unique iplus1 = +fun mk_map_comp0_tac ctxt map_comp0s ctor_maps unique iplus1 = let val i = iplus1 - 1; val unique' = Thm.permute_prems 0 i unique; val map_comp0s' = drop i map_comp0s @ take i map_comp0s; val ctor_maps' = drop i ctor_maps @ take i ctor_maps; fun mk_comp comp simp = - EVERY' [rtac @{thm ext}, rtac trans, rtac o_apply, rtac trans, rtac o_apply, - rtac trans, rtac (simp RS arg_cong), rtac trans, rtac simp, - rtac trans, rtac (comp RS arg_cong), rtac sym, rtac o_apply]; + EVERY' [rtac ctxt @{thm ext}, rtac ctxt trans, rtac ctxt o_apply, rtac ctxt trans, rtac ctxt o_apply, + rtac ctxt trans, rtac ctxt (simp RS arg_cong), rtac ctxt trans, rtac ctxt simp, + rtac ctxt trans, rtac ctxt (comp RS arg_cong), rtac ctxt sym, rtac ctxt o_apply]; in - (rtac sym THEN' rtac unique' THEN' EVERY' (map2 mk_comp map_comp0s' ctor_maps')) 1 + (rtac ctxt sym THEN' rtac ctxt unique' THEN' EVERY' (map2 mk_comp map_comp0s' ctor_maps')) 1 end; -fun mk_set_map0_tac set_nat = - EVERY' (map rtac [@{thm ext}, trans, o_apply, sym, trans, o_apply, set_nat]) 1; +fun mk_set_map0_tac ctxt set_nat = + EVERY' (map (rtac ctxt) [@{thm ext}, trans, o_apply, sym, trans, o_apply, set_nat]) 1; -fun mk_bd_card_order_tac bd_card_orders = - CONJ_WRAP_GEN' (rtac @{thm card_order_csum}) rtac bd_card_orders 1; +fun mk_bd_card_order_tac ctxt bd_card_orders = + CONJ_WRAP_GEN' (rtac ctxt @{thm card_order_csum}) (rtac ctxt) bd_card_orders 1; fun mk_wit_tac ctxt n ctor_set wit = REPEAT_DETERM (atac 1 ORELSE - EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac ctxt ctor_set, + EVERY' [dtac ctxt set_rev_mp, rtac ctxt equalityD1, resolve_tac ctxt ctor_set, REPEAT_DETERM o - (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN' + (TRY o REPEAT_DETERM o etac ctxt UnE THEN' TRY o etac ctxt @{thm UN_E} THEN' (eresolve_tac ctxt wit ORELSE' (dresolve_tac ctxt wit THEN' - (etac FalseE ORELSE' - EVERY' [hyp_subst_tac ctxt, dtac set_rev_mp, rtac equalityD1, resolve_tac ctxt ctor_set, - REPEAT_DETERM_N n o etac UnE]))))] 1); + (etac ctxt FalseE ORELSE' + EVERY' [hyp_subst_tac ctxt, dtac ctxt set_rev_mp, rtac ctxt equalityD1, resolve_tac ctxt ctor_set, + REPEAT_DETERM_N n o etac ctxt UnE]))))] 1); fun mk_ctor_rel_tac ctxt in_Irels i in_rel map_comp0 map_cong0 ctor_map ctor_sets ctor_inject ctor_dtor set_map0s ctor_set_incls ctor_set_set_inclss = @@ -647,91 +648,91 @@ val le_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS ord_eq_le_trans; val eq_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS trans; val if_tac = - EVERY' [dtac (in_Irel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], - rtac (in_rel RS iffD2), rtac exI, rtac conjI, rtac CollectI, + EVERY' [dtac ctxt (in_Irel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], + rtac ctxt (in_rel RS iffD2), rtac ctxt exI, rtac ctxt conjI, rtac ctxt CollectI, EVERY' (map2 (fn set_map0 => fn ctor_set_incl => - EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_map0, - rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply, - rtac (ctor_set_incl RS subset_trans), etac le_arg_cong_ctor_dtor]) + EVERY' [rtac ctxt conjI, rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, + rtac ctxt ord_eq_le_trans, rtac ctxt trans_fun_cong_image_id_id_apply, + rtac ctxt (ctor_set_incl RS subset_trans), etac ctxt le_arg_cong_ctor_dtor]) passive_set_map0s ctor_set_incls), CONJ_WRAP' (fn (in_Irel, (set_map0, ctor_set_set_incls)) => - EVERY' [rtac ord_eq_le_trans, rtac set_map0, rtac @{thm image_subsetI}, rtac CollectI, - rtac @{thm case_prodI}, rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI, + EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, + rtac ctxt @{thm case_prodI}, rtac ctxt (in_Irel RS iffD2), rtac ctxt exI, rtac ctxt conjI, rtac ctxt CollectI, CONJ_WRAP' (fn thm => - EVERY' (map etac [thm RS subset_trans, le_arg_cong_ctor_dtor])) + EVERY' (map (etac ctxt) [thm RS subset_trans, le_arg_cong_ctor_dtor])) ctor_set_set_incls, - rtac conjI, rtac refl, rtac refl]) + rtac ctxt conjI, rtac ctxt refl, rtac ctxt refl]) (in_Irels ~~ (active_set_map0s ~~ ctor_set_set_inclss)), CONJ_WRAP' (fn conv => - EVERY' [rtac trans, rtac map_comp0, rtac trans, rtac map_cong0, - REPEAT_DETERM_N m o rtac @{thm fun_cong[OF comp_id]}, - REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]), - rtac (ctor_inject RS iffD1), rtac trans, rtac sym, rtac ctor_map, - etac eq_arg_cong_ctor_dtor]) + EVERY' [rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0, + REPEAT_DETERM_N m o rtac ctxt @{thm fun_cong[OF comp_id]}, + REPEAT_DETERM_N n o EVERY' (map (rtac ctxt) [trans, o_apply, conv]), + rtac ctxt (ctor_inject RS iffD1), rtac ctxt trans, rtac ctxt sym, rtac ctxt ctor_map, + etac ctxt eq_arg_cong_ctor_dtor]) fst_snd_convs]; val only_if_tac = - EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], - rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI, + EVERY' [dtac ctxt (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], + rtac ctxt (in_Irel RS iffD2), rtac ctxt exI, rtac ctxt conjI, rtac ctxt CollectI, CONJ_WRAP' (fn (ctor_set, passive_set_map0) => - EVERY' [rtac ord_eq_le_trans, rtac ctor_set, rtac @{thm Un_least}, - rtac ord_eq_le_trans, rtac @{thm box_equals[OF _ refl]}, - rtac passive_set_map0, rtac trans_fun_cong_image_id_id_apply, atac, - CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least})) - (fn (active_set_map0, in_Irel) => EVERY' [rtac ord_eq_le_trans, - rtac @{thm SUP_cong[OF _ refl]}, rtac active_set_map0, rtac @{thm UN_least}, - dtac set_rev_mp, etac @{thm image_mono}, etac imageE, - dtac @{thm ssubst_mem[OF pair_collapse]}, + EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt ctor_set, rtac ctxt @{thm Un_least}, + rtac ctxt ord_eq_le_trans, rtac ctxt @{thm box_equals[OF _ refl]}, + rtac ctxt passive_set_map0, rtac ctxt trans_fun_cong_image_id_id_apply, atac, + CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 @{thm Un_least})) + (fn (active_set_map0, in_Irel) => EVERY' [rtac ctxt ord_eq_le_trans, + rtac ctxt @{thm SUP_cong[OF _ refl]}, rtac ctxt active_set_map0, rtac ctxt @{thm UN_least}, + dtac ctxt set_rev_mp, etac ctxt @{thm image_mono}, etac ctxt imageE, + dtac ctxt @{thm ssubst_mem[OF pair_collapse]}, REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE :: @{thms case_prodE iffD1[OF prod.inject, elim_format]}), hyp_subst_tac ctxt, - dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, + dtac ctxt (in_Irel RS iffD1), dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], atac]) (rev (active_set_map0s ~~ in_Irels))]) (ctor_sets ~~ passive_set_map0s), - rtac conjI, - REPEAT_DETERM_N 2 o EVERY' [rtac trans, rtac ctor_map, rtac (ctor_inject RS iffD2), - rtac trans, rtac map_comp0, rtac trans, rtac map_cong0, - REPEAT_DETERM_N m o rtac @{thm fun_cong[OF comp_id]}, - EVERY' (map (fn in_Irel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac, - dtac @{thm ssubst_mem[OF pair_collapse]}, + rtac ctxt conjI, + REPEAT_DETERM_N 2 o EVERY' [rtac ctxt trans, rtac ctxt ctor_map, rtac ctxt (ctor_inject RS iffD2), + rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0, + REPEAT_DETERM_N m o rtac ctxt @{thm fun_cong[OF comp_id]}, + EVERY' (map (fn in_Irel => EVERY' [rtac ctxt trans, rtac ctxt o_apply, dtac ctxt set_rev_mp, atac, + dtac ctxt @{thm ssubst_mem[OF pair_collapse]}, REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE :: @{thms case_prodE iffD1[OF prod.inject, elim_format]}), hyp_subst_tac ctxt, - dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) + dtac ctxt (in_Irel RS iffD1), dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE, atac]) in_Irels), atac]] in - EVERY' [rtac iffI, if_tac, only_if_tac] 1 + EVERY' [rtac ctxt iffI, if_tac, only_if_tac] 1 end; fun mk_ctor_rec_transfer_tac ctxt n m ctor_rec_defs ctor_fold_transfers pre_T_map_transfers ctor_rels = CONJ_WRAP (fn (ctor_rec_def, ctor_fold_transfer) => - REPEAT_DETERM (HEADGOAL (rtac rel_funI)) THEN + REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN unfold_thms_tac ctxt [ctor_rec_def, o_apply] THEN - HEADGOAL (rtac @{thm rel_funD[OF snd_transfer]} THEN' - etac (mk_rel_funDN_rotated (n + 1) ctor_fold_transfer) THEN' + HEADGOAL (rtac ctxt @{thm rel_funD[OF snd_transfer]} THEN' + etac ctxt (mk_rel_funDN_rotated (n + 1) ctor_fold_transfer) THEN' EVERY' (map2 (fn pre_T_map_transfer => fn ctor_rel => - etac (mk_rel_funDN_rotated 2 @{thm convol_transfer}) THEN' - rtac (mk_rel_funDN_rotated 2 @{thm comp_transfer}) THEN' - rtac (mk_rel_funDN (m + n) pre_T_map_transfer) THEN' - REPEAT_DETERM_N m o rtac @{thm id_transfer} THEN' - REPEAT_DETERM o rtac @{thm fst_transfer} THEN' - rtac rel_funI THEN' - etac (ctor_rel RS iffD2)) pre_T_map_transfers ctor_rels))) + etac ctxt (mk_rel_funDN_rotated 2 @{thm convol_transfer}) THEN' + rtac ctxt (mk_rel_funDN_rotated 2 @{thm comp_transfer}) THEN' + rtac ctxt (mk_rel_funDN (m + n) pre_T_map_transfer) THEN' + REPEAT_DETERM_N m o rtac ctxt @{thm id_transfer} THEN' + REPEAT_DETERM o rtac ctxt @{thm fst_transfer} THEN' + rtac ctxt rel_funI THEN' + etac ctxt (ctor_rel RS iffD2)) pre_T_map_transfers ctor_rels))) (ctor_rec_defs ~~ ctor_fold_transfers); fun mk_rel_induct_tac ctxt IHs m ctor_induct2 ks ctor_rels rel_mono_strong0s = let val n = length ks; in unfold_tac ctxt @{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} THEN - EVERY' [REPEAT_DETERM o rtac allI, rtac ctor_induct2, + EVERY' [REPEAT_DETERM o rtac ctxt allI, rtac ctxt ctor_induct2, EVERY' (@{map 3} (fn IH => fn ctor_rel => fn rel_mono_strong0 => - EVERY' [rtac impI, dtac (ctor_rel RS iffD1), rtac (IH RS @{thm spec2} RS mp), - etac rel_mono_strong0, - REPEAT_DETERM_N m o rtac @{thm ballI[OF ballI[OF imp_refl]]}, + EVERY' [rtac ctxt impI, dtac ctxt (ctor_rel RS iffD1), rtac ctxt (IH RS @{thm spec2} RS mp), + etac ctxt rel_mono_strong0, + REPEAT_DETERM_N m o rtac ctxt @{thm ballI[OF ballI[OF imp_refl]]}, EVERY' (map (fn j => - EVERY' [select_prem_tac n (dtac asm_rl) j, rtac @{thm ballI[OF ballI]}, + EVERY' [select_prem_tac ctxt n (dtac ctxt asm_rl) j, rtac ctxt @{thm ballI[OF ballI]}, Goal.assume_rule_tac ctxt]) ks)]) IHs ctor_rels rel_mono_strong0s)] 1 end; @@ -744,14 +745,14 @@ @{thms rel_fun_def_butlast all_conj_distrib[symmetric] imp_conjR[symmetric]} THEN unfold_thms_tac ctxt @{thms rel_fun_iff_leq_vimage2p} THEN HEADGOAL (EVERY' - [REPEAT_DETERM o resolve_tac ctxt [allI, impI], rtac ctor_rel_induct, + [REPEAT_DETERM o resolve_tac ctxt [allI, impI], rtac ctxt ctor_rel_induct, EVERY' (map (fn map_transfer => EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI, @{thm vimage2pI}], SELECT_GOAL (unfold_thms_tac ctxt folds), - etac @{thm predicate2D_vimage2p}, - rtac (funpow (m + n + 1) (fn thm => thm RS rel_funD) map_transfer), - REPEAT_DETERM_N m o rtac @{thm id_transfer}, - REPEAT_DETERM_N n o rtac @{thm vimage2p_rel_fun}, + etac ctxt @{thm predicate2D_vimage2p}, + rtac ctxt (funpow (m + n + 1) (fn thm => thm RS rel_funD) map_transfer), + REPEAT_DETERM_N m o rtac ctxt @{thm id_transfer}, + REPEAT_DETERM_N n o rtac ctxt @{thm vimage2p_rel_fun}, atac]) map_transfers)]) end; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/BNF/bnf_tactics.ML --- a/src/HOL/Tools/BNF/bnf_tactics.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_tactics.ML Thu Jul 16 12:23:22 2015 +0200 @@ -10,12 +10,12 @@ sig include CTR_SUGAR_GENERAL_TACTICS - val fo_rtac: thm -> Proof.context -> int -> tactic + val fo_rtac: Proof.context -> thm -> int -> tactic val subst_tac: Proof.context -> int list option -> thm list -> int -> tactic val subst_asm_tac: Proof.context -> int list option -> thm list -> int -> tactic - val mk_rotate_eq_tac: (int -> tactic) -> thm -> thm -> thm -> thm -> ''a list -> ''a list -> - int -> tactic + val mk_rotate_eq_tac: Proof.context -> (int -> tactic) -> thm -> thm -> thm -> thm -> ''a list -> + ''a list -> int -> tactic val mk_pointfree: Proof.context -> thm -> thm @@ -24,7 +24,7 @@ val mk_map_comp_id_tac: Proof.context -> thm -> tactic val mk_map_cong0_tac: Proof.context -> int -> thm -> tactic - val mk_map_cong0L_tac: int -> thm -> thm -> tactic + val mk_map_cong0L_tac: Proof.context -> int -> thm -> thm -> tactic end; structure BNF_Tactics : BNF_TACTICS = @@ -34,14 +34,14 @@ open BNF_Util (*stolen from Christian Urban's Cookbook (and adapted slightly)*) -fun fo_rtac thm = Subgoal.FOCUS (fn {concl, ...} => +fun fo_rtac ctxt thm = Subgoal.FOCUS (fn {concl, context = ctxt, ...} => let val concl_pat = Drule.strip_imp_concl (Thm.cprop_of thm) val insts = Thm.first_order_match (concl_pat, concl) in - rtac (Drule.instantiate_normalize insts thm) 1 + rtac ctxt (Drule.instantiate_normalize insts thm) 1 end - handle Pattern.MATCH => no_tac); + handle Pattern.MATCH => no_tac) ctxt; (*unlike "unfold_thms_tac", it succeed when the RHS contains schematic variables not in the LHS*) fun subst_tac ctxt = EqSubst.eqsubst_tac ctxt o the_default [0]; @@ -54,16 +54,16 @@ |> apply2 (dest_comb #> apsnd (dest_comb #> fst) #> HOLogic.mk_comp) |> mk_Trueprop_eq |> (fn goal => Goal.prove_sorry ctxt [] [] goal - (K (rtac @{thm ext} 1 THEN + (K (rtac ctxt @{thm ext} 1 THEN unfold_thms_tac ctxt ([o_apply, unfold_thms ctxt [o_apply] (mk_sym thm)]) THEN - rtac refl 1))) + rtac ctxt refl 1))) |> Thm.close_derivation; (* Theorems for open typedefs with UNIV as representing set *) fun mk_Abs_inj_thm inj = inj OF (replicate 2 UNIV_I); -fun mk_Abs_bij_thm ctxt Abs_inj_thm surj = rule_by_tactic ctxt ((rtac surj THEN' etac exI) 1) +fun mk_Abs_bij_thm ctxt Abs_inj_thm surj = rule_by_tactic ctxt ((rtac ctxt surj THEN' etac ctxt exI) 1) (Abs_inj_thm RS @{thm bijI'}); @@ -71,20 +71,20 @@ (* General tactic generators *) (*applies assoc rule to the lhs of an equation as long as possible*) -fun mk_flatten_assoc_tac refl_tac trans assoc cong = rtac trans 1 THEN - REPEAT_DETERM (CHANGED ((FIRST' [rtac trans THEN' rtac assoc, rtac cong THEN' refl_tac]) 1)) THEN +fun mk_flatten_assoc_tac ctxt refl_tac trans assoc cong = rtac ctxt trans 1 THEN + REPEAT_DETERM (CHANGED ((FIRST' [rtac ctxt trans THEN' rtac ctxt assoc, rtac ctxt cong THEN' refl_tac]) 1)) THEN refl_tac 1; (*proves two sides of an equation to be equal assuming both are flattened and rhs can be obtained from lhs by the given permutation of monoms*) -fun mk_rotate_eq_tac refl_tac trans assoc com cong = +fun mk_rotate_eq_tac ctxt refl_tac trans assoc com cong = let fun gen_tac [] [] = K all_tac | gen_tac [x] [y] = if x = y then refl_tac else error "mk_rotate_eq_tac: different lists" | gen_tac (x :: xs) (y :: ys) = if x = y - then rtac cong THEN' refl_tac THEN' gen_tac xs ys - else rtac trans THEN' rtac com THEN' - K (mk_flatten_assoc_tac refl_tac trans assoc cong) THEN' + then rtac ctxt cong THEN' refl_tac THEN' gen_tac xs ys + else rtac ctxt trans THEN' rtac ctxt com THEN' + K (mk_flatten_assoc_tac ctxt refl_tac trans assoc cong) THEN' gen_tac (xs @ [x]) (y :: ys) | gen_tac _ _ = error "mk_rotate_eq_tac: different lists"; in @@ -92,15 +92,15 @@ end; fun mk_map_comp_id_tac ctxt map_comp0 = - (rtac trans THEN' rtac map_comp0 THEN' K (unfold_thms_tac ctxt @{thms comp_id}) THEN' rtac refl) 1; + (rtac ctxt trans THEN' rtac ctxt map_comp0 THEN' K (unfold_thms_tac ctxt @{thms comp_id}) THEN' rtac ctxt refl) 1; fun mk_map_cong0_tac ctxt m map_cong0 = - EVERY' [rtac mp, rtac map_cong0, - CONJ_WRAP' (K (rtac ballI THEN' Goal.assume_rule_tac ctxt)) (1 upto m)] 1; + EVERY' [rtac ctxt mp, rtac ctxt map_cong0, + CONJ_WRAP' (K (rtac ctxt ballI THEN' Goal.assume_rule_tac ctxt)) (1 upto m)] 1; -fun mk_map_cong0L_tac passive map_cong0 map_id = - (rtac trans THEN' rtac map_cong0 THEN' EVERY' (replicate passive (rtac refl))) 1 THEN - REPEAT_DETERM (EVERY' [rtac trans, etac bspec, atac, rtac sym, rtac @{thm id_apply}] 1) THEN - rtac map_id 1; +fun mk_map_cong0L_tac ctxt passive map_cong0 map_id = + (rtac ctxt trans THEN' rtac ctxt map_cong0 THEN' EVERY' (replicate passive (rtac ctxt refl))) 1 THEN + REPEAT_DETERM (EVERY' [rtac ctxt trans, etac ctxt bspec, atac, rtac ctxt sym, rtac ctxt @{thm id_apply}] 1) THEN + rtac ctxt map_id 1; end; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Thu Jul 16 12:23:22 2015 +0200 @@ -708,7 +708,8 @@ HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u', Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs)); in - Goal.prove_sorry lthy [] [] goal (fn _ => mk_nchotomy_tac n exhaust_thm) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_nchotomy_tac ctxt n exhaust_thm) |> Thm.close_derivation end; @@ -737,7 +738,8 @@ val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase)); in (Goal.prove_sorry lthy [] [] goal (fn _ => mk_case_cong_tac lthy uexhaust_thm case_thms), - Goal.prove_sorry lthy [] [] weak_goal (K (etac arg_cong 1))) + Goal.prove_sorry lthy [] [] weak_goal (fn {context = ctxt, prems = _} => + etac ctxt arg_cong 1)) |> apply2 (singleton (Proof_Context.export names_lthy lthy) #> Thm.close_derivation) end; @@ -823,7 +825,8 @@ val m = the_single ms; val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs); in - Goal.prove_sorry lthy [] [] goal (fn _ => mk_unique_disc_def_tac m uexhaust_thm) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_unique_disc_def_tac ctxt m uexhaust_thm) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation end; @@ -886,7 +889,7 @@ HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))]; fun prove tac goal = - Goal.prove_sorry lthy [] [] goal (K tac) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => tac ctxt) |> Thm.close_derivation; val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs)); @@ -894,12 +897,14 @@ val half_goalss = map mk_goal half_pairss; val half_thmss = @{map 3} (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] => - fn disc_thm => [prove (mk_half_distinct_disc_tac lthy m discD disc_thm) goal]) + fn disc_thm => [prove (fn ctxt => + mk_half_distinct_disc_tac ctxt m discD disc_thm) goal]) half_goalss half_pairss (flat disc_thmss'); val other_half_goalss = map (mk_goal o map swap) half_pairss; val other_half_thmss = - map2 (map2 (prove o mk_other_half_distinct_disc_tac)) half_thmss + map2 (map2 (fn thm => prove (fn ctxt => + mk_other_half_distinct_disc_tac ctxt thm))) half_thmss other_half_goalss; in join_halves n half_thmss other_half_thmss ||> `transpose @@ -911,8 +916,8 @@ fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs)); in - Goal.prove_sorry lthy [] [] goal (fn _ => - mk_exhaust_disc_tac n exhaust_thm discI_thms) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_exhaust_disc_tac ctxt n exhaust_thm discI_thms) |> Thm.close_derivation end; @@ -946,8 +951,8 @@ fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs)); in - Goal.prove_sorry lthy [] [] goal (fn _ => - mk_exhaust_sel_tac n exhaust_disc_thm swapped_all_collapse_thms) + Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => + mk_exhaust_sel_tac ctxt n exhaust_disc_thm swapped_all_collapse_thms) |> Thm.close_derivation end; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML Thu Jul 16 12:23:22 2015 +0200 @@ -7,7 +7,7 @@ signature CTR_SUGAR_GENERAL_TACTICS = sig - val select_prem_tac: int -> (int -> tactic) -> int -> int -> tactic + val select_prem_tac: Proof.context -> int -> (int -> tactic) -> int -> int -> tactic val unfold_thms_tac: Proof.context -> thm list -> tactic end; @@ -24,17 +24,17 @@ val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic val mk_disc_eq_case_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic - val mk_exhaust_disc_tac: int -> thm -> thm list -> tactic - val mk_exhaust_sel_tac: int -> thm -> thm list -> tactic + val mk_exhaust_disc_tac: Proof.context -> int -> thm -> thm list -> tactic + val mk_exhaust_sel_tac: Proof.context -> int -> thm -> thm list -> tactic val mk_expand_tac: Proof.context -> int -> int list -> thm -> thm -> thm list -> thm list list list -> thm list list list -> tactic val mk_half_distinct_disc_tac: Proof.context -> int -> thm -> thm -> tactic - val mk_nchotomy_tac: int -> thm -> tactic - val mk_other_half_distinct_disc_tac: thm -> tactic + val mk_nchotomy_tac: Proof.context -> int -> thm -> tactic + val mk_other_half_distinct_disc_tac: Proof.context -> thm -> tactic val mk_split_tac: Proof.context -> int list -> thm -> thm list -> thm list list -> thm list list -> thm list list list -> tactic val mk_split_asm_tac: Proof.context -> thm -> tactic - val mk_unique_disc_def_tac: int -> thm -> tactic + val mk_unique_disc_def_tac: Proof.context -> int -> thm -> tactic end; structure Ctr_Sugar_Tactics : CTR_SUGAR_TACTICS = @@ -44,57 +44,58 @@ val meta_mp = @{thm meta_mp}; -fun select_prem_tac n tac k = DETERM o (EVERY' [REPEAT_DETERM_N (k - 1) o etac thin_rl, - tac, REPEAT_DETERM_N (n - k) o etac thin_rl]); +fun select_prem_tac ctxt n tac k = DETERM o (EVERY' [REPEAT_DETERM_N (k - 1) o etac ctxt thin_rl, + tac, REPEAT_DETERM_N (n - k) o etac ctxt thin_rl]); fun unfold_thms_tac _ [] = all_tac | unfold_thms_tac ctxt thms = Local_Defs.unfold_tac ctxt (distinct Thm.eq_thm_prop thms); fun if_P_or_not_P_OF pos thm = thm RS (if pos then @{thm if_P} else @{thm if_not_P}); -fun mk_nchotomy_tac n exhaust = - HEADGOAL (rtac allI THEN' rtac exhaust THEN' - EVERY' (maps (fn k => [rtac (mk_disjIN n k), REPEAT_DETERM o rtac exI, atac]) (1 upto n))); +fun mk_nchotomy_tac ctxt n exhaust = + HEADGOAL (rtac ctxt allI THEN' rtac ctxt exhaust THEN' + EVERY' (maps (fn k => [rtac ctxt (mk_disjIN n k), REPEAT_DETERM o rtac ctxt exI, atac]) (1 upto n))); -fun mk_unique_disc_def_tac m uexhaust = - HEADGOAL (EVERY' [rtac iffI, rtac uexhaust, REPEAT_DETERM_N m o rtac exI, atac, rtac refl]); +fun mk_unique_disc_def_tac ctxt m uexhaust = + HEADGOAL (EVERY' [rtac ctxt iffI, rtac ctxt uexhaust, REPEAT_DETERM_N m o rtac ctxt exI, atac, rtac ctxt refl]); fun mk_alternate_disc_def_tac ctxt k other_disc_def distinct uexhaust = - HEADGOAL (EVERY' ([rtac (other_disc_def RS @{thm arg_cong[of _ _ Not]} RS trans), - rtac @{thm iffI_np}, REPEAT_DETERM o etac exE, - hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [not_ex]), REPEAT_DETERM o rtac allI, - rtac distinct, rtac uexhaust] @ - (([etac notE, REPEAT_DETERM o rtac exI, atac], [REPEAT_DETERM o rtac exI, atac]) + HEADGOAL (EVERY' ([rtac ctxt (other_disc_def RS @{thm arg_cong[of _ _ Not]} RS trans), + rtac ctxt @{thm iffI_np}, REPEAT_DETERM o etac ctxt exE, + hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [not_ex]), REPEAT_DETERM o rtac ctxt allI, + rtac ctxt distinct, rtac ctxt uexhaust] @ + (([etac ctxt notE, REPEAT_DETERM o rtac ctxt exI, atac], [REPEAT_DETERM o rtac ctxt exI, atac]) |> k = 1 ? swap |> op @))); fun mk_half_distinct_disc_tac ctxt m discD disc' = - HEADGOAL (dtac discD THEN' REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN' - rtac disc'); - -fun mk_other_half_distinct_disc_tac half = HEADGOAL (etac @{thm contrapos_pn} THEN' etac half); + HEADGOAL (dtac ctxt discD THEN' REPEAT_DETERM_N m o etac ctxt exE THEN' hyp_subst_tac ctxt THEN' + rtac ctxt disc'); -fun mk_exhaust_disc_or_sel_tac n exhaust destIs = - HEADGOAL (rtac exhaust THEN' - EVERY' (map2 (fn k => fn destI => dtac destI THEN' - select_prem_tac n (etac meta_mp) k THEN' atac) (1 upto n) destIs)); +fun mk_other_half_distinct_disc_tac ctxt half = + HEADGOAL (etac ctxt @{thm contrapos_pn} THEN' etac ctxt half); + +fun mk_exhaust_disc_or_sel_tac ctxt n exhaust destIs = + HEADGOAL (rtac ctxt exhaust THEN' + EVERY' (map2 (fn k => fn destI => dtac ctxt destI THEN' + select_prem_tac ctxt n (etac ctxt meta_mp) k THEN' atac) (1 upto n) destIs)); val mk_exhaust_disc_tac = mk_exhaust_disc_or_sel_tac; -fun mk_exhaust_sel_tac n exhaust_disc collapses = - mk_exhaust_disc_or_sel_tac n exhaust_disc collapses ORELSE - HEADGOAL (etac meta_mp THEN' resolve0_tac collapses); +fun mk_exhaust_sel_tac ctxt n exhaust_disc collapses = + mk_exhaust_disc_or_sel_tac ctxt n exhaust_disc collapses ORELSE + HEADGOAL (etac ctxt meta_mp THEN' resolve0_tac collapses); fun mk_collapse_tac ctxt m discD sels = - HEADGOAL (dtac discD THEN' + HEADGOAL (dtac ctxt discD THEN' (if m = 0 then atac else - REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN' - SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac refl)); + REPEAT_DETERM_N m o etac ctxt exE THEN' hyp_subst_tac ctxt THEN' + SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac ctxt refl)); fun mk_disc_eq_case_tac ctxt ct exhaust discs distincts cases = HEADGOAL Goal.conjunction_tac THEN - ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW + ALLGOALS (rtac ctxt (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW (hyp_subst_tac ctxt THEN' SELECT_GOAL (unfold_thms_tac ctxt (@{thms not_True_eq_False not_False_eq_True} @ cases @ ((refl :: discs @ distincts) RL [eqTrueI, eqFalseI]))) THEN' @@ -103,87 +104,87 @@ fun mk_expand_tac ctxt n ms uexhaust_disc vexhaust_disc uncollapses distinct_discsss distinct_discsss' = if ms = [0] then - HEADGOAL (rtac (@{thm trans_sym} OF (replicate 2 (the_single uncollapses))) THEN' - TRY o EVERY' [rtac uexhaust_disc, atac, rtac vexhaust_disc, atac]) + HEADGOAL (rtac ctxt (@{thm trans_sym} OF (replicate 2 (the_single uncollapses))) THEN' + TRY o EVERY' [rtac ctxt uexhaust_disc, atac, rtac ctxt vexhaust_disc, atac]) else let val ks = 1 upto n in - HEADGOAL (rtac uexhaust_disc THEN' + HEADGOAL (rtac ctxt uexhaust_disc THEN' EVERY' (@{map 5} (fn k => fn m => fn distinct_discss => fn distinct_discss' => fn uuncollapse => - EVERY' [rtac (uuncollapse RS trans) THEN' TRY o atac, rtac sym, rtac vexhaust_disc, + EVERY' [rtac ctxt (uuncollapse RS trans) THEN' TRY o atac, rtac ctxt sym, rtac ctxt vexhaust_disc, EVERY' (@{map 4} (fn k' => fn distinct_discs => fn distinct_discs' => fn vuncollapse => EVERY' (if k' = k then - [rtac (vuncollapse RS trans), TRY o atac] @ + [rtac ctxt (vuncollapse RS trans), TRY o atac] @ (if m = 0 then - [rtac refl] + [rtac ctxt refl] else - [if n = 1 then K all_tac else EVERY' [dtac meta_mp, atac, dtac meta_mp, atac], - REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE, + [if n = 1 then K all_tac else EVERY' [dtac ctxt meta_mp, atac, dtac ctxt meta_mp, atac], + REPEAT_DETERM_N (Int.max (0, m - 1)) o etac ctxt conjE, asm_simp_tac (ss_only [] ctxt)]) else - [dtac (the_single (if k = n then distinct_discs else distinct_discs')), - etac (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}), + [dtac ctxt (the_single (if k = n then distinct_discs else distinct_discs')), + etac ctxt (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}), atac, atac])) ks distinct_discss distinct_discss' uncollapses)]) ks ms distinct_discsss distinct_discsss' uncollapses)) end; fun mk_case_same_ctr_tac ctxt injects = - REPEAT_DETERM o etac exE THEN' etac conjE THEN' + REPEAT_DETERM o etac ctxt exE THEN' etac ctxt conjE THEN' (case injects of [] => atac - | [inject] => dtac (inject RS iffD1) THEN' REPEAT_DETERM o etac conjE THEN' - hyp_subst_tac ctxt THEN' rtac refl); + | [inject] => dtac ctxt (inject RS iffD1) THEN' REPEAT_DETERM o etac ctxt conjE THEN' + hyp_subst_tac ctxt THEN' rtac ctxt refl); fun mk_case_distinct_ctrs_tac ctxt distincts = - REPEAT_DETERM o etac exE THEN' etac conjE THEN' full_simp_tac (ss_only distincts ctxt); + REPEAT_DETERM o etac ctxt exE THEN' etac ctxt conjE THEN' full_simp_tac (ss_only distincts ctxt); fun mk_case_tac ctxt n k case_def injects distinctss = let val case_def' = mk_unabs_def (n + 1) (case_def RS meta_eq_to_obj_eq); val ks = 1 upto n; in - HEADGOAL (rtac (case_def' RS trans) THEN' rtac @{thm the_equality} THEN' - rtac (mk_disjIN n k) THEN' REPEAT_DETERM o rtac exI THEN' rtac conjI THEN' rtac refl THEN' - rtac refl THEN' + HEADGOAL (rtac ctxt (case_def' RS trans) THEN' rtac ctxt @{thm the_equality} THEN' + rtac ctxt (mk_disjIN n k) THEN' REPEAT_DETERM o rtac ctxt exI THEN' rtac ctxt conjI THEN' rtac ctxt refl THEN' + rtac ctxt refl THEN' EVERY' (map2 (fn k' => fn distincts => - (if k' < n then etac disjE else K all_tac) THEN' + (if k' < n then etac ctxt disjE else K all_tac) THEN' (if k' = k then mk_case_same_ctr_tac ctxt injects else mk_case_distinct_ctrs_tac ctxt distincts)) ks distinctss)) end; fun mk_case_distrib_tac ctxt ct exhaust cases = - HEADGOAL (rtac (cterm_instantiate_pos [SOME ct] exhaust)) THEN - ALLGOALS (hyp_subst_tac ctxt THEN' SELECT_GOAL (unfold_thms_tac ctxt cases) THEN' rtac refl); + HEADGOAL (rtac ctxt (cterm_instantiate_pos [SOME ct] exhaust)) THEN + ALLGOALS (hyp_subst_tac ctxt THEN' SELECT_GOAL (unfold_thms_tac ctxt cases) THEN' rtac ctxt refl); fun mk_case_cong_tac ctxt uexhaust cases = - HEADGOAL (rtac uexhaust THEN' - EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex] ctxt)]) cases)); + HEADGOAL (rtac ctxt uexhaust THEN' + EVERY' (maps (fn casex => [dtac ctxt sym, asm_simp_tac (ss_only [casex] ctxt)]) cases)); fun mk_case_eq_if_tac ctxt n uexhaust cases discss' selss = - HEADGOAL (rtac uexhaust THEN' + HEADGOAL (rtac ctxt uexhaust THEN' EVERY' (@{map 3} (fn casex => fn if_discs => fn sels => EVERY' [hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt (if_discs @ sels)), - rtac casex]) + rtac ctxt casex]) cases (map2 (seq_conds if_P_or_not_P_OF n) (1 upto n) discss') selss)); fun mk_split_tac ctxt ms uexhaust cases selss injectss distinctsss = let val depth = fold Integer.max ms 0 in - HEADGOAL (rtac uexhaust) THEN + HEADGOAL (rtac ctxt uexhaust) THEN ALLGOALS (fn k => (hyp_subst_tac ctxt THEN' simp_tac (ss_only (@{thms simp_thms} @ cases @ nth selss (k - 1) @ nth injectss (k - 1) @ flat (nth distinctsss (k - 1))) ctxt)) k) THEN - ALLGOALS (etac thin_rl THEN' rtac iffI THEN' - REPEAT_DETERM o rtac allI THEN' rtac impI THEN' REPEAT_DETERM o etac conjE THEN' - hyp_subst_tac ctxt THEN' atac THEN' REPEAT_DETERM o etac allE THEN' etac impE THEN' - REPEAT_DETERM o (rtac conjI THEN' rtac refl) THEN' rtac refl THEN' atac) + ALLGOALS (etac ctxt thin_rl THEN' rtac ctxt iffI THEN' + REPEAT_DETERM o rtac ctxt allI THEN' rtac ctxt impI THEN' REPEAT_DETERM o etac ctxt conjE THEN' + hyp_subst_tac ctxt THEN' atac THEN' REPEAT_DETERM o etac ctxt allE THEN' etac ctxt impE THEN' + REPEAT_DETERM o (rtac ctxt conjI THEN' rtac ctxt refl) THEN' rtac ctxt refl THEN' atac) end; val split_asm_thms = @{thms imp_conv_disj de_Morgan_conj de_Morgan_disj not_not not_ex}; fun mk_split_asm_tac ctxt split = - HEADGOAL (rtac (split RS trans)) THEN unfold_thms_tac ctxt split_asm_thms THEN - HEADGOAL (rtac refl); + HEADGOAL (rtac ctxt (split RS trans)) THEN unfold_thms_tac ctxt split_asm_thms THEN + HEADGOAL (rtac ctxt refl); end; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML Thu Jul 16 12:23:22 2015 +0200 @@ -54,6 +54,11 @@ val rapp: term -> term -> term + val rtac: Proof.context -> thm -> int -> tactic + val etac: Proof.context -> thm -> int -> tactic + val dtac: Proof.context -> thm -> int -> tactic + val ftac: Proof.context -> thm -> int -> tactic + val list_all_free: term list -> term -> term val list_exists_free: term list -> term -> term @@ -273,4 +278,9 @@ fun CONJ_WRAP gen_tac = CONJ_WRAP_GEN (rtac conjI 1) gen_tac; fun CONJ_WRAP' gen_tac = CONJ_WRAP_GEN' (rtac conjI) gen_tac; +fun rtac ctxt thm = resolve_tac ctxt [thm]; +fun etac ctxt thm = eresolve_tac ctxt [thm]; +fun dtac ctxt thm = dresolve_tac ctxt [thm]; +fun ftac ctxt thm = forward_tac ctxt [thm]; + end; diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/Lifting/lifting_bnf.ML --- a/src/HOL/Tools/Lifting/lifting_bnf.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/Lifting/lifting_bnf.ML Thu Jul 16 12:23:22 2015 +0200 @@ -31,11 +31,11 @@ val rel_conversep_sym = rel_conversep_of_bnf bnf RS sym in EVERY' [SELECT_GOAL (Local_Defs.unfold_tac ctxt [@{thm Quotient_alt_def5}]), - REPEAT_DETERM o (etac conjE), rtac conjI, SELECT_GOAL (Local_Defs.unfold_tac ctxt [rel_Grp_UNIV_sym]), - rtac rel_mono THEN_ALL_NEW assume_tac ctxt, rtac conjI, SELECT_GOAL (Local_Defs.unfold_tac ctxt - [rel_conversep_sym, rel_Grp_UNIV_sym]), rtac rel_mono THEN_ALL_NEW assume_tac ctxt, + REPEAT_DETERM o (etac ctxt conjE), rtac ctxt conjI, SELECT_GOAL (Local_Defs.unfold_tac ctxt [rel_Grp_UNIV_sym]), + rtac ctxt rel_mono THEN_ALL_NEW assume_tac ctxt, rtac ctxt conjI, SELECT_GOAL (Local_Defs.unfold_tac ctxt + [rel_conversep_sym, rel_Grp_UNIV_sym]), rtac ctxt rel_mono THEN_ALL_NEW assume_tac ctxt, SELECT_GOAL (Local_Defs.unfold_tac ctxt [rel_conversep_sym, rel_OO_of_bnf bnf RS sym]), - hyp_subst_tac ctxt, rtac refl] i + hyp_subst_tac ctxt, rtac ctxt refl] i end fun mk_Quotient args = diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/Lifting/lifting_def_code_dt.ML --- a/src/HOL/Tools/Lifting/lifting_def_code_dt.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/Lifting/lifting_def_code_dt.ML Thu Jul 16 12:23:22 2015 +0200 @@ -218,7 +218,7 @@ end; fun case_tac rule ctxt i st = - (Subgoal.FOCUS_PARAMS (fn {params, ...} => HEADGOAL(rtac + (Subgoal.FOCUS_PARAMS (fn {params, ...} => HEADGOAL(rtac ctxt (Ctr_Sugar_Util.cterm_instantiate_pos [SOME (params |> hd |> snd)] rule))) ctxt i st); fun bundle_name_of_bundle_binding binding phi context = @@ -283,7 +283,8 @@ val rsp_norm = Conv.fconv_rule rel_eq_onps_conv rsp val f'_rsp_goal = HOLogic.mk_Trueprop (f'_rsp_rel $ rhs $ rhs); val f'_rsp = Goal.prove_sorry lthy [] [] f'_rsp_goal - (K (HEADGOAL (CONVERSION (rel_eq_onps_conv) THEN' rtac rsp_norm))) + (fn {context = ctxt, prems = _} => + HEADGOAL (CONVERSION (rel_eq_onps_conv) THEN' rtac ctxt rsp_norm)) |> Thm.close_derivation val (f'_lift_def, lthy) = add_lift_def ld_no_notes f'_var f'_qty rhs f'_rsp [] lthy val f'_lift_def = inst_of_lift_def lthy f'_qty f'_lift_def @@ -295,7 +296,7 @@ val f_alt_def_goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (f_alt_def_goal_lhs, f_alt_def_goal_rhs)); fun f_alt_def_tac ctxt i = EVERY' [Transfer.gen_frees_tac [] ctxt, DETERM o Transfer.transfer_tac true ctxt, - SELECT_GOAL (Local_Defs.unfold_tac ctxt [id_apply]), rtac refl] i; + SELECT_GOAL (Local_Defs.unfold_tac ctxt [id_apply]), rtac ctxt refl] i; val rep_isom_transfer = transfer_of_rep_isom_data rep_isom_data val (_, transfer_lthy) = Proof_Context.note_thmss "" [((Binding.empty, []), [([rep_isom_transfer], [Transfer.transfer_add])])] lthy @@ -322,7 +323,7 @@ fun eq_onp_to_top_tac ctxt = SELECT_GOAL (Local_Defs.unfold_tac ctxt (@{thm eq_onp_top_eq_eq[symmetric]} :: Lifting_Info.get_relator_eq_onp_rules ctxt)) fun lift_isom_tac ctxt = HEADGOAL (eq_onp_to_top_tac ctxt - THEN' (rtac @{thm id_transfer})); + THEN' (rtac ctxt @{thm id_transfer})); val (rep_isom_lift_def, lthy) = lift_def ld_no_notes (Binding.qualified true "Rep_isom" uTname, NoSyn) (qty_isom --> qty) (HOLogic.id_const rty) lift_isom_tac [] lthy @@ -408,11 +409,11 @@ fun lift_sel_tac exhaust_rule dt_rules wits ctxt i = (Method.insert_tac wits THEN' eq_onp_to_top_tac ctxt THEN' (* normalize *) - rtac unfold_lift_sel_rsp THEN' + rtac ctxt unfold_lift_sel_rsp THEN' case_tac exhaust_rule ctxt THEN_ALL_NEW ( EVERY' [hyp_subst_tac ctxt, (* does not kill wits because = was rewritten to eq_onp top *) Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq dt_rules), - REPEAT_DETERM o etac conjE, atac])) i + REPEAT_DETERM o etac ctxt conjE, atac])) i val pred_simps = Transfer.lookup_pred_data lthy (Tname rty) |> the |> Transfer.pred_simps val sel_tac = lift_sel_tac (#exhaust ctr_sugar) (#case_thms ctr_sugar @ pred_simps) val sel_names = map (fn (k, xs) => map (fn k' => Binding.qualified true @@ -437,7 +438,7 @@ SELECT_GOAL (Local_Defs.unfold_tac ctxt @{thms eq_onp_top_eq_eq}) (* normalize *), Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq @{thms id_apply simp_thms Ball_def}), - rtac TrueI] i; + rtac ctxt TrueI] i; val (_, transfer_lthy) = Proof_Context.note_thmss "" [((Binding.empty, []), [(@{thms right_total_UNIV_transfer},[Transfer.transfer_add]), @@ -508,7 +509,7 @@ in EVERY' [Transfer.gen_frees_tac [] ctxt, DETERM o (Transfer.transfer_tac true ctxt), case_tac exhaust ctxt THEN_ALL_NEW EVERY' [hyp_subst_tac ctxt, - Raw_Simplifier.rewrite_goal_tac ctxt simp_rules, rtac TrueI ]] i + Raw_Simplifier.rewrite_goal_tac ctxt simp_rules, rtac ctxt TrueI ]] i end end @@ -777,7 +778,7 @@ let val internal_rsp_thm = Goal.prove lthy [] [] prsp_tm (fn {context = ctxt, ...} => - rtac readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac ctxt [thm] 1) + rtac ctxt readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac ctxt [thm] 1) in after_qed internal_rsp_thm lthy end diff -r 53697011b03a -r 26ffdb966759 src/HOL/Tools/Transfer/transfer_bnf.ML --- a/src/HOL/Tools/Transfer/transfer_bnf.ML Thu Jul 16 10:48:20 2015 +0200 +++ b/src/HOL/Tools/Transfer/transfer_bnf.ML Thu Jul 16 12:23:22 2015 +0200 @@ -84,14 +84,14 @@ val thms = constr_defs @ map mk_sym [rel_eq_of_bnf bnf, rel_conversep_of_bnf bnf, rel_OO_of_bnf bnf] in - (SELECT_GOAL (Local_Defs.unfold_tac ctxt thms) THEN' rtac (rel_mono_of_bnf bnf) + (SELECT_GOAL (Local_Defs.unfold_tac ctxt thms) THEN' rtac ctxt (rel_mono_of_bnf bnf) THEN_ALL_NEW assume_tac ctxt) i end fun bi_constraint_tac constr_iff sided_constr_intros ctxt i = (SELECT_GOAL (Local_Defs.unfold_tac ctxt [constr_iff]) THEN' - CONJ_WRAP' (fn thm => rtac thm THEN_ALL_NEW - (REPEAT_DETERM o etac conjE THEN' assume_tac ctxt)) sided_constr_intros) i + CONJ_WRAP' (fn thm => rtac ctxt thm THEN_ALL_NEW + (REPEAT_DETERM o etac ctxt conjE THEN' assume_tac ctxt)) sided_constr_intros) i fun generate_relation_constraint_goal ctxt bnf constraint_def = let @@ -204,24 +204,24 @@ val n = live_of_bnf bnf val set_map's = set_map_of_bnf bnf in - EVERY' [rtac ext, SELECT_GOAL (Local_Defs.unfold_tac ctxt [@{thm Domainp.simps}, - in_rel_of_bnf bnf, pred_def]), rtac iffI, + EVERY' [rtac ctxt ext, SELECT_GOAL (Local_Defs.unfold_tac ctxt [@{thm Domainp.simps}, + in_rel_of_bnf bnf, pred_def]), rtac ctxt iffI, REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], hyp_subst_tac ctxt, - CONJ_WRAP' (fn set_map => EVERY' [rtac ballI, dtac (set_map RS equalityD1 RS set_mp), - etac imageE, dtac set_rev_mp, assume_tac ctxt, + CONJ_WRAP' (fn set_map => EVERY' [rtac ctxt ballI, dtac ctxt (set_map RS equalityD1 RS set_mp), + etac ctxt imageE, dtac ctxt set_rev_mp, assume_tac ctxt, REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm case_prodE}], - hyp_subst_tac ctxt, rtac @{thm iffD2[OF arg_cong2[of _ _ _ _ Domainp, OF refl fst_conv]]}, - etac @{thm DomainPI}]) set_map's, - REPEAT_DETERM o etac conjE, + hyp_subst_tac ctxt, rtac ctxt @{thm iffD2[OF arg_cong2[of _ _ _ _ Domainp, OF refl fst_conv]]}, + etac ctxt @{thm DomainPI}]) set_map's, + REPEAT_DETERM o etac ctxt conjE, REPEAT_DETERM o resolve_tac ctxt [exI, (refl RS conjI), rotate_prems 1 conjI], - rtac refl, rtac (box_equals OF [map_cong0_of_bnf bnf, map_comp_of_bnf bnf RS sym, + rtac ctxt refl, rtac ctxt (box_equals OF [map_cong0_of_bnf bnf, map_comp_of_bnf bnf RS sym, map_id_of_bnf bnf]), - REPEAT_DETERM_N n o (EVERY' [rtac @{thm box_equals[OF _ sym[OF o_apply] sym[OF id_apply]]}, - rtac @{thm fst_conv}]), rtac CollectI, - CONJ_WRAP' (fn set_map => EVERY' [rtac (set_map RS @{thm ord_eq_le_trans}), + REPEAT_DETERM_N n o (EVERY' [rtac ctxt @{thm box_equals[OF _ sym[OF o_apply] sym[OF id_apply]]}, + rtac ctxt @{thm fst_conv}]), rtac ctxt CollectI, + CONJ_WRAP' (fn set_map => EVERY' [rtac ctxt (set_map RS @{thm ord_eq_le_trans}), REPEAT_DETERM o resolve_tac ctxt [@{thm image_subsetI}, CollectI, @{thm case_prodI}], - dtac (rotate_prems 1 bspec), assume_tac ctxt, - etac @{thm DomainpE}, etac @{thm someI}]) set_map's + dtac ctxt (rotate_prems 1 bspec), assume_tac ctxt, + etac ctxt @{thm DomainpE}, etac ctxt @{thm someI}]) set_map's ] i end @@ -251,7 +251,7 @@ fun pred_eq_onp_tac bnf pred_def ctxt i = (SELECT_GOAL (Local_Defs.unfold_tac ctxt [@{thm eq_onp_Grp}, @{thm Ball_Collect}, pred_def]) THEN' CONVERSION (subst_conv (map_id0_of_bnf bnf RS sym)) - THEN' rtac (rel_Grp_of_bnf bnf)) i + THEN' rtac ctxt (rel_Grp_of_bnf bnf)) i fun prove_rel_eq_onp ctxt bnf pred_def = let