berghofe@22313: (* Title: HOL/Nominal/nominal_inductive.ML berghofe@22313: Author: Stefan Berghofer, TU Muenchen berghofe@22313: berghofe@22530: Infrastructure for proving equivariance and strong induction theorems berghofe@22530: for inductive predicates involving nominal datatypes. berghofe@22313: *) berghofe@22313: berghofe@22313: signature NOMINAL_INDUCTIVE = berghofe@22313: sig berghofe@30087: val prove_strong_ind: string -> (string * string list) list -> local_theory -> Proof.state berghofe@30087: val prove_eqvt: string -> string list -> local_theory -> local_theory berghofe@22313: end berghofe@22313: berghofe@22313: structure NominalInductive : NOMINAL_INDUCTIVE = berghofe@22313: struct berghofe@22313: wenzelm@59940: val inductive_forall_def = @{thm HOL.induct_forall_def}; urbanc@33772: val inductive_atomize = @{thms induct_atomize}; urbanc@33772: val inductive_rulify = @{thms induct_rulify}; berghofe@24570: wenzelm@41228: fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify []; berghofe@24570: wenzelm@51717: fun atomize_conv ctxt = wenzelm@41228: Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) wenzelm@51717: (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize); wenzelm@51717: fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)); wenzelm@24832: fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 wenzelm@51717: (Conv.params_conv ~1 (K (Conv.prems_conv ~1 (atomize_conv ctxt))) ctxt)); berghofe@24570: haftmann@44929: fun preds_of ps t = inter (op = o apsnd dest_Free) ps (Term.add_frees t []); berghofe@30087: wenzelm@39159: val fresh_prod = @{thm fresh_prod}; berghofe@22530: haftmann@44689: val perm_bool = mk_meta_eq @{thm perm_bool_def}; wenzelm@39159: val perm_boolI = @{thm perm_boolI}; berghofe@22313: val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb wenzelm@59582: (Drule.strip_imp_concl (Thm.cprop_of perm_boolI)))); berghofe@22313: berghofe@25824: fun mk_perm_bool pi th = th RS Drule.cterm_instantiate berghofe@25824: [(perm_boolI_pi, pi)] perm_boolI; berghofe@25824: wenzelm@38715: fun mk_perm_bool_simproc names = Simplifier.simproc_global_i wenzelm@59582: (Thm.theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn ctxt => haftmann@44868: fn Const (@{const_name Nominal.perm}, _) $ _ $ t => haftmann@36692: if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t)) berghofe@24570: then SOME perm_bool else NONE berghofe@24570: | _ => NONE); berghofe@24570: berghofe@22313: fun transp ([] :: _) = [] berghofe@22313: | transp xs = map hd xs :: transp (map tl xs); berghofe@22313: berghofe@22530: fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of berghofe@22530: (Const (s, T), ts) => (case strip_type T of berghofe@22530: (Ts, Type (tname, _)) => haftmann@31938: (case NominalDatatype.get_nominal_datatype thy tname of berghofe@22530: NONE => fold (add_binders thy i) ts bs berghofe@22530: | SOME {descr, index, ...} => (case AList.lookup op = berghofe@22530: (#3 (the (AList.lookup op = descr index))) s of berghofe@22530: NONE => fold (add_binders thy i) ts bs berghofe@22530: | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') => berghofe@22530: let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs' berghofe@22530: in (add_binders thy i u berghofe@22530: (fold (fn (u, T) => berghofe@22530: if exists (fn j => j < i) (loose_bnos u) then I wenzelm@59058: else insert (op aconv o apply2 fst) berghofe@22530: (incr_boundvars (~i) u, T)) cargs1 bs'), cargs2) berghofe@22530: end) cargs (bs, ts ~~ Ts)))) berghofe@22530: | _ => fold (add_binders thy i) ts bs) berghofe@22530: | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs)) berghofe@22530: | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs berghofe@22530: | add_binders thy i _ bs = bs; berghofe@22530: haftmann@38795: fun split_conj f names (Const (@{const_name HOL.conj}, _) $ p $ q) _ = (case head_of p of berghofe@24570: Const (name, _) => haftmann@36692: if member (op =) names name then SOME (f p q) else NONE berghofe@24570: | _ => NONE) berghofe@24570: | split_conj _ _ _ _ = NONE; berghofe@24570: berghofe@24570: fun strip_all [] t = t haftmann@38558: | strip_all (_ :: xs) (Const (@{const_name All}, _) $ Abs (s, T, t)) = strip_all xs t; berghofe@24570: berghofe@24570: (*********************************************************************) berghofe@24570: (* maps R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)) *) berghofe@27722: (* or ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t) *) berghofe@27722: (* to R ... & id (ALL z. P z (pi_1 o ... o pi_n o t)) *) berghofe@27722: (* or id (ALL z. P z (pi_1 o ... o pi_n o t)) *) berghofe@24570: (* *) berghofe@24570: (* where "id" protects the subformula from simplification *) berghofe@24570: (*********************************************************************) berghofe@24570: haftmann@38795: fun inst_conj_all names ps pis (Const (@{const_name HOL.conj}, _) $ p $ q) _ = berghofe@24570: (case head_of p of berghofe@24570: Const (name, _) => haftmann@36692: if member (op =) names name then SOME (HOLogic.mk_conj (p, haftmann@44868: Const (@{const_name Fun.id}, HOLogic.boolT --> HOLogic.boolT) $ berghofe@24570: (subst_bounds (pis, strip_all pis q)))) berghofe@24570: else NONE berghofe@24570: | _ => NONE) berghofe@24570: | inst_conj_all names ps pis t u = berghofe@24570: if member (op aconv) ps (head_of u) then haftmann@44868: SOME (Const (@{const_name Fun.id}, HOLogic.boolT --> HOLogic.boolT) $ berghofe@24570: (subst_bounds (pis, strip_all pis t))) berghofe@24570: else NONE berghofe@24570: | inst_conj_all _ _ _ _ _ = NONE; berghofe@24570: wenzelm@51717: fun inst_conj_all_tac ctxt k = EVERY berghofe@24570: [TRY (EVERY [etac conjE 1, rtac conjI 1, atac 1]), berghofe@24570: REPEAT_DETERM_N k (etac allE 1), wenzelm@51717: simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1]; berghofe@24570: berghofe@24570: fun map_term f t u = (case f t u of berghofe@24570: NONE => map_term' f t u | x => x) berghofe@24570: and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of berghofe@24570: (NONE, NONE) => NONE berghofe@24570: | (SOME t'', NONE) => SOME (t'' $ u) berghofe@24570: | (NONE, SOME u'') => SOME (t $ u'') berghofe@24570: | (SOME t'', SOME u'') => SOME (t'' $ u'')) berghofe@24570: | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of berghofe@24570: NONE => NONE berghofe@24570: | SOME t'' => SOME (Abs (s, T, t''))) berghofe@24570: | map_term' _ _ _ = NONE; berghofe@24570: berghofe@24570: (*********************************************************************) berghofe@24570: (* Prove F[f t] from F[t], where F is monotone *) berghofe@24570: (*********************************************************************) berghofe@24570: berghofe@24570: fun map_thm ctxt f tac monos opt th = berghofe@24570: let wenzelm@59582: val prop = Thm.prop_of th; berghofe@24570: fun prove t = berghofe@24570: Goal.prove ctxt [] [] t (fn _ => berghofe@24570: EVERY [cut_facts_tac [th] 1, etac rev_mp 1, wenzelm@59498: REPEAT_DETERM (FIRSTGOAL (resolve_tac ctxt monos)), berghofe@24570: REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))]) berghofe@24570: in Option.map prove (map_term f prop (the_default prop opt)) end; berghofe@24570: berghofe@27352: val eta_contract_cterm = Thm.dest_arg o Thm.cprop_of o Thm.eta_conversion; berghofe@27352: berghofe@25824: fun first_order_matchs pats objs = Thm.first_order_match berghofe@27352: (eta_contract_cterm (Conjunction.mk_conjunction_balanced pats), berghofe@27352: eta_contract_cterm (Conjunction.mk_conjunction_balanced objs)); berghofe@25824: berghofe@25824: fun first_order_mrs ths th = ths MRS wenzelm@59582: Thm.instantiate (first_order_matchs (cprems_of th) (map Thm.cprop_of ths)) th; berghofe@25824: berghofe@30087: fun prove_strong_ind s avoids ctxt = berghofe@22313: let wenzelm@42361: val thy = Proof_Context.theory_of ctxt; berghofe@25824: val ({names, ...}, {raw_induct, intrs, elims, ...}) = haftmann@31723: Inductive.the_inductive ctxt (Sign.intern_const thy s); haftmann@31723: val ind_params = Inductive.params_of raw_induct; wenzelm@24832: val raw_induct = atomize_induct ctxt raw_induct; berghofe@25824: val elims = map (atomize_induct ctxt) elims; haftmann@31723: val monos = Inductive.get_monos ctxt; urbanc@24571: val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt; haftmann@33040: val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of berghofe@22788: [] => () berghofe@22788: | xs => error ("Missing equivariance theorem for predicate(s): " ^ berghofe@22788: commas_quote xs)); nipkow@44045: val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the wenzelm@24861: (Induct.lookup_inductP ctxt (hd names))))); wenzelm@59582: val ([raw_induct'], ctxt') = Variable.import_terms false [Thm.prop_of raw_induct] ctxt; berghofe@22530: val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> berghofe@22530: HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb); berghofe@22530: val ps = map (fst o snd) concls; berghofe@22530: wenzelm@59058: val _ = (case duplicates (op = o apply2 fst) avoids of berghofe@22530: [] => () berghofe@22530: | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))); wenzelm@57884: val _ = avoids |> forall (fn (a, xs) => null (duplicates (op =) xs) orelse wenzelm@57884: error ("Duplicate variable names for case " ^ quote a)); haftmann@33040: val _ = (case subtract (op =) induct_cases (map fst avoids) of berghofe@22530: [] => () berghofe@22530: | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)); berghofe@25824: val avoids' = if null induct_cases then replicate (length intrs) ("", []) berghofe@25824: else map (fn name => berghofe@25824: (name, the_default [] (AList.lookup op = avoids name))) induct_cases; berghofe@22530: fun mk_avoids params (name, ps) = berghofe@22530: let val k = length params - 1 berghofe@22530: in map (fn x => case find_index (equal x o fst) params of berghofe@22530: ~1 => error ("No such variable in case " ^ quote name ^ berghofe@22530: " of inductive definition: " ^ quote x) berghofe@22530: | i => (Bound (k - i), snd (nth params i))) ps berghofe@22530: end; berghofe@22530: berghofe@22530: val prems = map (fn (prem, avoid) => berghofe@22530: let berghofe@22530: val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem); berghofe@22530: val concl = incr_boundvars 1 (Logic.strip_assums_concl prem); berghofe@22530: val params = Logic.strip_params prem berghofe@22530: in berghofe@22530: (params, berghofe@22530: fold (add_binders thy 0) (prems @ [concl]) [] @ berghofe@22530: map (apfst (incr_boundvars 1)) (mk_avoids params avoid), berghofe@22530: prems, strip_comb (HOLogic.dest_Trueprop concl)) berghofe@22530: end) (Logic.strip_imp_prems raw_induct' ~~ avoids'); berghofe@22530: berghofe@22530: val atomTs = distinct op = (maps (map snd o #2) prems); wenzelm@56253: val ind_sort = if null atomTs then @{sort type} wenzelm@36428: else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn T => Sign.intern_class thy wenzelm@36428: ("fs_" ^ Long_Name.base_name (fst (dest_Type T)))) atomTs)); wenzelm@43326: val (fs_ctxt_tyname, _) = Name.variant "'n" (Variable.names_of ctxt'); berghofe@30087: val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt'; berghofe@22530: val fsT = TFree (fs_ctxt_tyname, ind_sort); berghofe@22530: berghofe@24570: val inductive_forall_def' = Drule.instantiate' wenzelm@59621: [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def; berghofe@24570: berghofe@22530: fun lift_pred' t (Free (s, T)) ts = berghofe@22530: list_comb (Free (s, fsT --> T), t :: ts); berghofe@22530: val lift_pred = lift_pred' (Bound 0); berghofe@22530: berghofe@24570: fun lift_prem (t as (f $ u)) = berghofe@22530: let val (p, ts) = strip_comb t berghofe@22530: in haftmann@44868: if member (op =) ps p then HOLogic.mk_induct_forall fsT $ haftmann@44868: Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts)) berghofe@22530: else lift_prem f $ lift_prem u berghofe@22530: end berghofe@22530: | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t) berghofe@22530: | lift_prem t = t; berghofe@22530: berghofe@22530: fun mk_distinct [] = [] wenzelm@32952: | mk_distinct ((x, T) :: xs) = map_filter (fn (y, U) => berghofe@22530: if T = U then SOME (HOLogic.mk_Trueprop berghofe@22530: (HOLogic.mk_not (HOLogic.eq_const T $ x $ y))) berghofe@22530: else NONE) xs @ mk_distinct xs; berghofe@22530: berghofe@22530: fun mk_fresh (x, T) = HOLogic.mk_Trueprop haftmann@31938: (NominalDatatype.fresh_const T fsT $ x $ Bound 0); berghofe@22530: berghofe@22530: val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) => berghofe@22530: let berghofe@22530: val params' = params @ [("y", fsT)]; berghofe@22530: val prem = Logic.list_implies berghofe@22530: (map mk_fresh bvars @ mk_distinct bvars @ berghofe@22530: map (fn prem => berghofe@30087: if null (preds_of ps prem) then prem berghofe@22530: else lift_prem prem) prems, berghofe@22530: HOLogic.mk_Trueprop (lift_pred p ts)); wenzelm@29276: val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term prem params') berghofe@22530: in wenzelm@46218: (Logic.list_all (params', prem), (rev vs, subst_bounds (vs, prem))) berghofe@22530: end) prems); berghofe@22530: berghofe@22530: val ind_vars = blanchet@58112: (Old_Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ berghofe@22530: map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; berghofe@22530: val ind_Ts = rev (map snd ind_vars); berghofe@22530: berghofe@22530: val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj berghofe@22530: (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, berghofe@22530: HOLogic.list_all (ind_vars, lift_pred p haftmann@31938: (map (fold_rev (NominalDatatype.mk_perm ind_Ts) berghofe@22530: (map Bound (length atomTs downto 1))) ts)))) concls)); berghofe@22530: berghofe@22530: val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj berghofe@22530: (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, berghofe@22530: lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls)); berghofe@22530: berghofe@22530: val vc_compat = map (fn (params, bvars, prems, (p, ts)) => wenzelm@46218: map (fn q => Logic.list_all (params, incr_boundvars ~1 (Logic.list_implies wenzelm@32952: (map_filter (fn prem => berghofe@30087: if null (preds_of ps prem) then SOME prem berghofe@24570: else map_term (split_conj (K o I) names) prem prem) prems, q)))) berghofe@22530: (mk_distinct bvars @ berghofe@22530: maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop haftmann@31938: (NominalDatatype.fresh_const U T $ u $ t)) bvars) berghofe@22530: (ts ~~ binder_types (fastype_of p)))) prems; berghofe@22530: wenzelm@39557: val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; wenzelm@39557: val pt2_atoms = map (fn aT => Global_Theory.get_thm thy wenzelm@30364: ("pt_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "2")) atomTs; wenzelm@54895: val eqvt_ss = simpset_of (put_simpset HOL_basic_ss (Proof_Context.init_global thy) berghofe@27352: addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) haftmann@44868: addsimprocs [mk_perm_bool_simproc [@{const_name Fun.id}], wenzelm@51717: NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]); wenzelm@39557: val fresh_bij = Global_Theory.get_thms thy "fresh_bij"; wenzelm@39557: val perm_bij = Global_Theory.get_thms thy "perm_bij"; wenzelm@39557: val fs_atoms = map (fn aT => Global_Theory.get_thm thy wenzelm@30364: ("fs_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "1")) atomTs; wenzelm@39557: val exists_fresh' = Global_Theory.get_thms thy "exists_fresh'"; wenzelm@39557: val fresh_atm = Global_Theory.get_thms thy "fresh_atm"; wenzelm@39557: val swap_simps = Global_Theory.get_thms thy "swap_simps"; wenzelm@39557: val perm_fresh_fresh = Global_Theory.get_thms thy "perm_fresh_fresh"; berghofe@22530: berghofe@22530: fun obtain_fresh_name ts T (freshs1, freshs2, ctxt) = berghofe@22530: let berghofe@27722: (** protect terms to avoid that fresh_prod interferes with **) berghofe@22530: (** pairs used in introduction rules of inductive predicate **) berghofe@22530: fun protect t = haftmann@44868: let val T = fastype_of t in Const (@{const_name Fun.id}, T --> T) $ t end; berghofe@22530: val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1); berghofe@22530: val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop berghofe@22530: (HOLogic.exists_const T $ Abs ("x", T, haftmann@31938: NominalDatatype.fresh_const T (fastype_of p) $ berghofe@22530: Bound 0 $ p))) berghofe@22530: (fn _ => EVERY wenzelm@59498: [resolve_tac ctxt exists_fresh' 1, wenzelm@59498: resolve_tac ctxt fs_atoms 1]); wenzelm@32202: val (([(_, cx)], ths), ctxt') = Obtain.result wenzelm@51717: (fn ctxt' => EVERY berghofe@22530: [etac exE 1, wenzelm@51717: full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1, wenzelm@51717: full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1, berghofe@22530: REPEAT (etac conjE 1)]) berghofe@22530: [ex] ctxt wenzelm@59582: in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end; berghofe@22530: berghofe@30087: fun mk_ind_proof ctxt' thss = berghofe@30087: Goal.prove ctxt' [] prems' concl' (fn {prems = ihyps, context = ctxt} => berghofe@22530: let val th = Goal.prove ctxt [] [] concl (fn {context, ...} => berghofe@22530: rtac raw_induct 1 THEN berghofe@22530: EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) => wenzelm@51717: [REPEAT (rtac allI 1), simp_tac (put_simpset eqvt_ss context) 1, berghofe@22530: SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} => berghofe@22530: let berghofe@22530: val (params', (pis, z)) = wenzelm@59582: chop (length params - length atomTs - 1) (map (Thm.term_of o #2) params) ||> berghofe@22530: split_last; berghofe@22530: val bvars' = map berghofe@22530: (fn (Bound i, T) => (nth params' (length params' - i), T) berghofe@22530: | (t, T) => (t, T)) bvars; berghofe@22530: val pi_bvars = map (fn (t, _) => haftmann@31938: fold_rev (NominalDatatype.mk_perm []) pis t) bvars'; wenzelm@59582: val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl)); berghofe@22530: val (freshs1, freshs2, ctxt'') = fold berghofe@22530: (obtain_fresh_name (ts @ pi_bvars)) berghofe@22530: (map snd bvars') ([], [], ctxt'); haftmann@31938: val freshs2' = NominalDatatype.mk_not_sym freshs2; haftmann@31938: val pis' = map NominalDatatype.perm_of_pair (pi_bvars ~~ freshs1); berghofe@24570: fun concat_perm pi1 pi2 = berghofe@24570: let val T = fastype_of pi1 berghofe@24570: in if T = fastype_of pi2 then wenzelm@56253: Const (@{const_name append}, T --> T --> T) $ pi1 $ pi2 berghofe@24570: else pi2 berghofe@24570: end; berghofe@24570: val pis'' = fold (concat_perm #> map) pis' pis; wenzelm@59582: val env = Pattern.first_order_match thy (ihypt, Thm.prop_of ihyp) berghofe@22530: (Vartab.empty, Vartab.empty); wenzelm@59621: val ihyp' = Thm.instantiate ([], map (apply2 (Thm.global_cterm_of thy)) wenzelm@32035: (map (Envir.subst_term env) vs ~~ haftmann@31938: map (fold_rev (NominalDatatype.mk_perm []) berghofe@22530: (rev pis' @ pis)) params' @ [z])) ihyp; berghofe@24570: fun mk_pi th = wenzelm@54983: Simplifier.simplify (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}] haftmann@31938: addsimprocs [NominalDatatype.perm_simproc]) wenzelm@54983: (Simplifier.simplify (put_simpset eqvt_ss ctxt') wenzelm@59621: (fold_rev (mk_perm_bool o Thm.global_cterm_of thy) berghofe@25824: (rev pis' @ pis) th)); berghofe@24570: val (gprems1, gprems2) = split_list berghofe@24570: (map (fn (th, t) => berghofe@30087: if null (preds_of ps t) then (SOME th, mk_pi th) berghofe@24570: else wenzelm@54983: (map_thm ctxt' (split_conj (K o I) names) berghofe@24570: (etac conjunct1 1) monos NONE th, wenzelm@54983: mk_pi (the (map_thm ctxt' (inst_conj_all names ps (rev pis'')) wenzelm@54983: (inst_conj_all_tac ctxt' (length pis'')) monos (SOME t) th)))) wenzelm@32952: (gprems ~~ oprems)) |>> map_filter I; berghofe@22530: val vc_compat_ths' = map (fn th => berghofe@22530: let berghofe@25824: val th' = first_order_mrs gprems1 th; wenzelm@59582: val (bop, lhs, rhs) = (case Thm.concl_of th' of berghofe@22530: _ $ (fresh $ lhs $ rhs) => berghofe@22530: (fn t => fn u => fresh $ t $ u, lhs, rhs) berghofe@22530: | _ $ (_ $ (_ $ lhs $ rhs)) => berghofe@22530: (curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs)); berghofe@22530: val th'' = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop haftmann@31938: (bop (fold_rev (NominalDatatype.mk_perm []) pis lhs) haftmann@31938: (fold_rev (NominalDatatype.mk_perm []) pis rhs))) wenzelm@51717: (fn _ => simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps berghofe@22530: (fresh_bij @ perm_bij)) 1 THEN rtac th' 1) wenzelm@51717: in Simplifier.simplify (put_simpset eqvt_ss ctxt'' addsimps fresh_atm) th'' end) berghofe@22530: vc_compat_ths; haftmann@31938: val vc_compat_ths'' = NominalDatatype.mk_not_sym vc_compat_ths'; berghofe@27449: (** Since swap_simps simplifies (pi :: 'a prm) o (x :: 'b) to x **) berghofe@27449: (** we have to pre-simplify the rewrite rules **) wenzelm@51717: val swap_simps_simpset = put_simpset HOL_ss ctxt'' addsimps swap_simps @ wenzelm@51717: map (Simplifier.simplify (put_simpset HOL_ss ctxt'' addsimps swap_simps)) berghofe@22530: (vc_compat_ths'' @ freshs2'); berghofe@22530: val th = Goal.prove ctxt'' [] [] berghofe@22530: (HOLogic.mk_Trueprop (list_comb (P $ hd ts, haftmann@31938: map (fold (NominalDatatype.mk_perm []) pis') (tl ts)))) wenzelm@51717: (fn _ => EVERY ([simp_tac (put_simpset eqvt_ss ctxt'') 1, rtac ihyp' 1, wenzelm@59582: REPEAT_DETERM_N (Thm.nprems_of ihyp - length gprems) wenzelm@51717: (simp_tac swap_simps_simpset 1), berghofe@22530: REPEAT_DETERM_N (length gprems) wenzelm@51717: (simp_tac (put_simpset HOL_basic_ss ctxt'' berghofe@27847: addsimps [inductive_forall_def'] haftmann@31938: addsimprocs [NominalDatatype.perm_simproc]) 1 THEN wenzelm@59498: resolve_tac ctxt'' gprems2 1)])); wenzelm@59582: val final = Goal.prove ctxt'' [] [] (Thm.term_of concl) wenzelm@51717: (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss ctxt'' berghofe@22530: addsimps vc_compat_ths'' @ freshs2' @ berghofe@22530: perm_fresh_fresh @ fresh_atm) 1); wenzelm@42361: val final' = Proof_Context.export ctxt'' ctxt' [final]; wenzelm@59498: in resolve_tac ctxt' final' 1 end) context 1]) berghofe@22530: (prems ~~ thss ~~ ihyps ~~ prems''))) berghofe@22530: in berghofe@22530: cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN wenzelm@59498: REPEAT (REPEAT (resolve_tac ctxt [conjI, impI] 1) THEN wenzelm@27228: etac impE 1 THEN atac 1 THEN REPEAT (etac @{thm allE_Nil} 1) THEN wenzelm@51717: asm_full_simp_tac ctxt 1) wenzelm@42361: end) |> singleton (Proof_Context.export ctxt' ctxt); berghofe@22530: berghofe@25824: (** strong case analysis rule **) berghofe@25824: berghofe@25824: val cases_prems = map (fn ((name, avoids), rule) => berghofe@25824: let wenzelm@59582: val ([rule'], ctxt') = Variable.import_terms false [Thm.prop_of rule] ctxt; berghofe@30087: val prem :: prems = Logic.strip_imp_prems rule'; berghofe@30087: val concl = Logic.strip_imp_concl rule' berghofe@25824: in berghofe@25824: (prem, berghofe@25824: List.drop (snd (strip_comb (HOLogic.dest_Trueprop prem)), length ind_params), berghofe@25824: concl, berghofe@30087: fold_map (fn (prem, (_, avoid)) => fn ctxt => berghofe@25824: let berghofe@25824: val prems = Logic.strip_assums_hyp prem; berghofe@25824: val params = Logic.strip_params prem; berghofe@25824: val bnds = fold (add_binders thy 0) prems [] @ mk_avoids params avoid; berghofe@30087: fun mk_subst (p as (s, T)) (i, j, ctxt, ps, qs, is, ts) = berghofe@25824: if member (op = o apsnd fst) bnds (Bound i) then berghofe@25824: let berghofe@30087: val ([s'], ctxt') = Variable.variant_fixes [s] ctxt; berghofe@25824: val t = Free (s', T) berghofe@30087: in (i + 1, j, ctxt', ps, (t, T) :: qs, i :: is, t :: ts) end berghofe@30087: else (i + 1, j + 1, ctxt, p :: ps, qs, is, Bound j :: ts); berghofe@30087: val (_, _, ctxt', ps, qs, is, ts) = fold_rev mk_subst params berghofe@30087: (0, 0, ctxt, [], [], [], []) berghofe@25824: in berghofe@30087: ((ps, qs, is, map (curry subst_bounds (rev ts)) prems), ctxt') berghofe@30087: end) (prems ~~ avoids) ctxt') berghofe@25824: end) haftmann@31723: (Inductive.partition_rules' raw_induct (intrs ~~ avoids') ~~ berghofe@30087: elims); berghofe@25824: berghofe@25824: val cases_prems' = berghofe@30087: map (fn (prem, args, concl, (prems, _)) => berghofe@25824: let berghofe@25824: fun mk_prem (ps, [], _, prems) = wenzelm@46218: Logic.list_all (ps, Logic.list_implies (prems, concl)) berghofe@25824: | mk_prem (ps, qs, _, prems) = wenzelm@46218: Logic.list_all (ps, Logic.mk_implies berghofe@25824: (Logic.list_implies berghofe@25824: (mk_distinct qs @ berghofe@25824: maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop haftmann@31938: (NominalDatatype.fresh_const T (fastype_of u) $ t $ u)) berghofe@25824: args) qs, berghofe@25824: HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj berghofe@25824: (map HOLogic.dest_Trueprop prems))), berghofe@25824: concl)) berghofe@25824: in map mk_prem prems end) cases_prems; berghofe@25824: wenzelm@54895: val cases_eqvt_simpset = put_simpset HOL_ss (Proof_Context.init_global thy) berghofe@27449: addsimps eqvt_thms @ swap_simps @ perm_pi_simp berghofe@27352: addsimprocs [NominalPermeq.perm_simproc_app, berghofe@27352: NominalPermeq.perm_simproc_fun]; berghofe@27352: berghofe@27352: val simp_fresh_atm = map wenzelm@54895: (Simplifier.simplify (put_simpset HOL_basic_ss (Proof_Context.init_global thy) wenzelm@54895: addsimps fresh_atm)); berghofe@25824: berghofe@30087: fun mk_cases_proof ((((name, thss), elim), (prem, args, concl, (prems, ctxt'))), berghofe@25824: prems') = berghofe@30087: (name, Goal.prove ctxt' [] (prem :: prems') concl wenzelm@26711: (fn {prems = hyp :: hyps, context = ctxt1} => berghofe@25824: EVERY (rtac (hyp RS elim) 1 :: berghofe@25824: map (fn (((_, vc_compat_ths), case_hyp), (_, qs, is, _)) => berghofe@25824: SUBPROOF (fn {prems = case_hyps, params, context = ctxt2, concl, ...} => berghofe@25824: if null qs then berghofe@25824: rtac (first_order_mrs case_hyps case_hyp) 1 berghofe@25824: else berghofe@25824: let wenzelm@59582: val params' = map (Thm.term_of o #2 o nth (rev params)) is; berghofe@25824: val tab = params' ~~ map fst qs; berghofe@25824: val (hyps1, hyps2) = chop (length args) case_hyps; berghofe@25824: (* turns a = t and [x1 # t, ..., xn # t] *) berghofe@25824: (* into [x1 # a, ..., xn # a] *) berghofe@25824: fun inst_fresh th' ths = berghofe@25824: let val (ths1, ths2) = chop (length qs) ths berghofe@25824: in berghofe@25824: (map (fn th => berghofe@25824: let berghofe@25824: val (cf, ct) = wenzelm@59582: Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th)); berghofe@25824: val arg_cong' = Drule.instantiate' wenzelm@59586: [SOME (Thm.ctyp_of_cterm ct)] berghofe@25824: [NONE, SOME ct, SOME cf] (arg_cong RS iffD2); berghofe@25824: val inst = Thm.first_order_match (ct, wenzelm@59582: Thm.dest_arg (Thm.dest_arg (Thm.cprop_of th'))) berghofe@25824: in [th', th] MRS Thm.instantiate inst arg_cong' berghofe@25824: end) ths1, berghofe@25824: ths2) berghofe@25824: end; berghofe@25824: val (vc_compat_ths1, vc_compat_ths2) = berghofe@25824: chop (length vc_compat_ths - length args * length qs) berghofe@25824: (map (first_order_mrs hyps2) vc_compat_ths); berghofe@25824: val vc_compat_ths' = haftmann@31938: NominalDatatype.mk_not_sym vc_compat_ths1 @ berghofe@25824: flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2)); berghofe@25824: val (freshs1, freshs2, ctxt3) = fold berghofe@25824: (obtain_fresh_name (args @ map fst qs @ params')) berghofe@25824: (map snd qs) ([], [], ctxt2); haftmann@31938: val freshs2' = NominalDatatype.mk_not_sym freshs2; haftmann@31938: val pis = map (NominalDatatype.perm_of_pair) berghofe@25824: ((freshs1 ~~ map fst qs) @ (params' ~~ freshs1)); wenzelm@59621: val mk_pis = fold_rev mk_perm_bool (map (Thm.global_cterm_of thy) pis); wenzelm@59621: val obj = Thm.global_cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms berghofe@25824: (fn x as Free _ => haftmann@36692: if member (op =) args x then x berghofe@25824: else (case AList.lookup op = tab x of berghofe@25824: SOME y => y haftmann@31938: | NONE => fold_rev (NominalDatatype.mk_perm []) pis x) wenzelm@59582: | x => x) o HOLogic.dest_Trueprop o Thm.prop_of) case_hyps)); berghofe@25824: val inst = Thm.first_order_match (Thm.dest_arg berghofe@25824: (Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj); wenzelm@59582: val th = Goal.prove ctxt3 [] [] (Thm.term_of concl) berghofe@25824: (fn {context = ctxt4, ...} => berghofe@25824: rtac (Thm.instantiate inst case_hyp) 1 THEN berghofe@25824: SUBPROOF (fn {prems = fresh_hyps, ...} => berghofe@25824: let haftmann@31938: val fresh_hyps' = NominalDatatype.mk_not_sym fresh_hyps; wenzelm@51717: val case_simpset = cases_eqvt_simpset addsimps freshs2' @ berghofe@27352: simp_fresh_atm (vc_compat_ths' @ fresh_hyps'); wenzelm@51717: val fresh_fresh_simpset = case_simpset addsimps perm_fresh_fresh; berghofe@25824: val hyps1' = map wenzelm@51717: (mk_pis #> Simplifier.simplify fresh_fresh_simpset) hyps1; berghofe@25824: val hyps2' = map wenzelm@51717: (mk_pis #> Simplifier.simplify case_simpset) hyps2; berghofe@27449: val case_hyps' = hyps1' @ hyps2' berghofe@25824: in wenzelm@51717: simp_tac case_simpset 1 THEN berghofe@25824: REPEAT_DETERM (TRY (rtac conjI 1) THEN wenzelm@59498: resolve_tac ctxt4 case_hyps' 1) berghofe@25824: end) ctxt4 1) wenzelm@42361: val final = Proof_Context.export ctxt3 ctxt2 [th] wenzelm@59498: in resolve_tac ctxt2 final 1 end) ctxt1 1) berghofe@30087: (thss ~~ hyps ~~ prems))) |> wenzelm@42361: singleton (Proof_Context.export ctxt' ctxt)) berghofe@25824: berghofe@22530: in berghofe@30087: ctxt'' |> wenzelm@50771: Proof.theorem NONE (fn thss => fn ctxt => (* FIXME ctxt/ctxt' should be called lthy/lthy' *) berghofe@22530: let wenzelm@30364: val rec_name = space_implode "_" (map Long_Name.base_name names); wenzelm@30223: val rec_qualified = Binding.qualify false rec_name; wenzelm@33368: val ind_case_names = Rule_Cases.case_names induct_cases; haftmann@31723: val induct_cases' = Inductive.partition_rules' raw_induct berghofe@25824: (intrs ~~ induct_cases); wenzelm@51717: val thss' = map (map (atomize_intr ctxt)) thss; haftmann@31723: val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss'); berghofe@24570: val strong_raw_induct = wenzelm@51717: mk_ind_proof ctxt thss' |> Inductive.rulify ctxt; wenzelm@51717: val strong_cases = map (mk_cases_proof ##> Inductive.rulify ctxt) berghofe@25824: (thsss ~~ elims ~~ cases_prems ~~ cases_prems'); wenzelm@50771: val strong_induct_atts = wenzelm@50771: map (Attrib.internal o K) wenzelm@50771: [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))]; berghofe@22530: val strong_induct = wenzelm@50771: if length names > 1 then strong_raw_induct wenzelm@50771: else strong_raw_induct RSN (2, rev_mp); wenzelm@33671: val ((_, [strong_induct']), ctxt') = ctxt |> Local_Theory.note wenzelm@50771: ((rec_qualified (Binding.name "strong_induct"), strong_induct_atts), [strong_induct]); berghofe@22530: val strong_inducts = wenzelm@33670: Project_Rule.projects ctxt (1 upto length names) strong_induct'; berghofe@22530: in berghofe@30087: ctxt' |> wenzelm@50771: Local_Theory.notes wenzelm@50771: [((rec_qualified (Binding.name "strong_inducts"), []), wenzelm@50771: strong_inducts |> map (fn th => ([th], wenzelm@50771: [Attrib.internal (K ind_case_names), wenzelm@50771: Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd |> wenzelm@33671: Local_Theory.notes (map (fn ((name, elim), (_, cases)) => wenzelm@30435: ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "strong_cases"), wenzelm@33368: [Attrib.internal (K (Rule_Cases.case_names (map snd cases))), wenzelm@50771: Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of elim)))]), [([elim], [])])) berghofe@30087: (strong_cases ~~ induct_cases')) |> snd berghofe@30087: end) berghofe@24570: (map (map (rulify_term thy #> rpair [])) vc_compat) berghofe@22530: end; berghofe@22530: wenzelm@54991: fun prove_eqvt s xatoms ctxt = (* FIXME ctxt should be called lthy *) berghofe@22530: let wenzelm@42361: val thy = Proof_Context.theory_of ctxt; berghofe@22788: val ({names, ...}, {raw_induct, intrs, elims, ...}) = haftmann@31723: Inductive.the_inductive ctxt (Sign.intern_const thy s); wenzelm@24832: val raw_induct = atomize_induct ctxt raw_induct; wenzelm@24832: val elims = map (atomize_induct ctxt) elims; wenzelm@51717: val intrs = map (atomize_intr ctxt) intrs; haftmann@31723: val monos = Inductive.get_monos ctxt; haftmann@31723: val intrs' = Inductive.unpartition_rules intrs berghofe@22788: (map (fn (((s, ths), (_, k)), th) => haftmann@31723: (s, ths ~~ Inductive.infer_intro_vars th k ths)) haftmann@31723: (Inductive.partition_rules raw_induct intrs ~~ haftmann@31723: Inductive.arities_of raw_induct ~~ elims)); haftmann@31723: val k = length (Inductive.params_of raw_induct); berghofe@22730: val atoms' = NominalAtoms.atoms_of thy; berghofe@22730: val atoms = berghofe@22730: if null xatoms then atoms' else berghofe@22730: let val atoms = map (Sign.intern_type thy) xatoms berghofe@22730: in berghofe@22730: (case duplicates op = atoms of berghofe@22730: [] => () berghofe@22730: | xs => error ("Duplicate atoms: " ^ commas xs); haftmann@33040: case subtract (op =) atoms' atoms of berghofe@22730: [] => () berghofe@22730: | xs => error ("No such atoms: " ^ commas xs); berghofe@22730: atoms) berghofe@22730: end; wenzelm@39557: val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; berghofe@30087: val (([t], [pi]), ctxt') = ctxt |> wenzelm@59582: Variable.import_terms false [Thm.concl_of raw_induct] ||>> berghofe@30087: Variable.variant_fixes ["pi"]; wenzelm@54991: val eqvt_simpset = put_simpset HOL_basic_ss ctxt' addsimps wenzelm@54991: (NominalThmDecls.get_eqvt_thms ctxt' @ perm_pi_simp) addsimprocs wenzelm@54991: [mk_perm_bool_simproc names, wenzelm@54991: NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; berghofe@22313: val ps = map (fst o HOLogic.dest_imp) berghofe@22313: (HOLogic.dest_conj (HOLogic.dest_Trueprop t)); wenzelm@54991: fun eqvt_tac pi (intr, vs) st = berghofe@22544: let berghofe@30087: fun eqvt_err s = wenzelm@59582: let val ([t], ctxt'') = Variable.import_terms true [Thm.prop_of intr] ctxt' berghofe@30087: in error ("Could not prove equivariance for introduction rule\n" ^ wenzelm@54991: Syntax.string_of_term ctxt'' t ^ "\n" ^ s) berghofe@30087: end; wenzelm@54983: val res = SUBPROOF (fn {context = ctxt'', prems, params, ...} => berghofe@22788: let wenzelm@54983: val prems' = map (fn th => the_default th (map_thm ctxt'' berghofe@24570: (split_conj (K I) names) (etac conjunct2 1) monos NONE th)) prems; wenzelm@51717: val prems'' = map (fn th => Simplifier.simplify eqvt_simpset wenzelm@59621: (mk_perm_bool (Thm.global_cterm_of thy pi) th)) prems'; wenzelm@59621: val intr' = Drule.cterm_instantiate (map (Thm.global_cterm_of thy) vs ~~ wenzelm@59621: map (Thm.global_cterm_of thy o NominalDatatype.mk_perm [] pi o Thm.term_of o #2) params) berghofe@22788: intr wenzelm@59498: in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac ctxt'' prems'')) 1 berghofe@30087: end) ctxt' 1 st berghofe@22544: in berghofe@22544: case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of berghofe@22544: NONE => eqvt_err ("Rule does not match goal\n" ^ wenzelm@59582: Syntax.string_of_term ctxt' (hd (Thm.prems_of st))) berghofe@22544: | SOME (th, _) => Seq.single th berghofe@22544: end; berghofe@22313: val thss = map (fn atom => berghofe@25824: let val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, []))) berghofe@22530: in map (fn th => zero_var_indexes (th RS mp)) blanchet@58112: (Old_Datatype_Aux.split_conj_thm (Goal.prove ctxt' [] [] berghofe@22313: (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p => berghofe@30087: let berghofe@30087: val (h, ts) = strip_comb p; berghofe@30087: val (ts1, ts2) = chop k ts berghofe@30087: in berghofe@30087: HOLogic.mk_imp (p, list_comb (h, ts1 @ haftmann@31938: map (NominalDatatype.mk_perm [] pi') ts2)) berghofe@30087: end) ps))) wenzelm@54991: (fn _ => EVERY (rtac raw_induct 1 :: map (fn intr_vs => wenzelm@51717: full_simp_tac eqvt_simpset 1 THEN wenzelm@54991: eqvt_tac pi' intr_vs) intrs')) |> wenzelm@42361: singleton (Proof_Context.export ctxt' ctxt))) berghofe@22544: end) atoms berghofe@22544: in berghofe@30087: ctxt |> wenzelm@33671: Local_Theory.notes (map (fn (name, ths) => wenzelm@30435: ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "eqvt"), berghofe@30087: [Attrib.internal (K NominalThmDecls.eqvt_add)]), [(ths, [])])) berghofe@30087: (names ~~ transp thss)) |> snd berghofe@22544: end; berghofe@22313: berghofe@22313: berghofe@22313: (* outer syntax *) berghofe@22313: wenzelm@24867: val _ = wenzelm@59936: Outer_Syntax.local_theory_to_proof @{command_keyword nominal_inductive} wenzelm@36960: "prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes" wenzelm@46949: (Parse.xname -- Scan.optional (@{keyword "avoids"} |-- Parse.and_list1 (Parse.name -- wenzelm@46949: (@{keyword ":"} |-- Scan.repeat1 Parse.name))) [] >> (fn (name, avoids) => berghofe@30087: prove_strong_ind name avoids)); berghofe@22313: wenzelm@24867: val _ = wenzelm@59936: Outer_Syntax.local_theory @{command_keyword equivariance} wenzelm@46961: "prove equivariance for inductive predicate involving nominal datatypes" wenzelm@46949: (Parse.xname -- Scan.optional (@{keyword "["} |-- Parse.list1 Parse.name --| @{keyword "]"}) [] >> berghofe@30087: (fn (name, atoms) => prove_eqvt name atoms)); berghofe@22530: berghofe@22313: end