src/HOL/Tools/datatype_abs_proofs.ML
author wenzelm
Tue, 02 Sep 2008 14:10:45 +0200
changeset 28083 103d9282a946
parent 27691 ce171cbd4b93
child 28084 a05ca48ef263
permissions -rw-r--r--
explicit type Name.binding for higher-specification elements;
     1 (*  Title:      HOL/Tools/datatype_abs_proofs.ML
     2     ID:         $Id$
     3     Author:     Stefan Berghofer, TU Muenchen
     4 
     5 Proofs and defintions independent of concrete representation
     6 of datatypes  (i.e. requiring only abstract properties such as
     7 injectivity / distinctness of constructors and induction)
     8 
     9  - case distinction (exhaustion) theorems
    10  - characteristic equations for primrec combinators
    11  - characteristic equations for case combinators
    12  - equations for splitting "P (case ...)" expressions
    13   - "nchotomy" and "case_cong" theorems for TFL
    14 
    15 *)
    16 
    17 signature DATATYPE_ABS_PROOFS =
    18 sig
    19   val prove_casedist_thms : string list ->
    20     DatatypeAux.descr list -> (string * sort) list -> thm ->
    21     attribute list -> theory -> thm list * theory
    22   val prove_primrec_thms : bool -> string list ->
    23     DatatypeAux.descr list -> (string * sort) list ->
    24       DatatypeAux.datatype_info Symtab.table -> thm list list -> thm list list ->
    25         simpset -> thm -> theory -> (string list * thm list) * theory
    26   val prove_case_thms : bool -> string list ->
    27     DatatypeAux.descr list -> (string * sort) list ->
    28       string list -> thm list -> theory -> (thm list list * string list) * theory
    29   val prove_split_thms : string list ->
    30     DatatypeAux.descr list -> (string * sort) list ->
    31       thm list list -> thm list list -> thm list -> thm list list -> theory ->
    32         (thm * thm) list * theory
    33   val prove_nchotomys : string list -> DatatypeAux.descr list ->
    34     (string * sort) list -> thm list -> theory -> thm list * theory
    35   val prove_weak_case_congs : string list -> DatatypeAux.descr list ->
    36     (string * sort) list -> theory -> thm list * theory
    37   val prove_case_congs : string list ->
    38     DatatypeAux.descr list -> (string * sort) list ->
    39       thm list -> thm list list -> theory -> thm list * theory
    40 end;
    41 
    42 structure DatatypeAbsProofs: DATATYPE_ABS_PROOFS =
    43 struct
    44 
    45 open DatatypeAux;
    46 
    47 (************************ case distinction theorems ***************************)
    48 
    49 fun prove_casedist_thms new_type_names descr sorts induct case_names_exhausts thy =
    50   let
    51     val _ = message "Proving case distinction theorems ...";
    52 
    53     val descr' = List.concat descr;
    54     val recTs = get_rec_types descr' sorts;
    55     val newTs = Library.take (length (hd descr), recTs);
    56 
    57     val {maxidx, ...} = rep_thm induct;
    58     val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
    59 
    60     fun prove_casedist_thm ((i, t), T) =
    61       let
    62         val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
    63           Abs ("z", T', Const ("True", T''))) induct_Ps;
    64         val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $
    65           Var (("P", 0), HOLogic.boolT))
    66         val insts = Library.take (i, dummyPs) @ (P::(Library.drop (i + 1, dummyPs)));
    67         val cert = cterm_of thy;
    68         val insts' = (map cert induct_Ps) ~~ (map cert insts);
    69         val induct' = refl RS ((List.nth
    70           (split_conj_thm (cterm_instantiate insts' induct), i)) RSN (2, rev_mp))
    71 
    72       in
    73         SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
    74           (fn {prems, ...} => EVERY
    75             [rtac induct' 1,
    76              REPEAT (rtac TrueI 1),
    77              REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
    78              REPEAT (rtac TrueI 1)])
    79       end;
    80 
    81     val casedist_thms = map prove_casedist_thm ((0 upto (length newTs - 1)) ~~
    82       (DatatypeProp.make_casedists descr sorts) ~~ newTs)
    83   in
    84     thy
    85     |> store_thms_atts "exhaust" new_type_names (map single case_names_exhausts) casedist_thms
    86   end;
    87 
    88 
    89 (*************************** primrec combinators ******************************)
    90 
    91 fun prove_primrec_thms flat_names new_type_names descr sorts
    92     (dt_info : datatype_info Symtab.table) constr_inject dist_rewrites dist_ss induct thy =
    93   let
    94     val _ = message "Constructing primrec combinators ...";
    95 
    96     val big_name = space_implode "_" new_type_names;
    97     val thy0 = add_path flat_names big_name thy;
    98 
    99     val descr' = List.concat descr;
   100     val recTs = get_rec_types descr' sorts;
   101     val used = foldr add_typ_tfree_names [] recTs;
   102     val newTs = Library.take (length (hd descr), recTs);
   103 
   104     val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
   105 
   106     val big_rec_name' = big_name ^ "_rec_set";
   107     val rec_set_names' =
   108       if length descr' = 1 then [big_rec_name'] else
   109         map ((curry (op ^) (big_rec_name' ^ "_")) o string_of_int)
   110           (1 upto (length descr'));
   111     val rec_set_names = map (Sign.full_name thy0) rec_set_names';
   112 
   113     val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used;
   114 
   115     val rec_set_Ts = map (fn (T1, T2) =>
   116       reccomb_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
   117 
   118     val rec_fns = map (uncurry (mk_Free "f"))
   119       (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
   120     val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
   121       (rec_set_names' ~~ rec_set_Ts);
   122     val rec_sets = map (fn c => list_comb (Const c, rec_fns))
   123       (rec_set_names ~~ rec_set_Ts);
   124 
   125     (* introduction rules for graph of primrec function *)
   126 
   127     fun make_rec_intr T rec_set ((rec_intr_ts, l), (cname, cargs)) =
   128       let
   129         fun mk_prem ((dt, U), (j, k, prems, t1s, t2s)) =
   130           let val free1 = mk_Free "x" U j
   131           in (case (strip_dtyp dt, strip_type U) of
   132              ((_, DtRec m), (Us, _)) =>
   133                let
   134                  val free2 = mk_Free "y" (Us ---> List.nth (rec_result_Ts, m)) k;
   135                  val i = length Us
   136                in (j + 1, k + 1, HOLogic.mk_Trueprop (HOLogic.list_all
   137                      (map (pair "x") Us, List.nth (rec_sets', m) $
   138                        app_bnds free1 i $ app_bnds free2 i)) :: prems,
   139                    free1::t1s, free2::t2s)
   140                end
   141            | _ => (j + 1, k, prems, free1::t1s, t2s))
   142           end;
   143 
   144         val Ts = map (typ_of_dtyp descr' sorts) cargs;
   145         val (_, _, prems, t1s, t2s) = foldr mk_prem (1, 1, [], [], []) (cargs ~~ Ts)
   146 
   147       in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop
   148         (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
   149           list_comb (List.nth (rec_fns, l), t1s @ t2s)))], l + 1)
   150       end;
   151 
   152     val (rec_intr_ts, _) = Library.foldl (fn (x, ((d, T), set_name)) =>
   153       Library.foldl (make_rec_intr T set_name) (x, #3 (snd d)))
   154         (([], 0), descr' ~~ recTs ~~ rec_sets');
   155 
   156     val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
   157         InductivePackage.add_inductive_global (serial_string ())
   158           {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
   159             alt_name = Name.binding big_rec_name', coind = false, no_elim = false, no_ind = true,
   160             skip_mono = true}
   161           (map (fn (s, T) => ((Name.binding s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
   162           (map dest_Free rec_fns)
   163           (map (fn x => ((Name.no_binding, []), x)) rec_intr_ts) [] thy0;
   164 
   165     (* prove uniqueness and termination of primrec combinators *)
   166 
   167     val _ = message "Proving termination and uniqueness of primrec functions ...";
   168 
   169     fun mk_unique_tac ((tac, intrs), ((((i, (tname, _, constrs)), elim), T), T')) =
   170       let
   171         val distinct_tac =
   172           (if i < length newTs then
   173              full_simp_tac (HOL_ss addsimps (List.nth (dist_rewrites, i))) 1
   174            else full_simp_tac dist_ss 1);
   175 
   176         val inject = map (fn r => r RS iffD1)
   177           (if i < length newTs then List.nth (constr_inject, i)
   178             else #inject (the (Symtab.lookup dt_info tname)));
   179 
   180         fun mk_unique_constr_tac n ((tac, intr::intrs, j), (cname, cargs)) =
   181           let
   182             val k = length (List.filter is_rec_type cargs)
   183 
   184           in (EVERY [DETERM tac,
   185                 REPEAT (etac ex1E 1), rtac ex1I 1,
   186                 DEPTH_SOLVE_1 (ares_tac [intr] 1),
   187                 REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1),
   188                 etac elim 1,
   189                 REPEAT_DETERM_N j distinct_tac,
   190                 TRY (dresolve_tac inject 1),
   191                 REPEAT (etac conjE 1), hyp_subst_tac 1,
   192                 REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
   193                 TRY (hyp_subst_tac 1),
   194                 rtac refl 1,
   195                 REPEAT_DETERM_N (n - j - 1) distinct_tac],
   196               intrs, j + 1)
   197           end;
   198 
   199         val (tac', intrs', _) = Library.foldl (mk_unique_constr_tac (length constrs))
   200           ((tac, intrs, 0), constrs);
   201 
   202       in (tac', intrs') end;
   203 
   204     val rec_unique_thms =
   205       let
   206         val rec_unique_ts = map (fn (((set_t, T1), T2), i) =>
   207           Const ("Ex1", (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
   208             absfree ("y", T2, set_t $ mk_Free "x" T1 i $ Free ("y", T2)))
   209               (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
   210         val cert = cterm_of thy1
   211         val insts = map (fn ((i, T), t) => absfree ("x" ^ (string_of_int i), T, t))
   212           ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
   213         val induct' = cterm_instantiate ((map cert induct_Ps) ~~
   214           (map cert insts)) induct;
   215         val (tac, _) = Library.foldl mk_unique_tac
   216           (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1
   217               THEN rewtac (mk_meta_eq choice_eq), rec_intrs),
   218             descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts);
   219 
   220       in split_conj_thm (SkipProof.prove_global thy1 [] []
   221         (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
   222       end;
   223 
   224     val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
   225 
   226     (* define primrec combinators *)
   227 
   228     val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
   229     val reccomb_names = map (Sign.full_name thy1)
   230       (if length descr' = 1 then [big_reccomb_name] else
   231         (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
   232           (1 upto (length descr'))));
   233     val reccombs = map (fn ((name, T), T') => list_comb
   234       (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
   235         (reccomb_names ~~ recTs ~~ rec_result_Ts);
   236 
   237     val (reccomb_defs, thy2) =
   238       thy1
   239       |> Sign.add_consts_i (map (fn ((name, T), T') =>
   240           (Sign.base_name name, reccomb_fn_Ts @ [T] ---> T', NoSyn))
   241           (reccomb_names ~~ recTs ~~ rec_result_Ts))
   242       |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
   243           ((Sign.base_name name) ^ "_def", Logic.mk_equals (comb, absfree ("x", T,
   244            Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
   245              set $ Free ("x", T) $ Free ("y", T'))))))
   246                (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
   247       ||> parent_path flat_names;
   248 
   249 
   250     (* prove characteristic equations for primrec combinators *)
   251 
   252     val _ = message "Proving characteristic theorems for primrec combinators ..."
   253 
   254     val rec_thms = map (fn t => SkipProof.prove_global thy2 [] [] t
   255       (fn _ => EVERY
   256         [rewrite_goals_tac reccomb_defs,
   257          rtac the1_equality 1,
   258          resolve_tac rec_unique_thms 1,
   259          resolve_tac rec_intrs 1,
   260          REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
   261            (DatatypeProp.make_primrecs new_type_names descr sorts thy2)
   262 
   263   in
   264     thy2
   265     |> Sign.add_path (space_implode "_" new_type_names)
   266     |> PureThy.add_thmss [(("recs", rec_thms), [])]
   267     ||> Sign.parent_path
   268     |-> (fn thms => pair (reccomb_names, Library.flat thms))
   269   end;
   270 
   271 
   272 (***************************** case combinators *******************************)
   273 
   274 fun prove_case_thms flat_names new_type_names descr sorts reccomb_names primrec_thms thy =
   275   let
   276     val _ = message "Proving characteristic theorems for case combinators ...";
   277 
   278     val thy1 = add_path flat_names (space_implode "_" new_type_names) thy;
   279 
   280     val descr' = List.concat descr;
   281     val recTs = get_rec_types descr' sorts;
   282     val used = foldr add_typ_tfree_names [] recTs;
   283     val newTs = Library.take (length (hd descr), recTs);
   284     val T' = TFree (Name.variant used "'t", HOLogic.typeS);
   285 
   286     fun mk_dummyT dt = binder_types (typ_of_dtyp descr' sorts dt) ---> T';
   287 
   288     val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
   289       let
   290         val Ts = map (typ_of_dtyp descr' sorts) cargs;
   291         val Ts' = map mk_dummyT (List.filter is_rec_type cargs)
   292       in Const ("arbitrary", Ts @ Ts' ---> T')
   293       end) constrs) descr';
   294 
   295     val case_names = map (fn s => Sign.full_name thy1 (s ^ "_case")) new_type_names;
   296 
   297     (* define case combinators via primrec combinators *)
   298 
   299     val (case_defs, thy2) = Library.foldl (fn ((defs, thy),
   300       ((((i, (_, _, constrs)), T), name), recname)) =>
   301         let
   302           val (fns1, fns2) = ListPair.unzip (map (fn ((_, cargs), j) =>
   303             let
   304               val Ts = map (typ_of_dtyp descr' sorts) cargs;
   305               val Ts' = Ts @ map mk_dummyT (List.filter is_rec_type cargs);
   306               val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts'));
   307               val frees = Library.take (length cargs, frees');
   308               val free = mk_Free "f" (Ts ---> T') j
   309             in
   310              (free, list_abs_free (map dest_Free frees',
   311                list_comb (free, frees)))
   312             end) (constrs ~~ (1 upto length constrs)));
   313 
   314           val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T';
   315           val fns = (List.concat (Library.take (i, case_dummy_fns))) @
   316             fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns)));
   317           val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
   318           val decl = (Sign.base_name name, caseT, NoSyn);
   319           val def = ((Sign.base_name name) ^ "_def",
   320             Logic.mk_equals (list_comb (Const (name, caseT), fns1),
   321               list_comb (reccomb, (List.concat (Library.take (i, case_dummy_fns))) @
   322                 fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns))) )));
   323           val ([def_thm], thy') =
   324             thy
   325             |> Sign.declare_const [] decl |> snd
   326             |> (PureThy.add_defs false o map Thm.no_attributes) [def];
   327 
   328         in (defs @ [def_thm], thy')
   329         end) (([], thy1), (hd descr) ~~ newTs ~~ case_names ~~
   330           (Library.take (length newTs, reccomb_names)));
   331 
   332     val case_thms = map (map (fn t => SkipProof.prove_global thy2 [] [] t
   333       (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
   334           (DatatypeProp.make_cases new_type_names descr sorts thy2)
   335 
   336   in
   337     thy2
   338     |> parent_path flat_names
   339     |> store_thmss "cases" new_type_names case_thms
   340     |-> (fn thmss => pair (thmss, case_names))
   341   end;
   342 
   343 
   344 (******************************* case splitting *******************************)
   345 
   346 fun prove_split_thms new_type_names descr sorts constr_inject dist_rewrites
   347     casedist_thms case_thms thy =
   348   let
   349     val _ = message "Proving equations for case splitting ...";
   350 
   351     val descr' = List.concat descr;
   352     val recTs = get_rec_types descr' sorts;
   353     val newTs = Library.take (length (hd descr), recTs);
   354 
   355     fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'),
   356         exhaustion), case_thms'), T) =
   357       let
   358         val cert = cterm_of thy;
   359         val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
   360         val exhaustion' = cterm_instantiate
   361           [(cert lhs, cert (Free ("x", T)))] exhaustion;
   362         val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
   363           (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
   364       in
   365         (SkipProof.prove_global thy [] [] t1 tacf,
   366          SkipProof.prove_global thy [] [] t2 tacf)
   367       end;
   368 
   369     val split_thm_pairs = map prove_split_thms
   370       ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
   371         dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
   372 
   373     val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs
   374 
   375   in
   376     thy
   377     |> store_thms "split" new_type_names split_thms
   378     ||>> store_thms "split_asm" new_type_names split_asm_thms
   379     |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
   380   end;
   381 
   382 fun prove_weak_case_congs new_type_names descr sorts thy =
   383   let
   384     fun prove_weak_case_cong t =
   385        SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   386          (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1])
   387 
   388     val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs
   389       new_type_names descr sorts thy)
   390 
   391   in thy |> store_thms "weak_case_cong" new_type_names weak_case_congs end;
   392 
   393 (************************* additional theorems for TFL ************************)
   394 
   395 fun prove_nchotomys new_type_names descr sorts casedist_thms thy =
   396   let
   397     val _ = message "Proving additional theorems for TFL ...";
   398 
   399     fun prove_nchotomy (t, exhaustion) =
   400       let
   401         (* For goal i, select the correct disjunct to attack, then prove it *)
   402         fun tac i 0 = EVERY [TRY (rtac disjI1 i),
   403               hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
   404           | tac i n = rtac disjI2 i THEN tac i (n - 1)
   405       in 
   406         SkipProof.prove_global thy [] [] t (fn _ =>
   407           EVERY [rtac allI 1,
   408            exh_tac (K exhaustion) 1,
   409            ALLGOALS (fn i => tac i (i-1))])
   410       end;
   411 
   412     val nchotomys =
   413       map prove_nchotomy (DatatypeProp.make_nchotomys descr sorts ~~ casedist_thms)
   414 
   415   in thy |> store_thms "nchotomy" new_type_names nchotomys end;
   416 
   417 fun prove_case_congs new_type_names descr sorts nchotomys case_thms thy =
   418   let
   419     fun prove_case_cong ((t, nchotomy), case_rewrites) =
   420       let
   421         val (Const ("==>", _) $ tm $ _) = t;
   422         val (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ Ma)) = tm;
   423         val cert = cterm_of thy;
   424         val nchotomy' = nchotomy RS spec;
   425         val nchotomy'' = cterm_instantiate
   426           [(cert (hd (add_term_vars (concl_of nchotomy', []))), cert Ma)] nchotomy'
   427       in
   428         SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   429           (fn {prems, ...} => 
   430             let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
   431             in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
   432                 cut_facts_tac [nchotomy''] 1,
   433                 REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
   434                 REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
   435             end)
   436       end;
   437 
   438     val case_congs = map prove_case_cong (DatatypeProp.make_case_congs
   439       new_type_names descr sorts thy ~~ nchotomys ~~ case_thms)
   440 
   441   in thy |> store_thms "case_cong" new_type_names case_congs end;
   442 
   443 end;