src/HOL/Nominal/nominal_induct.ML
changeset 18583 96e1ef2f806f
parent 18311 b83b00cbaecf
child 18610 05a5e950d5f1
equal deleted inserted replaced
18582:4f4cc426b440 18583:96e1ef2f806f
     4 The nominal induct proof method.
     4 The nominal induct proof method.
     5 *)
     5 *)
     6 
     6 
     7 structure NominalInduct:
     7 structure NominalInduct:
     8 sig
     8 sig
     9   val nominal_induct_tac: Proof.context -> (string option * term) option list ->
     9   val nominal_induct_tac: Proof.context -> (string option * term) option list list ->
    10     (string * typ) list -> (string * typ) list list -> thm ->
    10     (string * typ) list -> (string * typ) list list -> thm list ->
    11     thm list -> int -> RuleCases.cases_tactic
    11     thm list -> int -> RuleCases.cases_tactic
    12   val nominal_induct_method: Method.src -> Proof.context -> Method.method
    12   val nominal_induct_method: Method.src -> Proof.context -> Method.method
    13 end =
    13 end =
    14 struct
    14 struct
    15 
       
    16 
    15 
    17 (* proper tuples -- nested left *)
    16 (* proper tuples -- nested left *)
    18 
    17 
    19 fun tupleT Ts = HOLogic.unitT |> fold (fn T => fn U => HOLogic.mk_prodT (U, T)) Ts;
    18 fun tupleT Ts = HOLogic.unitT |> fold (fn T => fn U => HOLogic.mk_prodT (U, T)) Ts;
    20 fun tuple ts = HOLogic.unit |> fold (fn t => fn u => HOLogic.mk_prod (u, t)) ts;
    19 fun tuple ts = HOLogic.unit |> fold (fn t => fn u => HOLogic.mk_prod (u, t)) ts;
    28     [split_conv, split_paired_all, unit_all_eq1, thm "fresh_unit_elim", thm "fresh_prod_elim"]);
    27     [split_conv, split_paired_all, unit_all_eq1, thm "fresh_unit_elim", thm "fresh_prod_elim"]);
    29 
    28 
    30 
    29 
    31 (* prepare rule *)
    30 (* prepare rule *)
    32 
    31 
    33 (*conclusion: ?P fresh_struct ... insts*)
    32 (*conclusions: ?P avoiding_struct ... insts*)
    34 fun inst_rule thy insts fresh rule =
    33 fun inst_mutual_rule thy insts avoiding rules =
    35   let
    34   let
    36     val vars = InductAttrib.vars_of (Thm.concl_of rule);
    35     val (concls, rule) =
    37     val m = length vars and n = length insts;
    36       (case RuleCases.mutual_rule rules of
    38     val _ = if m >= n + 2 then () else error "Too few variables in conclusion of rule";
    37         NONE => error "Failed to join given rules into one mutual rule"
    39     val P :: x :: ys = vars;
    38       | SOME res => res);
    40     val zs = Library.drop (m - n - 2, ys);
    39     val (cases, consumes) = RuleCases.get rule;
    41 
    40 
    42     val subst =
    41     val l = length rules;
    43       (P, tuple_fun (map #2 fresh) (Term.dest_Var P)) ::
    42     val _ =
    44       (x, tuple (map Free fresh)) ::
    43       if length insts = l then ()
    45       List.mapPartial (fn (z, SOME t) => SOME (z, t) | _ => NONE) (zs ~~ insts);
    44       else error ("Bad number of instantiations for " ^ string_of_int l ^ " rules");
    46   in
    45 
    47     rule
    46     fun subst inst rule =
    48     |> Drule.cterm_instantiate (map (pairself (Thm.cterm_of thy)) subst)
    47       let
    49   end;
    48         val vars = InductAttrib.vars_of (Thm.concl_of rule);
       
    49         val m = length vars and n = length inst;
       
    50         val _ = if m >= n + 2 then () else error "Too few variables in conclusion of rule";
       
    51         val P :: x :: ys = vars;
       
    52         val zs = Library.drop (m - n - 2, ys);
       
    53       in
       
    54         (P, tuple_fun (map #2 avoiding) (Term.dest_Var P)) ::
       
    55         (x, tuple (map Free avoiding)) ::
       
    56         List.mapPartial (fn (z, SOME t) => SOME (z, t) | _ => NONE) (zs ~~ inst)
       
    57       end;
       
    58      val substs =
       
    59        map2 subst insts rules |> List.concat |> distinct
       
    60        |> map (pairself (Thm.cterm_of thy));
       
    61   in (((cases, concls), consumes), Drule.cterm_instantiate substs rule) end;
    50 
    62 
    51 fun rename_params_rule internal xs rule =
    63 fun rename_params_rule internal xs rule =
    52   let
    64   let
    53     val tune =
    65     val tune =
    54       if internal then Syntax.internal
    66       if internal then Syntax.internal
    68   in Thm.equal_elim (Thm.reflexive (Drule.cterm_fun rename_prems (Thm.cprop_of rule))) rule end;
    80   in Thm.equal_elim (Thm.reflexive (Drule.cterm_fun rename_prems (Thm.cprop_of rule))) rule end;
    69 
    81 
    70 
    82 
    71 (* nominal_induct_tac *)
    83 (* nominal_induct_tac *)
    72 
    84 
    73 fun nominal_induct_tac ctxt def_insts fresh fixing rule facts =
    85 fun nominal_induct_tac ctxt def_insts avoiding fixings rules facts =
    74   let
    86   let
    75     val thy = ProofContext.theory_of ctxt;
    87     val thy = ProofContext.theory_of ctxt;
    76     val cert = Thm.cterm_of thy;
    88     val cert = Thm.cterm_of thy;
    77 
    89 
    78     val ((insts, defs), defs_ctxt) = InductMethod.add_defs def_insts ctxt;
    90     val ((insts, defs), defs_ctxt) = fold_map InductMethod.add_defs def_insts ctxt |>> split_list;
    79     val atomized_defs = map ObjectLogic.atomize_thm defs;
    91     val atomized_defs = map (map ObjectLogic.atomize_thm) defs;
    80 
    92 
    81     val finish_rule =
    93     val finish_rule = PolyML.print #>
    82       split_all_tuples
    94       split_all_tuples
    83       #> rename_params_rule true (map (ProofContext.revert_skolem defs_ctxt o #1) fresh);
    95       #> rename_params_rule true (map (ProofContext.revert_skolem defs_ctxt o fst) avoiding) #> PolyML.print;
    84     fun rule_cases r = RuleCases.make true (SOME (Thm.prop_of r)) (InductMethod.rulified_term r);
    96     fun rule_cases r = RuleCases.make true (SOME (Thm.prop_of r)) (InductMethod.rulified_term r);
    85   in
    97   in
    86     (fn i => fn st =>
    98     (fn i => fn st =>
    87       rule
    99       rules
    88       |> `RuleCases.get
   100       |> inst_mutual_rule thy insts avoiding
    89       ||> inst_rule thy insts fresh
   101       |> RuleCases.consume (List.concat defs) facts
    90       |> RuleCases.consume defs facts
   102       |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) =>
    91       |> Seq.maps (fn ((cases, (k, more_facts)), r) =>
   103         (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j =>
    92         (CONJUNCTS (ALLGOALS (fn j =>
   104           (CONJUNCTS (ALLGOALS
    93             Method.insert_tac (more_facts @ atomized_defs) j
   105             (Method.insert_tac (more_facts @ nth_list atomized_defs (j - 1))
    94             THEN InductMethod.fix_tac defs_ctxt k (Library.nth_list fixing (j - 1)) j))
   106               THEN' InductMethod.fix_tac defs_ctxt
    95           THEN' InductMethod.atomize_tac) i st |> Seq.maps (fn st' =>
   107                 (nth concls (j - 1) + more_consumes)
    96             InductMethod.guess_instance (finish_rule r) i st'
   108                 (nth_list fixings (j - 1))))
    97             |> Seq.maps (fn r' =>
   109           THEN' InductMethod.inner_atomize_tac) j))
    98               CASES (rule_cases r' cases)
   110         THEN' InductMethod.atomize_tac) i st |> Seq.maps (fn st' =>
    99                 (Tactic.rtac (rename_params_rule false [] r') i THEN
   111             InductMethod.guess_instance (finish_rule (InductMethod.internalize more_consumes rule)) i (PolyML.print st')
       
   112             |> Seq.maps (fn rule' =>
       
   113               CASES (rule_cases (PolyML.print rule') cases)
       
   114                 (Tactic.rtac (rename_params_rule false [] rule') i THEN
   100                   PRIMSEQ (ProofContext.exports defs_ctxt ctxt)) st'))))
   115                   PRIMSEQ (ProofContext.exports defs_ctxt ctxt)) st'))))
   101     THEN_ALL_NEW_CASES InductMethod.rulify_tac
   116     THEN_ALL_NEW_CASES InductMethod.rulify_tac
   102   end;
   117   end;
   103 
   118 
   104 
   119 
   105 (* concrete syntax *)
   120 (* concrete syntax *)
   106 
   121 
   107 local
   122 local
   108 
   123 
   109 val freshN = "avoiding";
   124 val avoidingN = "avoiding";
   110 val fixingN = "fixing";
   125 val fixingN = "fixing";
   111 val ruleN = "rule";
   126 val ruleN = "rule";
   112 
   127 
   113 val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.local_term >> SOME;
   128 val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.local_term >> SOME;
   114 
   129 
   119 
   134 
   120 val free = Scan.state -- Args.local_term >> (fn (_, Free v) => v | (ctxt, t) =>
   135 val free = Scan.state -- Args.local_term >> (fn (_, Free v) => v | (ctxt, t) =>
   121   error ("Bad free variable: " ^ ProofContext.string_of_term ctxt t));
   136   error ("Bad free variable: " ^ ProofContext.string_of_term ctxt t));
   122 
   137 
   123 fun unless_more_args scan = Scan.unless (Scan.lift
   138 fun unless_more_args scan = Scan.unless (Scan.lift
   124   ((Args.$$$ freshN || Args.$$$ fixingN || Args.$$$ ruleN) -- Args.colon)) scan;
   139   ((Args.$$$ avoidingN || Args.$$$ fixingN || Args.$$$ ruleN) -- Args.colon)) scan;
   125 
   140 
   126 
   141 
   127 val def_insts = Scan.repeat (unless_more_args def_inst);
   142 val avoiding = Scan.optional (Scan.lift (Args.$$$ avoidingN -- Args.colon) |--
   128 
       
   129 val fresh = Scan.optional (Scan.lift (Args.$$$ freshN -- Args.colon) |--
       
   130   Scan.repeat (unless_more_args free)) [];
   143   Scan.repeat (unless_more_args free)) [];
   131 
   144 
   132 val fixing = Scan.optional (Scan.lift (Args.$$$ fixingN -- Args.colon) |--
   145 val fixing = Scan.optional (Scan.lift (Args.$$$ fixingN -- Args.colon) |--
   133   Args.and_list1 (Scan.repeat (unless_more_args free))) [];
   146   Args.and_list (Scan.repeat (unless_more_args free))) [];
   134 
   147 
   135 val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.colon) |-- Attrib.local_thm;
   148 val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.colon) |-- Attrib.local_thmss;
   136 
   149 
   137 in
   150 in
   138 
   151 
   139 fun nominal_induct_method src =
   152 fun nominal_induct_method src =
   140   Method.syntax (def_insts -- fresh -- fixing -- rule_spec) src
   153   Method.syntax
       
   154    (Args.and_list (Scan.repeat (unless_more_args def_inst)) --
       
   155     avoiding -- fixing -- rule_spec) src
   141   #> (fn (ctxt, (((x, y), z), w)) =>
   156   #> (fn (ctxt, (((x, y), z), w)) =>
   142     Method.RAW_METHOD_CASES (fn facts =>
   157     Method.RAW_METHOD_CASES (fn facts =>
   143       HEADGOAL (nominal_induct_tac ctxt x y z w facts)));
   158       HEADGOAL (nominal_induct_tac ctxt x y z w facts)));
   144 
   159 
   145 end;
   160 end;