Nominal/activities/tphols09/IDW/InductiveInternals.thy
changeset 415 f1be8028a4a9
equal deleted inserted replaced
414:05e5d68c9627 415:f1be8028a4a9
       
     1 theory InductiveInternals
       
     2 imports Main
       
     3 begin
       
     4 
       
     5 section {* Parsing *}
       
     6 
       
     7 ML {*
       
     8 fun parse p s =
       
     9   Scan.finite OuterLex.stopper (Scan.error (OuterParse.!!! p))
       
    10     (filter OuterLex.is_proper (OuterSyntax.scan Position.none s));
       
    11 *}
       
    12 
       
    13 ML {*
       
    14 local structure P = OuterParse and K = OuterKeyword in
       
    15 
       
    16 val ind_parser =
       
    17   P.fixes -- P.for_fixes --
       
    18   Scan.optional (P.$$$ "where" |--
       
    19     P.!!! (P.enum1 "|" (SpecParse.opt_thm_name ":" -- P.prop))) [];
       
    20 
       
    21 end;
       
    22 *}
       
    23 
       
    24 ML {*
       
    25 parse ind_parser
       
    26 "even and odd \
       
    27 \where\
       
    28 \  even0: \"even 0\"\
       
    29 \| evenS: \"odd n \<Longrightarrow> even (Suc n)\"\
       
    30 \| oddS: \"even n \<Longrightarrow> odd (Suc n)\""
       
    31 *}
       
    32 
       
    33 ML {*
       
    34 val (((preds_syn, params_syn), intro_srcs), []) =
       
    35 parse ind_parser
       
    36 "accpart for r :: \"'a \<Rightarrow> 'a \<Rightarrow> bool\"\
       
    37 \where\
       
    38 \  accpartI: \"(\<And>y. r y x \<Longrightarrow> accpart r y) \<Longrightarrow> accpart r x\""
       
    39 *}
       
    40 
       
    41 
       
    42 section {* Reading the Specification *}
       
    43 
       
    44 ML {*
       
    45 val lthy = TheoryTarget.context "-" @{theory}
       
    46 *}
       
    47 
       
    48 ML {*
       
    49 val ((vars, intrs), _) = Specification.read_spec
       
    50   (preds_syn @ params_syn) intro_srcs lthy;
       
    51 val (preds_syn, params_syn') = chop (length preds_syn) vars
       
    52 val params = map fst params_syn';
       
    53 *}
       
    54 
       
    55 
       
    56 section {* Making the definition *}
       
    57 
       
    58 ML {*
       
    59 val params' = map (fn (p, T) => Free (Binding.name_of p, T)) params;
       
    60 val preds = map (fn ((R, T), _) =>
       
    61   list_comb (Free (Binding.name_of R, T), params')) preds_syn;
       
    62 val Tss = map (binder_types o fastype_of) preds;
       
    63 
       
    64 val intrs' = map
       
    65   (ObjectLogic.atomize_term (ProofContext.theory_of lthy) o snd) intrs;
       
    66 *}
       
    67 
       
    68 ML {*
       
    69 fun mk_all x P = HOLogic.all_const (fastype_of x) $ lambda x P;
       
    70 *}
       
    71 
       
    72 ML {*
       
    73 val (defs, lthy1) = fold_map (fn ((((R, _), syn), pred), Ts) =>
       
    74   let val zs = map Free (Variable.variant_frees lthy intrs'
       
    75     (map (pair "z") Ts))
       
    76   in
       
    77     LocalTheory.define Thm.internalK
       
    78       ((R, syn), (Attrib.empty_binding, fold_rev lambda (params' @ zs)
       
    79         (fold_rev mk_all preds (fold_rev (curry HOLogic.mk_imp)
       
    80            intrs' (list_comb (pred, zs)))))) #>> snd #>> snd
       
    81    end) (preds_syn ~~ preds ~~ Tss) lthy;
       
    82 *}
       
    83 
       
    84 ML {*
       
    85 val (_, lthy2) = Variable.add_fixes
       
    86   (map (Binding.name_of o fst) params) lthy1;
       
    87 *}
       
    88 
       
    89 
       
    90 section {* Proving the induction rules *}
       
    91 
       
    92 ML {*
       
    93 val (Pnames, lthy3) =
       
    94   Variable.variant_fixes (replicate (length preds) "P") lthy2;
       
    95 val Ps = map (fn (s, Ts) => Free (s, Ts ---> HOLogic.boolT))
       
    96   (Pnames ~~ Tss);
       
    97 val cPs = map (cterm_of (ProofContext.theory_of lthy3)) Ps;
       
    98 val intrs'' = map (subst_free (preds ~~ Ps) o snd) intrs;
       
    99 *}
       
   100 
       
   101 ML {*
       
   102 fun inst_spec ct = Drule.instantiate'
       
   103   [SOME (ctyp_of_term ct)] [NONE, SOME ct] spec;
       
   104 *}
       
   105 
       
   106 ML {*
       
   107 fun prove_indrule ((R, P), Ts) =
       
   108   let
       
   109     val (znames, lthy4) =
       
   110       Variable.variant_fixes (replicate (length Ts) "z") lthy3;
       
   111     val zs = map Free (znames ~~ Ts)
       
   112   in
       
   113     Goal.prove lthy4 []
       
   114       [HOLogic.mk_Trueprop (list_comb (R, zs))]
       
   115       (Logic.list_implies (intrs'',
       
   116          HOLogic.mk_Trueprop (list_comb (P, zs))))
       
   117       (fn {prems, ...} => EVERY
       
   118          ([ObjectLogic.full_atomize_tac 1,
       
   119            cut_facts_tac prems 1,
       
   120            rewrite_goals_tac defs] @
       
   121           map (fn ct => dtac (inst_spec ct) 1) cPs @
       
   122           [assume_tac 1])) |>
       
   123     singleton (ProofContext.export lthy4 lthy1)
       
   124   end;
       
   125 *}
       
   126 
       
   127 ML {*
       
   128 val indrules = map prove_indrule (preds ~~ Ps ~~ Tss);
       
   129 *}
       
   130 
       
   131 
       
   132 section {* Proving the introduction rules *}
       
   133 
       
   134 ML {*
       
   135 val all_elims = fold (fn ct => fn th => th RS inst_spec ct);
       
   136 val imp_elims = fold (fn th => fn th' => [th', th] MRS mp);
       
   137 *}
       
   138 
       
   139 ML {*
       
   140 fun show_thms ctxt s ths = warning (s ^ ":\n" ^
       
   141   Pretty.string_of (Pretty.chunks (map (Display.pretty_thm ctxt) ths)));
       
   142 
       
   143 fun show_cterms ctxt s cts = warning (s ^ ":\n" ^
       
   144   Pretty.string_of (Pretty.list "[" "]"
       
   145     (map (Syntax.pretty_term ctxt o term_of) cts)));
       
   146 *}
       
   147 
       
   148 ML {*
       
   149 fun prove_intr (i, (_, r)) =
       
   150   Goal.prove lthy2 [] [] r
       
   151     (fn {prems, context = ctxt} => EVERY
       
   152        [ObjectLogic.rulify_tac 1,
       
   153         rewrite_goals_tac defs,
       
   154         REPEAT (resolve_tac [allI, impI] 1),
       
   155         SUBPROOF (fn {params, prems, context = ctxt', ...} =>
       
   156           let
       
   157             val (prems1, prems2) =
       
   158               chop (length prems - length intrs) prems;
       
   159             val _ = show_thms ctxt' "prems1" prems1;
       
   160             val _ = show_thms ctxt' "prems2" prems2;
       
   161             val (params1, params2) =
       
   162               chop (length params - length preds) (map snd params)
       
   163             val _ = show_cterms ctxt' "params1" params1;
       
   164             val _ = show_cterms ctxt' "params2" params2;
       
   165           in
       
   166             rtac (ObjectLogic.rulify
       
   167               (all_elims params1 (nth prems2 i))) 1 THEN
       
   168             EVERY (map (fn prem =>
       
   169               SUBPROOF (fn {prems = prems', concl, context = ctxt'', ...} =>
       
   170                 let
       
   171                   val prem' = prems' MRS prem;
       
   172                   val prem'' = case prop_of prem' of
       
   173                       _ $ (Const (@{const_name All}, _) $ _) =>
       
   174                         prem' |> all_elims params2 |>
       
   175                         imp_elims prems2
       
   176                     | _ => prem'
       
   177                   val _ = show_thms ctxt'' "prem" [prem];
       
   178                   val _ = show_thms ctxt'' "prem'" [prem'];
       
   179                   val _ = show_thms ctxt'' "prem''" [prem''];
       
   180                 in rtac prem'' 1 end) ctxt' 1) prems1)
       
   181           end) ctxt 1]) |>
       
   182   singleton (ProofContext.export lthy2 lthy1);
       
   183 *}
       
   184 
       
   185 ML {*
       
   186 val intr_ths = map_index prove_intr intrs;
       
   187 *}
       
   188 
       
   189 end