src/HOL/Tools/BNF/bnf_lift.ML
author wenzelm
Sun, 13 Dec 2015 21:56:15 +0100
changeset 61841 4d3527b94f2a
parent 61073 eea21f2ddf16
child 62137 b8dc1fd7d900
permissions -rw-r--r--
more general types Proof.method / context_tactic; proper context for Method.insert_tac; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
     1
(*  Title:      HOL/Tools/BNF/bnf_lift.ML
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
     2
    Author:     Julian Biendarra, TU Muenchen
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
     3
    Author:     Dmitriy Traytel, ETH Zurich
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
     4
    Copyright   2015
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
     5
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
     6
Lifting of BNFs through typedefs.
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
     7
*)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
     8
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
     9
signature BNF_LIFT =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    10
sig
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    11
  datatype lift_bnf_option =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    12
    Plugins_Option of Proof.context -> Plugin_Name.filter
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    13
  | No_Warn_Wits
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    14
  val copy_bnf:
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    15
    (((lift_bnf_option list * (binding option * (string * sort option)) list) *
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    16
      string) * thm option) * (binding * binding) ->
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    17
      local_theory -> local_theory
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    18
  val copy_bnf_cmd:
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    19
    (((lift_bnf_option list * (binding option * (string * string option)) list) *
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    20
      string) * (Facts.ref * Token.src list) option) * (binding * binding) ->
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    21
      local_theory -> local_theory
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    22
  val lift_bnf:
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    23
    (((lift_bnf_option list * (binding option * (string * sort option)) list) *
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    24
      string) * thm option) * (binding * binding) ->
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    25
      ({context: Proof.context, prems: thm list} -> tactic) list ->
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    26
      local_theory -> local_theory
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    27
  val lift_bnf_cmd:
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    28
     ((((lift_bnf_option list * (binding option * (string * string option)) list) *
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    29
       string) * string list) * (Facts.ref * Token.src list) option) * (binding * binding) ->
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    30
       local_theory -> Proof.state
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    31
end
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    32
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    33
structure BNF_Lift : BNF_LIFT =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    34
struct
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    35
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    36
open Ctr_Sugar_Tactics
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    37
open BNF_Util
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    38
open BNF_Comp
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    39
open BNF_Def
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    40
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    41
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    42
(* typedef_bnf *)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    43
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    44
datatype lift_bnf_option =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    45
  Plugins_Option of Proof.context -> Plugin_Name.filter
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    46
| No_Warn_Wits;
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    47
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    48
fun typedef_bnf thm wits specs map_b rel_b opts lthy =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    49
  let
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
    50
    val plugins =
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
    51
      get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts)
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    52
      |> the_default Plugin_Name.default_filter;
61072
f9be82413170 proper option, not catch-all pattern;
wenzelm
parents: 61067
diff changeset
    53
    val no_warn_wits = exists (fn No_Warn_Wits => true | _ => false) opts;
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    54
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    55
    (* extract Rep Abs F RepT AbsT *)
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
    56
    val (_, [Rep_G, Abs_G, F]) = Term.strip_comb (HOLogic.dest_Trueprop (Thm.prop_of thm));
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
    57
    val typ_Abs_G = dest_funT (fastype_of Abs_G);
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    58
    val RepT = fst typ_Abs_G; (* F *)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    59
    val AbsT = snd typ_Abs_G; (* G *)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    60
    val AbsT_name = fst (dest_Type AbsT);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    61
    val tvs = AbsT |> dest_Type |> snd |> map (fst o dest_TVar);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    62
    val alpha0s = map (TFree o snd) specs;
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    63
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    64
    (* instantiate the new type variables newtvs to oldtvs *)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    65
    val subst = subst_TVars (tvs ~~ alpha0s);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    66
    val typ_subst = typ_subst_TVars (tvs ~~ alpha0s);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    67
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    68
    val Rep_G = subst Rep_G;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    69
    val Abs_G = subst Abs_G;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    70
    val F = subst F;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    71
    val RepT = typ_subst RepT;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    72
    val AbsT = typ_subst AbsT;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    73
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
    74
    fun flatten_tyargs Ass =
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
    75
      map dest_TFree alpha0s
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
    76
      |> filter (fn T => exists (fn Ts => member (op =) Ts T) Ass);
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    77
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    78
    val Ds0 = filter (is_none o fst) specs |> map snd;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    79
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    80
    (* get the bnf for RepT *)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    81
    val ((bnf, (deads, alphas)),((_, unfolds), lthy)) =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    82
      bnf_of_typ Dont_Inline (Binding.qualify true AbsT_name) flatten_tyargs []
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    83
        Ds0 RepT ((empty_comp_cache, empty_unfolds), lthy);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    84
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
    85
    val set_bs =
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
    86
      map (fn T => find_index (fn U => T = U) alpha0s) alphas
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    87
      |> map (the_default Binding.empty o fst o nth specs);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    88
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
    89
    val _ = (case alphas of [] => error "No live variables" | _ => ());
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    90
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    91
    val defs = #map_unfolds unfolds @ flat (#set_unfoldss unfolds) @ #rel_unfolds unfolds;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    92
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    93
    (* number of live variables *)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    94
    val lives = length alphas;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    95
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    96
    (* state the three required properties *)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    97
    val sorts = map Type.sort_of_atyp alphas;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    98
    val names_lthy = fold Variable.declare_typ (alphas @ deads) lthy;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    99
    val (alphas', names_lthy) = mk_TFrees' sorts names_lthy;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   100
    val (betas, names_lthy) = mk_TFrees' sorts names_lthy;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   101
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   102
    val map_F = mk_map_of_bnf deads alphas betas bnf;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   103
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   104
    val (typ_fs, typ_aF) = fastype_of map_F |> strip_typeN lives ||> domain_type;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   105
    val typ_pairs = map HOLogic.mk_prodT (alphas ~~ alphas');
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   106
    val typ_subst_pair = typ_subst_atomic (alphas ~~ typ_pairs);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   107
    val typ_pair = typ_subst_pair RepT;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   108
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   109
    val subst_b = subst_atomic_types (alphas ~~ betas);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   110
    val subst_a' = subst_atomic_types (alphas ~~ alphas');
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   111
    val subst_pair = subst_atomic_types (alphas ~~ typ_pairs);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   112
    val aF_set = F;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   113
    val bF_set = subst_b F;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   114
    val aF_set' = subst_a' F;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   115
    val pairF_set = subst_pair F;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   116
    val map_F_fst = mk_map_of_bnf deads typ_pairs alphas bnf;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   117
    val map_F_snd = mk_map_of_bnf deads typ_pairs alphas' bnf;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   118
    val wits_F = mk_wits_of_bnf
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   119
      (replicate (nwits_of_bnf bnf) deads)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   120
      (replicate (nwits_of_bnf bnf) alphas) bnf;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   121
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   122
    (* val map_closed_F = @{term "\<And>f x. x \<in> F \<Longrightarrow> map_F f x \<in> F"}; *)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   123
    val (var_fs, names_lthy) = mk_Frees "f" typ_fs names_lthy;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   124
    val (var_x, names_lthy) = mk_Frees "x" [typ_aF] names_lthy |>> the_single;
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   125
    val mem_x = HOLogic.mk_Trueprop (HOLogic.mk_mem (var_x, aF_set));
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   126
    val map_f = list_comb (map_F, var_fs);
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   127
    val mem_map = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_f $ var_x, bF_set));
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   128
    val imp_map = Logic.mk_implies (mem_x, mem_map);
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   129
    val map_closed_F = fold_rev Logic.all var_fs (Logic.all var_x imp_map);
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   130
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   131
    (* val zip_closed_F = @{term "\<And>z. map_F fst z \<in> F \<Longrightarrow> map_F snd z \<in> F \<Longrightarrow> z \<in> F"}; *)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   132
    val (var_zs, names_lthy) = mk_Frees "z" [typ_pair] names_lthy;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   133
    val (pairs, names_lthy) = mk_Frees "tmp" typ_pairs names_lthy;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   134
    val var_z = hd var_zs;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   135
    val fsts = map (fst o Term.strip_comb o HOLogic.mk_fst) pairs;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   136
    val snds = map (fst o Term.strip_comb o HOLogic.mk_snd) pairs;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   137
    val map_fst = list_comb (list_comb (map_F_fst, fsts), var_zs);
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   138
    val mem_map_fst = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_fst, aF_set));
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   139
    val map_snd = list_comb (list_comb (map_F_snd, snds), var_zs);
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   140
    val mem_map_snd = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_snd, aF_set'));
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   141
    val mem_z = HOLogic.mk_Trueprop (HOLogic.mk_mem (var_z, pairF_set));
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   142
    val imp_zip = Logic.mk_implies (mem_map_fst, Logic.mk_implies (mem_map_snd, mem_z));
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   143
    val zip_closed_F = Logic.all var_z imp_zip;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   144
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   145
    (* val wit_closed_F = @{term "wit_F a \<in> F"}; *)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   146
    val (var_as, names_lthy) = mk_Frees "a" alphas names_lthy;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   147
    val (var_bs, _) = mk_Frees "a" alphas names_lthy;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   148
    val Iwits = the_default wits_F (Option.map (map (`(map (fn T =>
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   149
      find_index (fn U => T = U) alphas) o fst o strip_type o fastype_of))) wits);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   150
    val wit_closed_Fs =
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   151
      Iwits |> map (fn (I, wit_F) =>
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   152
        let
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   153
          val vars = map (nth var_as) I;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   154
          val wit_a = list_comb (wit_F, vars);
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   155
        in fold_rev Logic.all vars (HOLogic.mk_Trueprop (HOLogic.mk_mem (wit_a, aF_set))) end);
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   156
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   157
    val mk_wit_goals = mk_wit_goals var_as var_bs
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   158
      (mk_sets_of_bnf (replicate lives deads)  (replicate lives alphas) bnf);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   159
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   160
    val goals = [map_closed_F, zip_closed_F] @ wit_closed_Fs @
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   161
      (case wits of NONE => [] | _ => maps mk_wit_goals Iwits);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   162
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   163
    val lost_wits = filter_out (fn (J, _) => exists (fn (I, _) => I = J) Iwits) wits_F;
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   164
    val _ =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   165
      if null lost_wits orelse no_warn_wits then ()
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   166
      else
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   167
        lost_wits
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   168
        |> map (Syntax.pretty_typ lthy o fastype_of o snd)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   169
        |> Pretty.big_list
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   170
          "The following types of nonemptiness witnesses of the raw type's BNF were lost:"
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   171
        |> (fn pt => Pretty.chunks [pt,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   172
          Pretty.para "You can specify a liftable witness (e.g., a term of one of the above types\
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   173
            \ that satisfies the typedef's invariant)\
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   174
            \ using the annotation [wits: <term>]."])
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   175
        |> Pretty.string_of
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   176
        |> warning;
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   177
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   178
    fun after_qed ([map_closed_thm] :: [zip_closed_thm] :: wit_thmss) lthy =
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   179
          let
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   180
            val (wit_closed_thms, wit_thms) =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   181
              (case wits of
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   182
                NONE => (map the_single wit_thmss, wit_thms_of_bnf bnf)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   183
              | _ => chop (length wit_closed_Fs) (map the_single wit_thmss))
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   184
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   185
            (*  construct map set bd rel wit *)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   186
            (* val map_G = @{term "\<lambda>f. Abs_G o map_F f o Rep_G"}; *)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   187
            val Abs_Gb = subst_b Abs_G;
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   188
            val map_G =
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   189
              fold_rev HOLogic.tupled_lambda var_fs
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   190
                (HOLogic.mk_comp (HOLogic.mk_comp (Abs_Gb, map_f), Rep_G));
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   191
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   192
            (* val sets_G = [@{term "set_F o Rep_G"}]; *)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   193
            val sets_F = mk_sets_of_bnf (replicate lives deads) (replicate lives alphas) bnf;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   194
            val sets_G = map (fn set_F => HOLogic.mk_comp (set_F, Rep_G)) sets_F;
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   195
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   196
            (* val bd_G = @{term "bd_F"}; *)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   197
            val bd_F = mk_bd_of_bnf deads alphas bnf;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   198
            val bd_G = bd_F;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   199
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   200
            (* val rel_G = @{term "\<lambda>R. BNF_Def.vimage2p Rep_G Rep_G (rel_F R)"}; *)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   201
            val rel_F = mk_rel_of_bnf deads alphas betas bnf;
61073
eea21f2ddf16 misc tuning and simplification;
wenzelm
parents: 61072
diff changeset
   202
            val (typ_Rs, _) = strip_typeN lives (fastype_of rel_F);
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   203
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   204
            val (var_Rs, names_lthy) = mk_Frees "R" typ_Rs lthy;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   205
            val Rep_Gb = subst_b Rep_G;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   206
            val rel_G = fold_rev absfree (map dest_Free var_Rs)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   207
              (mk_vimage2p Rep_G Rep_Gb $ list_comb (rel_F, var_Rs));
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   208
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   209
            (* val wits_G = [@{term "Abs_G o wit_F"}]; *)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   210
            val (var_as, _) = mk_Frees "a" alphas names_lthy;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   211
            val wits_G =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   212
              map (fn (I, wit_F) =>
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   213
                let
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   214
                  val vs = map (nth var_as) I;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   215
                in fold_rev absfree (map dest_Free vs) (Abs_G $ (list_comb (wit_F, vs))) end)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   216
              Iwits;
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   217
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   218
            (* tactics *)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   219
            val Rep_thm = thm RS @{thm type_definition.Rep};
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   220
            val Abs_inverse_thm = thm RS @{thm type_definition.Abs_inverse};
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   221
            val Abs_inject_thm = thm RS @{thm type_definition.Abs_inject};
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   222
            val Rep_cases_thm = thm RS @{thm type_definition.Rep_cases};
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   223
            val Rep_inverse_thm = thm RS @{thm type_definition.Rep_inverse};
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   224
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   225
            fun map_id0_tac ctxt =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   226
              HEADGOAL (EVERY' [rtac ctxt ext,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   227
                SELECT_GOAL (unfold_thms_tac ctxt [map_id0_of_bnf bnf, id_apply, o_apply,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   228
                  Rep_inverse_thm]),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   229
                rtac ctxt refl]);
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   230
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   231
            fun map_comp0_tac ctxt =
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   232
              HEADGOAL (EVERY' [rtac ctxt ext,
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   233
                SELECT_GOAL (unfold_thms_tac ctxt [map_comp0_of_bnf bnf, o_apply,
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   234
                  Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]),
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   235
                rtac ctxt refl]);
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   236
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   237
            fun map_cong0_tac ctxt =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   238
              HEADGOAL (EVERY' ([SELECT_GOAL (unfold_thms_tac ctxt [o_apply]),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   239
                rtac ctxt (([Rep_thm RS map_closed_thm, Rep_thm RS map_closed_thm] MRS
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   240
                  Abs_inject_thm) RS iffD2),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   241
                rtac ctxt (map_cong0_of_bnf bnf)] @ replicate lives (Goal.assume_rule_tac ctxt)));
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   242
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   243
            val set_map0s_tac =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   244
              map (fn set_map => fn ctxt =>
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   245
                HEADGOAL (EVERY' [rtac ctxt ext,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   246
                  SELECT_GOAL (unfold_thms_tac ctxt [set_map, o_apply,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   247
                    Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   248
                  rtac ctxt refl]))
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   249
             (set_map_of_bnf bnf);
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   250
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   251
            fun card_order_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_card_order_of_bnf bnf));
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   252
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   253
            fun cinfinite_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_cinfinite_of_bnf bnf));
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   254
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   255
            val set_bds_tac =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   256
              map (fn set_bd => fn ctxt =>
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   257
                HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), rtac ctxt set_bd]))
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   258
              (set_bd_of_bnf bnf);
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   259
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   260
            fun le_rel_OO_tac ctxt =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   261
              HEADGOAL (EVERY' [rtac ctxt @{thm vimage2p_relcompp_mono},
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   262
                rtac ctxt ((rel_OO_of_bnf bnf RS sym) RS @{thm ord_eq_le_trans}),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   263
                rtac ctxt @{thm order_refl}]);
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   264
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   265
            fun rel_OO_Grp_tac ctxt =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   266
              HEADGOAL (EVERY' ([SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt ext))),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   267
                SELECT_GOAL (unfold_thms_tac ctxt [@{thm OO_Grp_alt}, mem_Collect_eq,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   268
                  o_apply, @{thm vimage2p_def}, in_rel_of_bnf bnf, Bex_def, mem_Collect_eq]),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   269
                rtac ctxt iffI,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   270
                SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve0_tac [exE,conjE]))),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   271
                rtac ctxt (zip_closed_thm OF (replicate 2 (Rep_thm RSN (2, @{thm ssubst_mem}))) RS
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   272
                  Rep_cases_thm),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   273
                assume_tac ctxt,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   274
                assume_tac ctxt,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   275
                hyp_subst_tac ctxt,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   276
                SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt exI))),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   277
                rtac ctxt conjI] @
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   278
                replicate (lives - 1) (rtac ctxt conjI THEN' assume_tac ctxt) @
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   279
                [assume_tac ctxt,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   280
                SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt conjI))),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   281
                REPEAT_DETERM_N 2 o
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   282
                  etac ctxt (trans OF [iffD2 OF [Abs_inject_thm OF
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   283
                    [map_closed_thm OF [Rep_thm], Rep_thm]], Rep_inverse_thm]),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   284
                SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve0_tac [exE,conjE]))),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   285
                rtac ctxt exI,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   286
                rtac ctxt conjI] @
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   287
                replicate (lives - 1) (rtac ctxt conjI THEN' assume_tac ctxt) @
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   288
                [assume_tac ctxt,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   289
                rtac ctxt conjI,
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   290
                REPEAT_DETERM_N 2 o EVERY'
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   291
                  [rtac ctxt (iffD1 OF [Abs_inject_thm OF [map_closed_thm OF [Rep_thm], Rep_thm]]),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   292
                  etac ctxt (Rep_inverse_thm RS sym RSN (2, trans))]]));
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   293
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   294
            fun wit_tac ctxt =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   295
              HEADGOAL (EVERY'
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   296
                (map (fn thm => (EVERY'
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   297
                  [SELECT_GOAL (unfold_thms_tac ctxt (o_apply ::
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   298
                    (wit_closed_thms RL [Abs_inverse_thm]))),
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   299
                  dtac ctxt thm, assume_tac ctxt]))
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   300
                wit_thms));
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   301
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   302
            val tactics = [map_id0_tac, map_comp0_tac, map_cong0_tac] @ set_map0s_tac @
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   303
              [card_order_bd_tac, cinfinite_bd_tac] @ set_bds_tac @ [le_rel_OO_tac, rel_OO_Grp_tac];
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   304
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   305
            val (bnf, lthy) = bnf_def Dont_Inline (user_policy Note_Some) false I
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   306
              tactics wit_tac NONE map_b rel_b set_bs
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   307
              ((((((Binding.empty, AbsT), map_G), sets_G), bd_G), wits_G), SOME rel_G)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   308
              lthy;
60928
141a1d485259 unfold intermediate definitions (stemming from composition) in lifted bnf operations
traytel
parents: 60918
diff changeset
   309
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   310
            val bnf = morph_bnf_defs (Morphism.thm_morphism "BNF" (unfold_thms lthy defs)) bnf;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   311
          in
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   312
            lthy |> BNF_Def.register_bnf plugins AbsT_name bnf
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   313
          end
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   314
      | after_qed _ _ = raise Match;
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   315
  in
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   316
    (goals, after_qed, defs, lthy)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   317
  end;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   318
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   319
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   320
(* main commands *)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   321
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   322
local
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   323
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   324
fun prepare_common prepare_name prepare_sort prepare_term prepare_thm
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   325
    (((((plugins, raw_specs), raw_Tname), raw_wits), xthm_opt), (map_b, rel_b)) lthy =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   326
  let
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   327
    val Tname = prepare_name lthy raw_Tname;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   328
    val input_thm =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   329
      (case xthm_opt of
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   330
        SOME xthm => prepare_thm lthy xthm
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   331
      | NONE => Typedef.get_info lthy Tname |> hd |> snd |> #type_definition);
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   332
    val wits = (Option.map o map) (prepare_term lthy) raw_wits;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   333
    val specs =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   334
      map (apsnd (apsnd (the_default @{sort type} o Option.map (prepare_sort lthy)))) raw_specs;
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   335
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   336
    val _ =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   337
      (case HOLogic.dest_Trueprop (Thm.prop_of input_thm) of
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   338
        Const (@{const_name type_definition}, _) $ _ $ _ $ _ => ()
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   339
      | _ => error "Unsupported type of a theorem: only type_definition is supported");
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   340
  in
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   341
    typedef_bnf input_thm wits specs map_b rel_b plugins lthy
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   342
  end;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   343
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   344
fun prepare_lift_bnf prepare_name prepare_sort prepare_term prepare_thm =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   345
  (fn (goals, after_qed, definitions, lthy) =>
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   346
    lthy
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   347
    |> Proof.theorem NONE after_qed (map (single o rpair []) goals)
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61073
diff changeset
   348
    |> Proof.refine_singleton
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61073
diff changeset
   349
        (Method.Basic (fn ctxt => SIMPLE_METHOD (unfold_thms_tac ctxt definitions)))
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61073
diff changeset
   350
    |> Proof.refine_singleton (Method.primitive_text (K I))) oo
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   351
  prepare_common prepare_name prepare_sort prepare_term prepare_thm o apfst (apfst (apsnd SOME));
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   352
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   353
fun prepare_solve prepare_name prepare_typ prepare_sort prepare_thm tacs =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   354
  (fn (goals, after_qed, _, lthy) =>
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   355
    lthy
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   356
    |> after_qed (map2 (single oo Goal.prove lthy [] []) goals (tacs (length goals)))) oo
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   357
  prepare_common prepare_name prepare_typ prepare_sort prepare_thm o apfst (apfst (rpair NONE));
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   358
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   359
in
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   360
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   361
val lift_bnf_cmd =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   362
  prepare_lift_bnf
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   363
    (fst o dest_Type oo Proof_Context.read_type_name {proper = true, strict = false})
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   364
    Syntax.read_sort Syntax.read_term (singleton o Attrib.eval_thms);
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   365
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   366
fun lift_bnf args tacs = prepare_solve (K I) (K I) (K I) (K I) (K tacs) args;
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   367
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   368
val copy_bnf =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   369
  prepare_solve (K I) (K I) (K I) (K I)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   370
    (fn n => replicate n (fn {context = ctxt, prems = _} => rtac ctxt UNIV_I 1));
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   371
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   372
val copy_bnf_cmd =
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   373
  prepare_solve
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   374
    (fst o dest_Type oo Proof_Context.read_type_name {proper = true, strict = false})
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   375
    Syntax.read_sort Syntax.read_term (singleton o Attrib.eval_thms)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   376
    (fn n => replicate n (fn {context = ctxt, prems = _} => rtac ctxt UNIV_I 1));
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   377
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   378
end;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   379
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   380
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   381
(* outer syntax *)
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   382
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   383
local
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   384
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   385
val parse_wits =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   386
  @{keyword "["} |-- (Parse.name --| @{keyword ":"} -- Scan.repeat Parse.term >>
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   387
    (fn ("wits", Ts) => Ts
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   388
      | (s, _) => error ("Unknown label " ^ quote s ^ " (expected \"wits\")"))) --|
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   389
  @{keyword "]"} || Scan.succeed [];
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   390
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   391
val parse_options =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   392
  Scan.optional (@{keyword "("} |--
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   393
    Parse.list1 (Parse.group (K "option")
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   394
      (Plugin_Name.parse_filter >> Plugins_Option
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   395
      || Parse.reserved "no_warn_wits" >> K No_Warn_Wits))
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   396
    --| @{keyword ")"}) [];
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   397
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   398
val parse_plugins =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   399
  Scan.optional (@{keyword "("} |-- Plugin_Name.parse_filter --| @{keyword ")"})
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   400
    (K Plugin_Name.default_filter) >> Plugins_Option >> single;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   401
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   402
val parse_typedef_thm = Scan.option (Parse.reserved "via" |-- Parse.xthm);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   403
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   404
in
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   405
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   406
val _ =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   407
  Outer_Syntax.local_theory_to_proof @{command_keyword lift_bnf}
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   408
    "register a subtype of a bounded natural functor (BNF) as a BNF"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   409
    ((parse_options -- parse_type_args_named_constrained -- Parse.type_const -- parse_wits --
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   410
      parse_typedef_thm -- parse_map_rel_bindings) >> lift_bnf_cmd);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   411
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   412
val _ =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   413
  Outer_Syntax.local_theory @{command_keyword copy_bnf}
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   414
    "register a type copy of a bounded natural functor (BNF) as a BNF"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   415
    ((parse_plugins -- parse_type_args_named_constrained -- Parse.type_const --
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   416
      parse_typedef_thm -- parse_map_rel_bindings) >> copy_bnf_cmd);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   417
61067
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   418
end;
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   419
180a20d4ae53 misc tuning and clarification;
wenzelm
parents: 60928
diff changeset
   420
end;