diff -r 3890483c674f -r e1e2ca92760b Nominal/Ex/Typing.thy --- a/Nominal/Ex/Typing.thy Tue Jan 04 13:47:38 2011 +0000 +++ b/Nominal/Ex/Typing.thy Wed Jan 05 16:51:27 2011 +0000 @@ -21,8 +21,11 @@ thm lam.size_eqvt ML {* +fun mk_cplus p q = Thm.capply (Thm.capply @{cterm "plus::perm \ perm \ perm"} p) q + fun mk_cminus p = Thm.capply @{cterm "uminus::perm \ perm"} p + fun minus_permute_intro_tac p = rtac (Drule.instantiate' [] [SOME (mk_cminus p)] @{thm permute_boolE}) @@ -48,8 +51,13 @@ |> HOLogic.mk_Trueprop |> (curry Logic.list_implies) prems |> (curry list_all_free) params + val finite_goal = avoid_trm + |> mk_finite + |> HOLogic.mk_Trueprop + |> (curry Logic.list_implies) prems + |> (curry list_all_free) params in - if null avoid then [] else [vc_goal] + if null avoid then [] else [vc_goal, finite_goal] end *} @@ -136,6 +144,12 @@ *} ML {* +fun map7 _ [] [] [] [] [] [] [] = [] + | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (r :: rs) (s :: ss) = + f x y z u v r s :: map7 f xs ys zs us vs rs ss +*} + +ML {* (* local abbreviations *) fun eqvt_stac ctxt = Nominal_Permeq.eqvt_strict_tac ctxt @{thms permute_minus_cancel} [] fun eqvt_srule ctxt = Nominal_Permeq.eqvt_strict_rule ctxt @{thms permute_minus_cancel} [] @@ -157,12 +171,16 @@ val prems' = prems |> map (minus_permute_elim p) |> map (eqvt_srule context) - + val prm' = (prems' MRS prm) |> flag ? (all_elims [p]) - |> flag ? (simplify (HOL_basic_ss addsimps @{thms permute_minus_cancel})) + |> flag ? (eqvt_srule context) + + val _ = tracing ("prm':" ^ @{make_string} prm') in - simp_tac (HOL_ss addsimps (prm' :: @{thms induct_forall_def })) 1 + print_tac "start helper" + THEN asm_full_simp_tac (HOL_ss addsimps (prm' :: @{thms induct_forall_def})) 1 + THEN print_tac "final helper" end) ctxt *} @@ -178,7 +196,7 @@ val prem' = cterm_instantiate (intr_cvars ~~ p_prms) prem (* for inductive-premises*) - fun tac1 prm = helper_tac true prm p context + fun tac1 prm = helper_tac true prm p context (* for non-inductive premises *) fun tac2 prm = @@ -189,63 +207,120 @@ fun select prm (t, i) = (if member same_name Ps (real_head_of t) then tac1 prm else tac2 prm) i in - EVERY1 [rtac prem', RANGE (map (SUBGOAL o select) prems) ] + EVERY1 [eqvt_stac ctxt, rtac prem', RANGE (map (SUBGOAL o select) prems) ] end) ctxt *} ML {* -fun fresh_thm ctxt fresh_thms p c prms avoid_trm = +fun fresh_thm ctxt user_thm p c concl_args avoid_trm = let val conj1 = mk_fresh_star (mk_perm (Bound 0) (mk_perm p avoid_trm)) c val conj2 = - mk_fresh_star_ty @{typ perm} (mk_supp (HOLogic.mk_tuple (map (mk_perm p) prms))) (Bound 0) + mk_fresh_star_ty @{typ perm} (mk_supp (HOLogic.mk_tuple (map (mk_perm p) concl_args))) (Bound 0) val fresh_goal = mk_exists ("q", @{typ perm}) (HOLogic.mk_conj (conj1, conj2)) |> HOLogic.mk_Trueprop - val _ = tracing ("fresh goal: " ^ Syntax.string_of_term ctxt fresh_goal) + val ss = @{thms finite_supp supp_Pair finite_Un permute_finite} @ + @{thms fresh_star_Pair fresh_star_permute_iff} + val simp = asm_full_simp_tac (HOL_ss addsimps ss) in Goal.prove ctxt [] [] fresh_goal - (K (HEADGOAL (rtac @{thm at_set_avoiding2}))) + (K (HEADGOAL (rtac @{thm at_set_avoiding2} + THEN_ALL_NEW EVERY' [cut_facts_tac user_thm, REPEAT o etac @{thm conjE}, simp]))) + end +*} + +ML {* +val supp_perm_eq' = + @{lemma "supp (p \ x) \* q ==> p \ x == (q + p) \ x" by (simp add: supp_perm_eq)} +val fresh_star_plus = + @{lemma "(q \ (p \ x)) \* c ==> ((q + p) \ x) \* c" by (simp add: permute_plus)} +*} + +ML {* +fun binder_tac prem intr_cvars param_trms Ps user_thm avoid avoid_trm concl_args ctxt = + Subgoal.FOCUS (fn {context = ctxt, params, prems, concl, ...} => + let + val thy = ProofContext.theory_of ctxt + val (prms, p, c) = split_last2 (map snd params) + val prm_trms = map term_of prms + val prm_tys = map fastype_of prm_trms + + val avoid_trm' = subst_free (param_trms ~~ prm_trms) avoid_trm + val concl_args' = map (subst_free (param_trms ~~ prm_trms)) concl_args + + val user_thm' = map (cterm_instantiate (intr_cvars ~~ prms)) user_thm + |> map (full_simplify (HOL_ss addsimps (@{thm fresh_star_Pair}::prems))) + + val fthm = fresh_thm ctxt user_thm' (term_of p) (term_of c) concl_args' avoid_trm' + + val (([(_, q)], fprop :: fresh_eqs), ctxt') = Obtain.result + (K (EVERY1 [etac @{thm exE}, + full_simp_tac (HOL_basic_ss addsimps @{thms supp_Pair fresh_star_Un}), + REPEAT o etac @{thm conjE}, + dtac fresh_star_plus, + REPEAT o dtac supp_perm_eq'])) [fthm] ctxt + + val expand_conv = Conv.try_conv (Conv.rewrs_conv fresh_eqs) + fun expand_conv_bot ctxt = Conv.bottom_conv (K expand_conv) ctxt + + val cperms = map (cterm_of thy o perm_const) prm_tys + val qp_prms = map2 (fn ct1 => fn ct2 => Thm.mk_binop ct1 (mk_cplus q p) ct2) cperms prms + val prem' = cterm_instantiate (intr_cvars ~~ qp_prms) prem + + val fprop' = eqvt_srule ctxt' fprop + val tac_fresh = simp_tac (HOL_basic_ss addsimps [fprop']) + + (* for inductive-premises*) + fun tac1 prm = helper_tac true prm (mk_cplus q p) ctxt' + + (* for non-inductive premises *) + fun tac2 prm = + EVERY' [ minus_permute_intro_tac (mk_cplus q p), + eqvt_stac ctxt, + helper_tac false prm (mk_cplus q p) ctxt' ] + + fun select prm (t, i) = + (if member same_name Ps (real_head_of t) then tac1 prm else tac2 prm) i + + val _ = tracing ("fthm:\n" ^ @{make_string} fthm) + val _ = tracing ("fr_eqs:\n" ^ cat_lines (map @{make_string} fresh_eqs)) + val _ = tracing ("fprop:\n" ^ @{make_string} fprop) + val _ = tracing ("fprop':\n" ^ @{make_string} fprop') + val _ = tracing ("fperm:\n" ^ @{make_string} q) + val _ = tracing ("prem':\n" ^ @{make_string} prem') + + val side_thm = Goal.prove ctxt' [] [] (term_of concl) + (fn {context, ...} => + EVERY1 [ CONVERSION (expand_conv_bot context), + eqvt_stac context, + rtac prem', + RANGE (tac_fresh :: map (SUBGOAL o select) prems), + K (print_tac "GOAL") ]) + |> singleton (ProofContext.export ctxt' ctxt) + in + rtac side_thm 1 + end) ctxt +*} + +ML {* +fun case_tac ctxt Ps avoid avoid_trm intr_cvars param_trms prem user_thm concl_args = + let + val tac1 = non_binder_tac prem intr_cvars Ps ctxt + val tac2 = binder_tac prem intr_cvars param_trms Ps user_thm avoid avoid_trm concl_args ctxt + in + EVERY' [ rtac @{thm allI}, rtac @{thm allI}, if null avoid then tac1 else tac2 ] end *} ML {* -fun binder_tac prem intr_cvars param_trms Ps fresh_thms avoid avoid_trm ctxt = - Subgoal.FOCUS (fn {context, params, ...} => - let - val thy = ProofContext.theory_of context - val (prms, p, c) = split_last2 (map snd params) - val prm_trms = map term_of prms - val prm_tys = map fastype_of prm_trms - val cperms = map (cterm_of thy o perm_const) prm_tys - val p_prms = map2 (fn ct1 => fn ct2 => Thm.mk_binop ct1 p ct2) cperms prms - val prem' = cterm_instantiate (intr_cvars ~~ p_prms) prem - val avoid_trm' = subst_free (param_trms ~~ prm_trms) avoid_trm - - val fthm = fresh_thm context fresh_thms (term_of p) (term_of c) (map term_of prms) avoid_trm' - in - Skip_Proof.cheat_tac thy - (* EVERY1 [rtac prem'] *) - end) ctxt -*} - -ML {* -fun case_tac ctxt fresh_thms Ps (avoid, avoid_trm) intr_cvars param_trms prem = +fun prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms concl_args + {prems, context} = let - val tac1 = non_binder_tac prem intr_cvars Ps ctxt - val tac2 = binder_tac prem intr_cvars param_trms Ps fresh_thms avoid avoid_trm ctxt - in - EVERY' [ rtac @{thm allI}, rtac @{thm allI}, eqvt_stac ctxt, - if null avoid then tac1 else tac2 ] - end -*} - -ML {* -fun prove_sinduct_tac raw_induct fresh_thms Ps avoids avoid_trms intr_cvars param_trms {prems, context} = - let - val cases_tac = map4 (case_tac context fresh_thms Ps) (avoids ~~avoid_trms) intr_cvars param_trms prems + val cases_tac = + map7 (case_tac context Ps) avoids avoid_trms intr_cvars param_trms prems user_thms concl_args in EVERY1 [ DETERM o rtac raw_induct, RANGE cases_tac ] end @@ -255,8 +330,10 @@ val normalise = @{lemma "(Q --> (!p c. P p c)) ==> (!!c. Q ==> P (0::perm) c)" by simp} *} +ML {* Local_Theory.note *} + ML {* -fun prove_strong_inductive rule_names avoids raw_induct intrs ctxt = +fun prove_strong_inductive pred_names rule_names avoids raw_induct intrs ctxt = let val thy = ProofContext.theory_of ctxt val ((_, [raw_induct']), ctxt') = Variable.import true [raw_induct] ctxt @@ -311,10 +388,10 @@ val ind_prems' = ind_prems |> map2 (prep_prem Ps c_name c_ty) (avoids ~~ avoid_trms) - fun after_qed ctxt_outside fresh_thms ctxt = + fun after_qed ctxt_outside user_thms ctxt = let - val thms = Goal.prove ctxt [] ind_prems' ind_concl' - (prove_sinduct_tac raw_induct fresh_thms Ps avoids avoid_trms intr_cvars param_trms) + val strong_ind_thms = Goal.prove ctxt [] ind_prems' ind_concl' + (prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms intr_concls_args) |> singleton (ProofContext.export ctxt ctxt_outside) |> Datatype_Aux.split_conj_thm |> map (fn thm => thm RS normalise) @@ -322,10 +399,22 @@ |> map (Drule.rotate_prems (length ind_prems')) |> map zero_var_indexes - val _ = tracing ("RESULTS\n" ^ cat_lines (map (Syntax.string_of_term ctxt o prop_of) thms)) + val qualified_thm_name = pred_names + |> map Long_Name.base_name + |> space_implode "_" + |> (fn s => Binding.qualify false s (Binding.name "strong_induct")) + + val attrs = + [ Attrib.internal (K (Rule_Cases.consumes 1)), + Attrib.internal (K (Rule_Cases.case_names rule_names)) ] + val _ = tracing ("RESULTS\n" ^ cat_lines (map (Syntax.string_of_term ctxt o prop_of) strong_ind_thms)) + val _ = tracing ("rule_names: " ^ commas rule_names) + val _ = tracing ("pred_names: " ^ commas pred_names) in ctxt - end + |> Local_Theory.note ((qualified_thm_name, attrs), strong_ind_thms) + |> snd + end in Proof.theorem NONE (after_qed ctxt) ((map o map) (rpair []) vc_compat_goals) ctxt'' end @@ -366,7 +455,7 @@ val avoid_trms = map2 read_avoids avoids_ordered intrs in - prove_strong_inductive rule_names avoid_trms raw_induct intrs ctxt + prove_strong_inductive names rule_names avoid_trms raw_induct intrs ctxt end *} @@ -422,6 +511,8 @@ nominal_inductive Acc . +thm Acc.strong_induct + section {* Typing *} nominal_datatype ty = @@ -441,8 +532,8 @@ inductive valid :: "(name \ ty) list \ bool" where - "valid []" -| "\atom x \ Gamma; valid Gamma\ \ valid ((x, T)#Gamma)" + v_Nil[intro]: "valid []" +| v_Cons[intro]: "\atom x \ Gamma; valid Gamma\ \ valid ((x, T)#Gamma)" inductive typing :: "(name\ty) list \ lam \ ty \ bool" ("_ \ _ : _" [60,60,60] 60) @@ -459,105 +550,66 @@ equivariance valid equivariance typing - nominal_inductive typing avoids t_Lam: "x" - (* | t_Var: "x" *) apply - apply(simp_all add: fresh_star_def ty_fresh lam.fresh)? done +thm typing.strong_induct +abbreviation + "sub_context" :: "(name \ ty) list \ (name \ ty) list \ bool" ("_ \ _" [60,60] 60) +where + "\1 \ \2 \ \x T. (x, T) \ set \1 \ (x, T) \ set \2" + +text {* Now it comes: The Weakening Lemma *} + +text {* + The first version is, after setting up the induction, + completely automatic except for use of atomize. *} -lemma - fixes c::"'a::fs" - assumes a: "\ \ t : T" - and a1: "\\ x T c. \valid \; (x, T) \ set \\ \ P c \ (Var x) T" - and a2: "\\ t1 T1 T2 t2 c. \\ \ t1 : T1 \ T2; \d. P d \ t1 T1 \ T2; \ \ t2 : T1; \d. P d \ t2 T1\ - \ P c \ (App t1 t2) T2" - and a3: "\x \ T1 t T2 c. \{atom x} \* c; atom x \ \; (x, T1) # \ \ t : T2; \d. P d ((x, T1) # \) t T2\ - \ P c \ (Lam x t) T1 \ T2" - shows "P c \ t T" -proof - - from a have "\p c. P c (p \ \) (p \ t) (p \ T)" - proof (induct) - case (t_Var \ x T p c) - then show ?case - apply - - apply(perm_strict_simp) - thm a1 - apply(rule a1) - apply(drule_tac p="p" in permute_boolI) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(rotate_tac 1) - apply(drule_tac p="p" in permute_boolI) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - done - next - case (t_App \ t1 T1 T2 t2 p c) - then show ?case - apply - - apply(perm_strict_simp) - apply(rule a2) - apply(drule_tac p="p" in permute_boolI) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(assumption) - apply(rotate_tac 2) - apply(drule_tac p="p" in permute_boolI) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(assumption) - done - next - case (t_Lam x \ T1 t T2 p c) - then show ?case - apply - - apply(subgoal_tac "\q. (q \ {p \ atom x}) \* c \ - supp (p \ \, p \ Lam x t, p \ (T1 \ T2)) \* q") - apply(erule exE) - apply(rule_tac t="p \ \" and s="(q + p) \ \" in subst) - apply(simp only: permute_plus) - apply(rule supp_perm_eq) - apply(simp add: supp_Pair fresh_star_Un) - apply(rule_tac t="p \ Lam x t" and s="(q + p) \ Lam x t" in subst) - apply(simp only: permute_plus) - apply(rule supp_perm_eq) - apply(simp add: supp_Pair fresh_star_Un) - apply(rule_tac t="p \ (T1 \ T2)" and s="(q + p) \ (T1 \ T2)" in subst) - apply(simp only: permute_plus) - apply(rule supp_perm_eq) - apply(simp add: supp_Pair fresh_star_Un) - (* apply(perm_simp) *) - apply(simp (no_asm) only: eqvts) - apply(rule a3) - apply(simp only: eqvts permute_plus) - apply(rule_tac p="- (q + p)" in permute_boolE) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(rule_tac p="- (q + p)" in permute_boolE) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(perm_strict_simp) - apply(simp only:) - thm at_set_avoiding2 - apply(rule at_set_avoiding2) - apply(simp add: finite_supp) - apply(simp add: finite_supp) - apply(simp add: finite_supp) - apply(rule_tac p="-p" in permute_boolE) - apply(perm_strict_simp add: permute_minus_cancel) - --"supplied by the user" - apply(simp add: fresh_star_Pair) - sorry - qed - then have "P c (0 \ \) (0 \ t) (0 \ T)" . - then show "P c \ t T" by simp -qed +lemma weakening_version2: + fixes \1 \2::"(name \ ty) list" + and t ::"lam" + and \ ::"ty" + assumes a: "\1 \ t : T" + and b: "valid \2" + and c: "\1 \ \2" + shows "\2 \ t : T" +using a b c +proof (nominal_induct \1 t T avoiding: \2 rule: typing.strong_induct) + case (t_Var \1 x T) (* variable case *) + have "\1 \ \2" by fact + moreover + have "valid \2" by fact + moreover + have "(x,T)\ set \1" by fact + ultimately show "\2 \ Var x : T" by auto +next + case (t_Lam x \1 T1 t T2) (* lambda case *) + have vc: "atom x \ \2" by fact (* variable convention *) + have ih: "\valid ((x, T1) # \2); (x, T1) # \1 \ (x, T1) # \2\ \ (x, T1) # \2 \ t : T2" by fact + have "\1 \ \2" by fact + then have "(x, T1) # \1 \ (x, T1) # \2" by simp + moreover + have "valid \2" by fact + then have "valid ((x, T1) # \2)" using vc by (simp add: v_Cons) + ultimately have "(x, T1) # \2 \ t : T2" using ih by simp + with vc show "\2 \ Lam x t : T1 \ T2" by auto +qed (auto) (* app case *) -*) +lemma weakening_version1: + fixes \1 \2::"(name \ ty) list" + assumes a: "\1 \ t : T" + and b: "valid \2" + and c: "\1 \ \2" + shows "\2 \ t : T" +using a b c +apply (nominal_induct \1 t T avoiding: \2 rule: typing.strong_induct) +apply (auto | atomize)+ +done + end