berghofe@30092: (* Author: Christian Urban and Makarius wenzelm@18283: wenzelm@18288: The nominal induct proof method. wenzelm@18283: *) wenzelm@18283: wenzelm@18283: structure NominalInduct: wenzelm@18283: sig haftmann@29581: val nominal_induct_tac: Proof.context -> (binding option * term) option list list -> wenzelm@18583: (string * typ) list -> (string * typ) list list -> thm list -> wenzelm@33368: thm list -> int -> Rule_Cases.cases_tactic wenzelm@30549: val nominal_induct_method: (Proof.context -> Proof.method) context_parser wenzelm@18283: end = wenzelm@18283: struct wenzelm@18283: wenzelm@18288: (* proper tuples -- nested left *) wenzelm@18283: wenzelm@18288: fun tupleT Ts = HOLogic.unitT |> fold (fn T => fn U => HOLogic.mk_prodT (U, T)) Ts; wenzelm@18288: fun tuple ts = HOLogic.unit |> fold (fn t => fn u => HOLogic.mk_prod (u, t)) ts; wenzelm@18288: wenzelm@18288: fun tuple_fun Ts (xi, T) = wenzelm@18288: Library.funpow (length Ts) HOLogic.mk_split wenzelm@18288: (Var (xi, (HOLogic.unitT :: Ts) ---> Term.range_type T)); wenzelm@18283: wenzelm@18288: val split_all_tuples = wenzelm@18288: Simplifier.full_simplify (HOL_basic_ss addsimps berghofe@30092: [split_conv, split_paired_all, unit_all_eq1, @{thm fresh_unit_elim}, @{thm fresh_prod_elim}] @ berghofe@30092: @{thms fresh_star_unit_elim} @ @{thms fresh_star_prod_elim}); wenzelm@18283: wenzelm@18288: wenzelm@18297: (* prepare rule *) wenzelm@18288: wenzelm@19903: fun inst_mutual_rule ctxt insts avoiding rules = wenzelm@18283: let wenzelm@33368: val (nconcls, joined_rule) = Rule_Cases.strict_mutual_rule ctxt rules; urbanc@22072: val concls = Logic.dest_conjunctions (Thm.concl_of joined_rule); wenzelm@33368: val (cases, consumes) = Rule_Cases.get joined_rule; wenzelm@18583: wenzelm@18583: val l = length rules; wenzelm@18583: val _ = wenzelm@18583: if length insts = l then () wenzelm@18583: else error ("Bad number of instantiations for " ^ string_of_int l ^ " rules"); wenzelm@18288: urbanc@22072: fun subst inst concl = wenzelm@18583: let wenzelm@24830: val vars = Induct.vars_of concl; wenzelm@18583: val m = length vars and n = length inst; wenzelm@18583: val _ = if m >= n + 2 then () else error "Too few variables in conclusion of rule"; wenzelm@18583: val P :: x :: ys = vars; haftmann@33955: val zs = (uncurry drop) (m - n - 2, ys); wenzelm@18583: in wenzelm@18583: (P, tuple_fun (map #2 avoiding) (Term.dest_Var P)) :: wenzelm@18583: (x, tuple (map Free avoiding)) :: wenzelm@32952: map_filter (fn (z, SOME t) => SOME (z, t) | _ => NONE) (zs ~~ inst) wenzelm@18583: end; wenzelm@18583: val substs = wenzelm@32952: map2 subst insts concls |> flat |> distinct (op =) wenzelm@19903: |> map (pairself (Thm.cterm_of (ProofContext.theory_of ctxt))); urbanc@22072: in urbanc@22072: (((cases, nconcls), consumes), Drule.cterm_instantiate substs joined_rule) urbanc@22072: end; wenzelm@18283: wenzelm@18299: fun rename_params_rule internal xs rule = wenzelm@18297: let wenzelm@18299: val tune = wenzelm@20072: if internal then Name.internal wenzelm@20072: else fn x => the_default x (try Name.dest_internal x); wenzelm@18299: val n = length xs; wenzelm@18299: fun rename prem = wenzelm@18299: let wenzelm@18299: val ps = Logic.strip_params prem; wenzelm@18299: val p = length ps; wenzelm@18299: val ys = wenzelm@18299: if p < n then [] haftmann@33955: else map (tune o #1) ((uncurry take) (p - n, ps)) @ xs; wenzelm@18299: in Logic.list_rename_params (ys, prem) end; wenzelm@18299: fun rename_prems prop = wenzelm@18299: let val (As, C) = Logic.strip_horn (Thm.prop_of rule) wenzelm@18299: in Logic.list_implies (map rename As, C) end; wenzelm@18299: in Thm.equal_elim (Thm.reflexive (Drule.cterm_fun rename_prems (Thm.cprop_of rule))) rule end; wenzelm@18297: wenzelm@18283: wenzelm@18288: (* nominal_induct_tac *) wenzelm@18283: wenzelm@18583: fun nominal_induct_tac ctxt def_insts avoiding fixings rules facts = wenzelm@18283: let wenzelm@18283: val thy = ProofContext.theory_of ctxt; wenzelm@18283: val cert = Thm.cterm_of thy; wenzelm@18283: wenzelm@24830: val ((insts, defs), defs_ctxt) = fold_map Induct.add_defs def_insts ctxt |>> split_list; wenzelm@23591: val atomized_defs = map (map (Conv.fconv_rule ObjectLogic.atomize)) defs; wenzelm@18283: wenzelm@19115: val finish_rule = wenzelm@18297: split_all_tuples wenzelm@26712: #> rename_params_rule true wenzelm@26712: (map (Name.clean o ProofContext.revert_skolem defs_ctxt o fst) avoiding); wenzelm@33368: fun rule_cases r = Rule_Cases.make_nested true (Thm.prop_of r) (Induct.rulified_term r); wenzelm@18283: in wenzelm@18297: (fn i => fn st => wenzelm@18583: rules wenzelm@19903: |> inst_mutual_rule ctxt insts avoiding wenzelm@33368: |> Rule_Cases.consume (flat defs) facts wenzelm@18583: |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) => wenzelm@18583: (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j => wenzelm@18583: (CONJUNCTS (ALLGOALS wenzelm@18583: (Method.insert_tac (more_facts @ nth_list atomized_defs (j - 1)) wenzelm@24830: THEN' Induct.fix_tac defs_ctxt wenzelm@18583: (nth concls (j - 1) + more_consumes) wenzelm@18583: (nth_list fixings (j - 1)))) wenzelm@24830: THEN' Induct.inner_atomize_tac) j)) wenzelm@24830: THEN' Induct.atomize_tac) i st |> Seq.maps (fn st' => wenzelm@26940: Induct.guess_instance ctxt wenzelm@24830: (finish_rule (Induct.internalize more_consumes rule)) i st' wenzelm@18583: |> Seq.maps (fn rule' => wenzelm@19115: CASES (rule_cases rule' cases) wenzelm@18583: (Tactic.rtac (rename_params_rule false [] rule') i THEN wenzelm@20288: PRIMITIVE (singleton (ProofContext.export defs_ctxt ctxt))) st')))) wenzelm@24830: THEN_ALL_NEW_CASES Induct.rulify_tac wenzelm@18283: end; wenzelm@18283: wenzelm@18283: wenzelm@18288: (* concrete syntax *) berghofe@17870: berghofe@17870: local berghofe@17870: wenzelm@27809: structure P = OuterParse; wenzelm@27809: wenzelm@18583: val avoidingN = "avoiding"; urbanc@20998: val fixingN = "arbitrary"; (* to be consistent with induct; hopefully this changes again *) wenzelm@18283: val ruleN = "rule"; berghofe@17870: berghofe@19036: val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.term >> SOME; berghofe@17870: wenzelm@18283: val def_inst = wenzelm@28083: ((Scan.lift (Args.binding --| (Args.$$$ "\" || Args.$$$ "==")) >> SOME) berghofe@19036: -- Args.term) >> SOME || wenzelm@18283: inst >> Option.map (pair NONE); urbanc@18099: wenzelm@27370: val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) => wenzelm@27370: error ("Bad free variable: " ^ Syntax.string_of_term ctxt t)); wenzelm@18283: wenzelm@18283: fun unless_more_args scan = Scan.unless (Scan.lift wenzelm@18583: ((Args.$$$ avoidingN || Args.$$$ fixingN || Args.$$$ ruleN) -- Args.colon)) scan; wenzelm@18283: berghofe@17870: wenzelm@18583: val avoiding = Scan.optional (Scan.lift (Args.$$$ avoidingN -- Args.colon) |-- wenzelm@18297: Scan.repeat (unless_more_args free)) []; berghofe@17870: wenzelm@18283: val fixing = Scan.optional (Scan.lift (Args.$$$ fixingN -- Args.colon) |-- wenzelm@27809: P.and_list' (Scan.repeat (unless_more_args free))) []; berghofe@17870: berghofe@19036: val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.colon) |-- Attrib.thms; berghofe@17870: berghofe@17870: in berghofe@17870: wenzelm@30549: val nominal_induct_method = wenzelm@30549: P.and_list' (Scan.repeat (unless_more_args def_inst)) -- wenzelm@30549: avoiding -- fixing -- rule_spec >> wenzelm@30549: (fn (((x, y), z), w) => fn ctxt => wenzelm@30510: RAW_METHOD_CASES (fn facts => wenzelm@18283: HEADGOAL (nominal_induct_tac ctxt x y z w facts))); berghofe@17870: berghofe@17870: end; wenzelm@18283: wenzelm@18283: end;