# HG changeset patch # User Christian Urban # Date 1294071672 0 # Node ID 0865caafbfe6072e420261b7da9ae8da39e9b0cd # Parent 64b4cb2c2bf805f4266546e2a2c05c5e2cf22f3e file with most of the strong rule induction development diff -r 64b4cb2c2bf8 -r 0865caafbfe6 Nominal/Ex/Typing.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex/Typing.thy Mon Jan 03 16:21:12 2011 +0000 @@ -0,0 +1,540 @@ +theory Lambda +imports "../Nominal2" +begin + + +atom_decl name + +nominal_datatype lam = + Var "name" +| App "lam" "lam" +| Lam x::"name" l::"lam" bind x in l + +thm lam.distinct +thm lam.induct +thm lam.exhaust lam.strong_exhaust +thm lam.fv_defs +thm lam.bn_defs +thm lam.perm_simps +thm lam.eq_iff +thm lam.fv_bn_eqvt +thm lam.size_eqvt + +ML {* +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}) + +fun minus_permute_elim p thm = + thm RS (Drule.instantiate' [] [NONE, SOME (mk_cminus p)] @{thm permute_boolI}) +*} + +ML {* +fun real_head_of (@{term Trueprop} $ t) = real_head_of t + | real_head_of (Const ("==>", _) $ _ $ t) = real_head_of t + | real_head_of (Const (@{const_name all}, _) $ Abs (_, _, t)) = real_head_of t + | real_head_of (Const (@{const_name All}, _) $ Abs (_, _, t)) = real_head_of t + | real_head_of (Const ("HOL.induct_forall", _) $ Abs (_, _, t)) = real_head_of t + | real_head_of t = head_of t +*} + +ML {* +fun mk_vc_compat (avoid, avoid_trm) prems concl_args params = + let + fun aux arg = arg + |> mk_fresh_star avoid_trm + |> HOLogic.mk_Trueprop + |> (curry Logic.list_implies) prems + |> (curry list_all_free) params + in + if null avoid then [] else map aux concl_args + end +*} + +ML {* +fun map_term prop f trm = + if prop trm + then f trm + else case trm of + (t1 $ t2) => map_term prop f t1 $ map_term prop f t2 + | Abs (x, T, t) => Abs (x, T, map_term prop f t) + | _ => trm +*} + +ML {* +fun add_p_c p (c, c_ty) trm = + let + val (P, args) = strip_comb trm + val (P_name, P_ty) = dest_Free P + val (ty_args, bool) = strip_type P_ty + val args' = map (mk_perm p) args + in + list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args') + |> (fn t => HOLogic.all_const c_ty $ lambda c t ) + |> (fn t => HOLogic.all_const @{typ perm} $ lambda p t) + end +*} + +ML {* +fun induct_forall_const T = Const ("HOL.induct_forall", (T --> @{typ bool}) --> @{typ bool}) +fun mk_induct_forall (a, T) t = induct_forall_const T $ Abs (a, T, t) +*} + +ML {* +fun add_c_prop qnt Ps (c, c_name, c_ty) trm = + let + fun add t = + let + val (P, args) = strip_comb t + val (P_name, P_ty) = dest_Free P + val (ty_args, bool) = strip_type P_ty + val args' = args + |> qnt ? map (incr_boundvars 1) + in + list_comb (Free (P_name, (c_ty :: ty_args) ---> bool), c :: args') + |> qnt ? mk_induct_forall (c_name, c_ty) + end + in + map_term (member (op =) Ps o head_of) add trm + end +*} + +ML {* +fun prep_prem Ps c_name c_ty (avoid, avoid_trm) (params, prems, concl) = + let + val prems' = prems + |> map (incr_boundvars 1) + |> map (add_c_prop true Ps (Bound 0, c_name, c_ty)) + + val avoid_trm' = avoid_trm + |> (curry list_abs_free) (params @ [(c_name, c_ty)]) + |> strip_abs_body + |> (fn t => mk_fresh_star_ty c_ty t (Bound 0)) + |> HOLogic.mk_Trueprop + + val prems'' = + if null avoid + then prems' + else avoid_trm' :: prems' + + val concl' = concl + |> incr_boundvars 1 + |> add_c_prop false Ps (Bound 0, c_name, c_ty) + in + mk_full_horn (params @ [(c_name, c_ty)]) prems'' concl' + end +*} + + +ML {* +fun same_name (Free (a1, _), Free (a2, _)) = (a1 = a2) + | same_name (Var (a1, _), Var (a2, _)) = (a1 = a2) + | same_name (Const (a1, _), Const (a2, _)) = (a1 = a2) + | same_name _ = false +*} + +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} [] +*} + +ML {* +val all_elims = + let + fun spec' ct = Drule.instantiate' [SOME (ctyp_of_term ct)] [NONE, SOME ct] @{thm spec} + in + fold (fn ct => fn th => th RS spec' ct) + end +*} + +ML {* +fun helper_tac flag prm p ctxt = + Subgoal.SUBPROOF (fn {context, prems, ...} => + let + 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})) + in + simp_tac (HOL_ss addsimps (prm' :: @{thms induct_forall_def })) 1 + end) ctxt +*} + +ML {* +fun non_binder_tac prem intr_cvars Ps ctxt = + Subgoal.SUBPROOF (fn {context, params, prems, ...} => + let + val thy = ProofContext.theory_of context + val (prms, p, _) = split_last2 (map snd params) + val prm_tys = map (fastype_of o term_of) prms + 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 + + (* for inductive-premises*) + fun tac1 prm = helper_tac true prm p context + + (* for non-inductive premises *) + fun tac2 prm = + EVERY' [ minus_permute_intro_tac p, + eqvt_stac context, + helper_tac false prm p context ] + + 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) ] + end) ctxt +*} + +ML {* +fun binder_tac prem intr_cvars Ps fresh_thms avoid avoid_trms ctxt = + Subgoal.FOCUS (fn {context, params, ...} => + let + val thy = ProofContext.theory_of context + val (prms, p, _) = split_last2 (map snd params) + val prm_tys = map (fastype_of o term_of) prms + 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 + in + Skip_Proof.cheat_tac thy + (* EVERY1 [rtac prem'] *) + end) ctxt +*} + +ML {* +fun case_tac ctxt fresh_thms Ps avoid avoid_trm intr_cvars prem = + let + val tac1 = non_binder_tac prem intr_cvars Ps ctxt + val tac2 = binder_tac prem intr_cvars 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 tac raw_induct fresh_thms Ps avoids avoid_trms intr_cvars {prems, context} = + let + val cases_tac = map4 (case_tac context fresh_thms Ps) avoids avoid_trms intr_cvars prems + in + EVERY1 [ DETERM o rtac raw_induct, RANGE cases_tac ] + end +*} + +ML {* +val normalise = @{lemma "(Q --> (!p c. P p c)) ==> (\c. Q ==> P (0::perm) c)" by simp} +*} + +ML {* +fun prove_strong_inductive rule_names avoids raw_induct intrs ctxt = + let + val thy = ProofContext.theory_of ctxt + val ((_, [raw_induct']), ctxt') = Variable.import true [raw_induct] ctxt + + val (ind_prems, ind_concl) = raw_induct' + |> prop_of + |> Logic.strip_horn + |>> map strip_full_horn + val params = map (fn (x, _, _) => x) ind_prems + val param_trms = (map o map) Free params + + val intr_vars_tys = map (fn t => rev (Term.add_vars (prop_of t) [])) intrs + val intr_vars = (map o map) fst intr_vars_tys + val intr_vars_substs = map2 (curry (op ~~)) intr_vars param_trms + val intr_cvars = (map o map) (cterm_of thy o Var) intr_vars_tys + + val (intr_prems, intr_concls) = intrs + |> map prop_of + |> map2 subst_Vars intr_vars_substs + |> map Logic.strip_horn + |> split_list + + val intr_concls_args = map (snd o strip_comb o HOLogic.dest_Trueprop) intr_concls + + val avoid_trms = avoids + |> (map o map) (setify ctxt') + |> map fold_union + + val vc_compat_goals = + map4 mk_vc_compat (avoids ~~ avoid_trms) intr_prems intr_concls_args params + + val ([c_name, a, p], ctxt'') = Variable.variant_fixes ["c", "'a", "p"] ctxt' + val c_ty = TFree (a, @{sort fs}) + val c = Free (c_name, c_ty) + val p = Free (p, @{typ perm}) + + val (preconds, ind_concls) = ind_concl + |> HOLogic.dest_Trueprop + |> HOLogic.dest_conj + |> map HOLogic.dest_imp + |> split_list + + val Ps = map (fst o strip_comb) ind_concls + + val ind_concl' = ind_concls + |> map (add_p_c p (c, c_ty)) + |> (curry (op ~~)) preconds + |> map HOLogic.mk_imp + |> fold_conj + |> HOLogic.mk_Trueprop + + val ind_prems' = ind_prems + |> map2 (prep_prem Ps c_name c_ty) (avoids ~~ avoid_trms) + + fun after_qed ctxt_outside fresh_thms ctxt = + let + val thms = Goal.prove ctxt [] ind_prems' ind_concl' + (tac raw_induct fresh_thms Ps avoids avoid_trms intr_cvars) + |> singleton (ProofContext.export ctxt ctxt_outside) + |> Datatype_Aux.split_conj_thm + |> map (fn thm => thm RS normalise) + |> map (asm_full_simplify (HOL_basic_ss addsimps @{thms permute_zero induct_rulify})) + |> 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)) + in + ctxt + end + in + Proof.theorem NONE (after_qed ctxt) ((map o map) (rpair []) vc_compat_goals) ctxt'' + end +*} + +ML {* +fun prove_strong_inductive_cmd (pred_name, avoids) ctxt = + let + val thy = ProofContext.theory_of ctxt; + val ({names, ...}, {raw_induct, intrs, ...}) = + Inductive.the_inductive ctxt (Sign.intern_const thy pred_name); + + val rule_names = + hd names + |> the o Induct.lookup_inductP ctxt + |> fst o Rule_Cases.get + |> map fst + + val _ = (case duplicates (op = o pairself fst) avoids of + [] => () + | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))) + + val _ = (case subtract (op =) rule_names (map fst avoids) of + [] => () + | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)) + + val avoids_ordered = order_default (op =) [] rule_names avoids + + fun read_avoids avoid_trms intr = + let + (* fixme hack *) + val (((_, ctrms), _), ctxt') = Variable.import true [intr] ctxt + val trms = map (term_of o snd) ctrms + val ctxt'' = fold Variable.declare_term trms ctxt' + in + map (Syntax.read_term ctxt'') avoid_trms + end + + val avoid_trms = map2 read_avoids avoids_ordered intrs + in + prove_strong_inductive rule_names avoid_trms raw_induct intrs ctxt + end +*} + +ML {* +(* outer syntax *) +local + structure P = Parse; + structure S = Scan + + val _ = Keyword.keyword "avoids" + + val single_avoid_parser = + P.name -- (P.$$$ ":" |-- P.and_list1 P.term) + + val avoids_parser = + S.optional (P.$$$ "avoids" |-- P.enum1 "|" single_avoid_parser) [] + + val main_parser = P.xname -- avoids_parser +in + val _ = + Outer_Syntax.local_theory_to_proof "nominal_inductive" + "prove strong induction theorem for inductive predicate involving nominal datatypes" + Keyword.thy_goal (main_parser >> prove_strong_inductive_cmd) +end +*} + +inductive + Acc :: "('a::pt \ 'a \ bool) \ 'a \ bool" +where + AccI: "(\y. R y x \ Acc R y) \ Acc R x" + +(* +equivariance Acc +*) + +lemma Acc_eqvt [eqvt]: + fixes p::"perm" + assumes a: "Acc R x" + shows "Acc (p \ R) (p \ x)" +using a +apply(induct) +apply(rule AccI) +apply(rotate_tac 1) +apply(drule_tac x="-p \ y" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(rule_tac p="p" in permute_boolE) +apply(perm_simp add: permute_minus_cancel) +apply(assumption) +apply(assumption) +done + + +nominal_inductive Acc . + +section {* Typing *} + +nominal_datatype ty = + TVar string +| TFun ty ty ("_ \ _") + +lemma ty_fresh: + fixes x::"name" + and T::"ty" + shows "atom x \ T" +apply (nominal_induct T rule: ty.strong_induct) +apply (simp_all add: ty.fresh pure_fresh) +done + + + +inductive + valid :: "(name \ ty) list \ bool" +where + "valid []" +| "\atom x \ Gamma; valid Gamma\ \ valid ((x, T)#Gamma)" + +inductive + typing :: "(name\ty) list \ lam \ ty \ bool" ("_ \ _ : _" [60,60,60] 60) +where + t_Var[intro]: "\valid \; (x, T) \ set \\ \ \ \ Var x : T" + | t_App[intro]: "\\ \ t1 : T1 \ T2; \ \ t2 : T1\ \ \ \ App t1 t2 : T2" + | t_Lam[intro]: "\atom x \ \; (x, T1) # \ \ t : T2\ \ \ \ Lam x t : T1 \ T2" + +thm typing.intros +thm typing.induct + + + +equivariance valid +equivariance typing + + +nominal_inductive typing + avoids t_Lam: "x" + apply - + apply(simp_all add: fresh_star_def ty_fresh lam.fresh)? + done + + + +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(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:) + 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 + +*) + + +end + + +