src/HOL/Tools/Lifting/lifting_def_code_dt.ML
author kuncar
Mon, 13 Apr 2015 15:27:34 +0200
changeset 60230 4857d553c52c
parent 60229 4cd6462c1fda
child 60231 0daab758e087
permissions -rw-r--r--
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
60229
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Lifting/lifting_def_code_dt.ML
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
     2
    Author:     Ondrej Kuncar
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
     3
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
     4
Workaround that allows us to execute lifted constants that have
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
     5
as a return type a datatype containing a subtype; lift_definition command
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
     6
*)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
     7
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
     8
signature LIFTING_DEF_CODE_DT =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
     9
sig
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    10
  type rep_isom_data
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    11
  val isom_of_rep_isom_data: rep_isom_data -> term
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    12
  val transfer_of_rep_isom_data: rep_isom_data -> thm
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    13
  val bundle_name_of_rep_isom_data: rep_isom_data -> string
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    14
  val pointer_of_rep_isom_data: rep_isom_data -> string
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    15
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    16
  type code_dt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    17
  val rty_of_code_dt: code_dt -> typ
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    18
  val qty_of_code_dt: code_dt -> typ
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    19
  val wit_of_code_dt: code_dt -> term
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    20
  val wit_thm_of_code_dt: code_dt -> thm
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    21
  val rep_isom_data_of_code_dt: code_dt -> rep_isom_data option
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    22
  val morph_code_dt: morphism -> code_dt -> code_dt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    23
  val mk_witness_of_code_dt: typ -> code_dt -> term
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    24
  val mk_rep_isom_of_code_dt: typ -> code_dt -> term option
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    25
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    26
  val code_dt_of: Proof.context -> typ * typ -> code_dt option
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    27
  val code_dt_of_global: theory -> typ * typ -> code_dt option
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    28
  val all_code_dt_of: Proof.context -> code_dt list
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    29
  val all_code_dt_of_global: theory -> code_dt list
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    30
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    31
  type config_code_dt = { code_dt: bool, lift_config: Lifting_Def.config }
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    32
  val default_config_code_dt: config_code_dt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    33
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    34
  val add_lift_def_code_dt:
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    35
    config_code_dt -> binding * mixfix -> typ -> term -> thm -> thm list -> local_theory ->
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    36
      Lifting_Def.lift_def * local_theory
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    37
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    38
  val lift_def_code_dt:
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    39
    config_code_dt -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list ->
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    40
    local_theory -> Lifting_Def.lift_def * local_theory
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    41
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    42
  val lift_def_cmd:
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    43
    string list * (binding * string option * mixfix) * string * (Facts.ref * Token.src list) list ->
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    44
    local_theory -> Proof.state
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    45
end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    46
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    47
structure Lifting_Def_Code_Dt: LIFTING_DEF_CODE_DT =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    48
struct
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    49
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    50
open Ctr_Sugar_Util BNF_Util BNF_FP_Util BNF_FP_Def_Sugar Lifting_Def Lifting_Util
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    51
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    52
(** data structures **)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    53
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    54
(* all type variables in qty are in rty *)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    55
datatype rep_isom_data = REP_ISOM of { isom: term, transfer: thm, bundle_name: string, pointer: string }
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    56
fun isom_of_rep_isom_data (REP_ISOM rep_isom) = #isom rep_isom;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    57
fun transfer_of_rep_isom_data (REP_ISOM rep_isom) = #transfer rep_isom;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    58
fun bundle_name_of_rep_isom_data (REP_ISOM rep_isom) = #bundle_name rep_isom;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    59
fun pointer_of_rep_isom_data (REP_ISOM rep_isom) = #pointer rep_isom;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    60
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    61
datatype code_dt = CODE_DT of { rty: typ, qty: typ, wit: term, wit_thm: thm,
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    62
  rep_isom_data: rep_isom_data option };
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    63
fun rty_of_code_dt (CODE_DT code_dt) = #rty code_dt;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    64
fun qty_of_code_dt (CODE_DT code_dt) = #qty code_dt;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    65
fun wit_of_code_dt (CODE_DT code_dt) = #wit code_dt;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    66
fun wit_thm_of_code_dt (CODE_DT code_dt) = #wit_thm code_dt;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    67
fun rep_isom_data_of_code_dt (CODE_DT code_dt) = #rep_isom_data code_dt;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    68
fun rty_equiv (T, U) = Type.raw_instance (T, U) andalso Type.raw_instance (U, T);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    69
val code_dt_eq = rty_equiv o apply2 rty_of_code_dt;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    70
fun term_of_code_dt code_dt = code_dt |> `rty_of_code_dt ||> qty_of_code_dt |> HOLogic.mk_prodT
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    71
  |> Net.encode_type |> single;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    72
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    73
(* modulo renaming, typ must contain TVars *)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    74
fun is_code_dt_of_type (rty, qty) code_dt = code_dt |> `rty_of_code_dt ||> qty_of_code_dt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    75
  |> HOLogic.mk_prodT |> curry rty_equiv (HOLogic.mk_prodT (rty, qty));
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    76
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    77
fun mk_rep_isom_data isom transfer bundle_name pointer =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    78
  REP_ISOM { isom = isom, transfer = transfer, bundle_name = bundle_name, pointer = pointer}
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    79
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    80
fun mk_code_dt rty qty wit wit_thm rep_isom_data =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    81
  CODE_DT { rty = rty, qty = qty, wit = wit, wit_thm = wit_thm, rep_isom_data = rep_isom_data };
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    82
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    83
fun map_rep_isom_data f1 f2 f3 f4
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    84
  (REP_ISOM { isom = isom, transfer = transfer, bundle_name = bundle_name, pointer = pointer }) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    85
  REP_ISOM { isom = f1 isom, transfer = f2 transfer, bundle_name = f3 bundle_name, pointer = f4 pointer };
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    86
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    87
fun map_code_dt f1 f2 f3 f4 f5 f6 f7 f8
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    88
  (CODE_DT {rty = rty, qty = qty, wit = wit, wit_thm = wit_thm, rep_isom_data = rep_isom_data}) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    89
  CODE_DT {rty = f1 rty, qty = f2 qty, wit = f3 wit, wit_thm = f4 wit_thm,
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    90
    rep_isom_data = Option.map (map_rep_isom_data f5 f6 f7 f8) rep_isom_data};
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    91
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    92
fun update_rep_isom isom transfer binding pointer i = mk_code_dt (rty_of_code_dt i) (qty_of_code_dt i)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    93
  (wit_of_code_dt i) (wit_thm_of_code_dt i) (SOME (mk_rep_isom_data isom transfer binding pointer))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    94
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    95
fun morph_code_dt phi =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    96
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    97
    val mty = Morphism.typ phi
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    98
    val mterm = Morphism.term phi
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
    99
    val mthm = Morphism.thm phi
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   100
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   101
    map_code_dt mty mty mterm mthm mterm mthm I I
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   102
  end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   103
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   104
val transfer_code_dt = morph_code_dt o Morphism.transfer_morphism;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   105
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   106
structure Data = Generic_Data
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   107
(
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   108
  type T = code_dt Item_Net.T
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   109
  val empty = Item_Net.init code_dt_eq term_of_code_dt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   110
  val extend = I
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   111
  val merge = Item_Net.merge
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   112
);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   113
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   114
fun code_dt_of_generic context (rty, qty) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   115
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   116
    val typ = HOLogic.mk_prodT (rty, qty)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   117
    val prefiltred = Item_Net.retrieve_matching (Data.get context) (Net.encode_type typ)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   118
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   119
    prefiltred |> filter (is_code_dt_of_type (rty, qty))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   120
    |> map (transfer_code_dt (Context.theory_of context)) |> find_first (fn _ => true)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   121
  end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   122
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   123
fun code_dt_of ctxt (rty, qty) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   124
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   125
    val sch_rty = Logic.type_map (singleton (Variable.polymorphic ctxt)) rty
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   126
    val sch_qty = Logic.type_map (singleton (Variable.polymorphic ctxt)) qty
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   127
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   128
    code_dt_of_generic (Context.Proof ctxt) (sch_rty, sch_qty)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   129
  end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   130
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   131
fun code_dt_of_global thy (rty, qty) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   132
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   133
    val sch_rty = Logic.varifyT_global rty
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   134
    val sch_qty = Logic.varifyT_global qty
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   135
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   136
    code_dt_of_generic (Context.Theory thy) (sch_rty, sch_qty)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   137
  end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   138
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   139
fun all_code_dt_of_generic context =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   140
    Item_Net.content (Data.get context) |> map (transfer_code_dt (Context.theory_of context));
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   141
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   142
val all_code_dt_of = all_code_dt_of_generic o Context.Proof;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   143
val all_code_dt_of_global = all_code_dt_of_generic o Context.Theory;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   144
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   145
fun update_code_dt code_dt =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   146
  Local_Theory.declaration {syntax = false, pervasive = true}
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   147
    (fn phi => Data.map (Item_Net.update (morph_code_dt phi code_dt)));
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   148
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   149
fun mk_match_of_code_dt qty code_dt = Vartab.empty |> Type.raw_match (qty_of_code_dt code_dt, qty)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   150
  |> Vartab.dest |> map (fn (x, (S, T)) => (TVar (x, S), T));
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   151
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   152
fun mk_witness_of_code_dt qty code_dt =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   153
  Term.subst_atomic_types (mk_match_of_code_dt qty code_dt) (wit_of_code_dt code_dt)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   154
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   155
fun mk_rep_isom_of_code_dt qty code_dt = Option.map
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   156
  (isom_of_rep_isom_data #> Term.subst_atomic_types (mk_match_of_code_dt qty code_dt))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   157
    (rep_isom_data_of_code_dt code_dt)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   158
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   159
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   160
(** unique name for a type **)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   161
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   162
fun var_name name sort = if sort = @{sort "{type}"} orelse sort = [] then ["x" ^ name]
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   163
  else "x" ^ name :: "x_" :: sort @ ["x_"];
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   164
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   165
fun concat_Tnames (Type (name, ts)) = name :: maps concat_Tnames ts
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   166
  | concat_Tnames (TFree (name, sort)) = var_name name sort
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   167
  | concat_Tnames (TVar ((name, _), sort)) = var_name name sort;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   168
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   169
fun unique_Tname (rty, qty) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   170
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   171
    val Tnames = map Long_Name.base_name (concat_Tnames rty @ ["x_x"] @ concat_Tnames qty);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   172
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   173
    fold (Binding.qualify false) (tl Tnames) (Binding.name (hd Tnames))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   174
  end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   175
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   176
(** witnesses **)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   177
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   178
fun mk_undefined T = Const (@{const_name undefined}, T);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   179
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   180
fun mk_witness quot_thm =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   181
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   182
    val wit_thm = quot_thm RS @{thm type_definition_Quotient_not_empty_witness}
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   183
    val wit = quot_thm_rep quot_thm $ mk_undefined (quot_thm_rty_qty quot_thm |> snd)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   184
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   185
    (wit, wit_thm)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   186
  end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   187
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   188
(** config **)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   189
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   190
type config_code_dt = { code_dt: bool, lift_config: config }
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   191
val default_config_code_dt = { code_dt = false, lift_config = default_config }
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   192
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   193
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   194
(** Main code **)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   195
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   196
val ld_no_notes = { notes = false }
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   197
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   198
fun comp_lift_error _ _ = error "Composition of abstract types has not been implemented yet."
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   199
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   200
fun lift qty (quot_thm, (lthy, rel_eq_onps)) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   201
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   202
    val quot_thm = Lifting_Term.force_qty_type lthy qty quot_thm
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   203
    val (rty, qty) = quot_thm_rty_qty quot_thm;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   204
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   205
    if is_none (code_dt_of lthy (rty, qty)) then
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   206
      let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   207
        val (wit, wit_thm) = (mk_witness quot_thm
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   208
          handle THM _ => error ("code_dt: " ^ quote (Tname qty) ^ " was not defined as a subtype."))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   209
        val code_dt = mk_code_dt rty qty wit wit_thm NONE
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   210
      in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   211
        (quot_thm, (update_code_dt code_dt lthy |> Local_Theory.restore, rel_eq_onps))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   212
      end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   213
    else
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   214
      (quot_thm, (lthy, rel_eq_onps))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   215
  end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   216
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   217
fun case_tac rule ctxt i st =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   218
  (Subgoal.FOCUS_PARAMS (fn {params, ...} => HEADGOAL(rtac
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   219
    (Ctr_Sugar_Util.cterm_instantiate_pos [SOME (params |> hd |> snd)] rule))) ctxt i st);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   220
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   221
fun bundle_name_of_bundle_binding binding phi context  =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   222
  Name_Space.full_name (Name_Space.naming_of context) (Morphism.binding phi binding);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   223
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   224
fun prove_schematic_quot_thm actions ctxt = Lifting_Term.prove_schematic_quot_thm actions
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   225
 (Lifting_Info.get_quotients ctxt) ctxt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   226
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   227
fun prove_code_dt (rty, qty) lthy =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   228
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   229
    val (fold_quot_thm: (local_theory * thm list) Lifting_Term.fold_quot_thm) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   230
      { constr = constr, lift = lift, comp_lift = comp_lift_error };
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   231
  in prove_schematic_quot_thm fold_quot_thm lthy (rty, qty) (lthy, []) |> snd end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   232
and add_lift_def_code_dt config var qty rhs rsp_thm par_thms lthy =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   233
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   234
    fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   235
    fun ret_rel_conv conv ctm =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   236
      case (Thm.term_of ctm) of
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   237
        Const (@{const_name "rel_fun"}, _) $ _ $ _ =>
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   238
          binop_conv2 Conv.all_conv conv ctm
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   239
        | _ => conv ctm
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   240
    fun R_conv rel_eq_onps = Transfer.top_sweep_rewr_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]}
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   241
      then_conv Transfer.bottom_rewr_conv rel_eq_onps
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   242
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   243
    val (ret_lift_def, lthy) = add_lift_def (#lift_config config) var qty rhs rsp_thm par_thms lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   244
  in
60230
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   245
    if (not (#code_dt config) orelse (code_eq_of_lift_def ret_lift_def <> NONE_EQ)
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   246
      andalso (code_eq_of_lift_def ret_lift_def <> UNKNOWN_EQ))
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   247
      (* Let us try even in case of UNKNOWN_EQ. If this leads to problems, the user can always
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   248
        say that they do not want this workaround. *)
60229
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   249
    then  (ret_lift_def, lthy)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   250
    else
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   251
      let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   252
        val lift_def = inst_of_lift_def lthy qty ret_lift_def
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   253
        val rty = rty_of_lift_def lift_def
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   254
        val rty_ret = body_type rty
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   255
        val qty_ret = body_type qty
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   256
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   257
        val (lthy, rel_eq_onps) = prove_code_dt (rty_ret, qty_ret) lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   258
        val code_dt = code_dt_of lthy (rty_ret, qty_ret)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   259
      in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   260
        if is_none code_dt orelse is_none (rep_isom_data_of_code_dt (the code_dt)) then (ret_lift_def, lthy)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   261
        else
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   262
          let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   263
            val code_dt = the code_dt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   264
            val rhs = dest_comb (rhs_of_lift_def lift_def) |> snd
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   265
            val rep_isom_data = code_dt |> rep_isom_data_of_code_dt |> the
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   266
            val qty_code_dt_bundle_name = bundle_name_of_rep_isom_data rep_isom_data
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   267
            val rep_isom = mk_rep_isom_of_code_dt qty_ret code_dt |> the
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   268
            val lthy = Bundle.includes [qty_code_dt_bundle_name] lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   269
            fun qty_isom_of_rep_isom rep = rep |> dest_Const |> snd |> domain_type
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   270
            val qty_isom = qty_isom_of_rep_isom rep_isom
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   271
            val f'_var = (Binding.suffix_name "_aux" (fst var), NoSyn);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   272
            val f'_qty = strip_type qty |> fst |> rpair qty_isom |> op --->
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   273
            val f'_rsp_rel = Lifting_Term.equiv_relation lthy (rty, f'_qty);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   274
            val rsp = rsp_thm_of_lift_def lift_def
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   275
            val rel_eq_onps_conv = HOLogic.Trueprop_conv (ret_rel_conv (R_conv rel_eq_onps))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   276
            val rsp_norm = Conv.fconv_rule rel_eq_onps_conv rsp
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   277
            val f'_rsp_goal = HOLogic.mk_Trueprop (f'_rsp_rel $ rhs $ rhs);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   278
            val f'_rsp = Goal.prove_sorry lthy [] [] f'_rsp_goal
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   279
              (K (HEADGOAL (CONVERSION (rel_eq_onps_conv) THEN' rtac rsp_norm)))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   280
              |> Thm.close_derivation
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   281
            val (f'_lift_def, lthy) = add_lift_def ld_no_notes f'_var f'_qty rhs f'_rsp [] lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   282
            val f'_lift_def = inst_of_lift_def lthy f'_qty f'_lift_def
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   283
            val f'_lift_const = mk_lift_const_of_lift_def f'_qty f'_lift_def
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   284
            val args_lthy = lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   285
            val (args, lthy) = mk_Frees "x" (binder_types qty) lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   286
            val f_alt_def_goal_lhs = list_comb (lift_const_of_lift_def lift_def, args);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   287
            val f_alt_def_goal_rhs = rep_isom $ list_comb (f'_lift_const, args);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   288
            val f_alt_def_goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (f_alt_def_goal_lhs, f_alt_def_goal_rhs));
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   289
            fun f_alt_def_tac ctxt i =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   290
              EVERY' [Transfer.gen_frees_tac [] ctxt, DETERM o Transfer.transfer_tac true ctxt,
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   291
                SELECT_GOAL (Local_Defs.unfold_tac ctxt [id_apply]), rtac refl] i;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   292
            val rep_isom_transfer = transfer_of_rep_isom_data rep_isom_data
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   293
            val (_, transfer_lthy) = Proof_Context.note_thmss "" [((Binding.empty, []),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   294
              [([rep_isom_transfer], [Transfer.transfer_add])])] lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   295
            val f_alt_def = Goal.prove_sorry transfer_lthy [] [] f_alt_def_goal
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   296
              (fn {context = ctxt, prems = _} => HEADGOAL (f_alt_def_tac ctxt))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   297
              |> Thm.close_derivation
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   298
              |> singleton (Variable.export lthy args_lthy)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   299
            val lthy = args_lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   300
            val lthy =  lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   301
              |> Local_Theory.note ((Binding.empty, @{attributes [code]}), [f_alt_def])
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   302
              |> snd
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   303
              |> Lifting_Setup.lifting_forget (pointer_of_rep_isom_data rep_isom_data)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   304
          in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   305
            (ret_lift_def, lthy)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   306
          end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   307
       end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   308
    end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   309
and mk_rep_isom qty_isom_bundle (rty, qty, qty_isom) lthy =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   310
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   311
    val (rty_name, typs) = dest_Type rty
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   312
    val (_, qty_typs) = dest_Type qty
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   313
    val fp = BNF_FP_Def_Sugar.fp_sugar_of lthy rty_name
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   314
    val fp = if is_some fp then the fp
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   315
      else error ("code_dt: " ^ quote rty_name ^ " is not a datatype.")
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   316
    val ctr_sugar = fp |> #fp_ctr_sugar |> #ctr_sugar
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   317
    val ctrs = map (Ctr_Sugar.mk_ctr typs) (#ctrs ctr_sugar);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   318
    val qty_ctrs = map (Ctr_Sugar.mk_ctr qty_typs) (#ctrs ctr_sugar);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   319
    val ctr_Tss = map (dest_Const #> snd #> binder_types) ctrs;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   320
    val qty_ctr_Tss = map (dest_Const #> snd #> binder_types) qty_ctrs;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   321
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   322
    fun lazy_prove_code_dt (rty, qty) lthy =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   323
      if is_none (code_dt_of lthy (rty, qty)) then prove_code_dt (rty, qty) lthy |> fst else lthy;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   324
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   325
    val lthy = fold2 (fold2 (lazy_prove_code_dt oo pair)) ctr_Tss qty_ctr_Tss lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   326
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   327
    val n = length ctrs;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   328
    val ks = 1 upto n;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   329
    val (xss, _) = mk_Freess "x" ctr_Tss lthy;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   330
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   331
    fun sel_retT (rty' as Type (s, rtys'), qty' as Type (s', qtys')) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   332
        if (rty', qty') = (rty, qty) then qty_isom else (if s = s'
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   333
          then Type (s, map sel_retT (rtys' ~~ qtys')) else qty')
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   334
      | sel_retT (_, qty') = qty';
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   335
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   336
    val sel_argss = @{map 5} (fn k => fn xs => @{map 3} (fn x => fn rty => fn qty =>
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   337
      (k, (qty, sel_retT (rty,qty)), (xs, x)))) ks xss xss ctr_Tss qty_ctr_Tss;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   338
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   339
    fun mk_sel_casex (_, _, (_, x)) = Ctr_Sugar.mk_case typs (x |> dest_Free |> snd) (#casex ctr_sugar);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   340
    val dis_casex = Ctr_Sugar.mk_case typs HOLogic.boolT (#casex ctr_sugar);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   341
    fun mk_sel_case_args lthy ctr_Tss ks (k, (qty, _), (xs, x)) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   342
      let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   343
        val T = x |> dest_Free |> snd;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   344
        fun gen_undef_wit Ts wits =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   345
          case code_dt_of lthy (T, qty) of
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   346
            SOME code_dt =>
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   347
              (fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_witness_of_code_dt qty code_dt),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   348
                wit_thm_of_code_dt code_dt :: wits)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   349
            | NONE => (fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T), wits)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   350
      in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   351
        @{fold_map 2} (fn Ts => fn k' => fn wits =>
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   352
          (if k = k' then (fold_rev Term.lambda xs x, wits) else gen_undef_wit Ts wits)) ctr_Tss ks []
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   353
      end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   354
    fun mk_sel_rhs arg =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   355
      let val (sel_rhs, wits) = mk_sel_case_args lthy ctr_Tss ks arg
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   356
      in (arg |> #2, wits, list_comb (mk_sel_casex arg, sel_rhs)) end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   357
    fun mk_dis_case_args args k  = map (fn (k', arg) => (if k = k'
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   358
      then fold_rev Term.lambda arg @{const True} else fold_rev Term.lambda arg @{const False})) args;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   359
    val sel_rhs = map (map mk_sel_rhs) sel_argss
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   360
    val dis_rhs = map (fn k => list_comb (dis_casex, mk_dis_case_args (ks ~~ xss) k)) ks
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   361
    val dis_qty = qty_isom --> HOLogic.boolT;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   362
    val uTname = unique_Tname (rty, qty)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   363
    val dis_names = map (fn k => Binding.qualified true ("dis" ^ string_of_int k) uTname) ks;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   364
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   365
    val (diss, lthy) = @{fold_map 2} (fn b => fn rhs => fn lthy =>
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   366
      lift_def ld_no_notes (b, NoSyn) dis_qty rhs (K all_tac) [] lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   367
      |> apfst (mk_lift_const_of_lift_def dis_qty)) dis_names dis_rhs lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   368
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   369
    fun lift_sel_tac exhaust_rule dt_rules wits ctxt i =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   370
      (Method.insert_tac wits THEN' case_tac exhaust_rule ctxt THEN_ALL_NEW (
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   371
      EVERY' [hyp_subst_tac ctxt, Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq dt_rules),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   372
        REPEAT_DETERM o etac conjE, atac])) i
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   373
    val pred_simps = Transfer.lookup_pred_data lthy (Tname rty) |> the |> Transfer.pred_simps
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   374
    val sel_tac = lift_sel_tac (#exhaust ctr_sugar) (#case_thms ctr_sugar @ pred_simps)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   375
    val sel_names = map (fn (k, xs) => map (fn k' => Binding.qualified true
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   376
      ("sel" ^ string_of_int k ^ string_of_int k') uTname) (1 upto length xs)) (ks ~~ ctr_Tss);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   377
    val (selss, lthy) = @{fold_map 2} (@{fold_map 2} (fn b => fn ((_, qty_ret), wits, rhs) => fn lthy =>
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   378
      lift_def_code_dt { code_dt = true, lift_config = ld_no_notes }
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   379
        (b, NoSyn) (qty_isom --> qty_ret) rhs (HEADGOAL o sel_tac wits) [] lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   380
      |> apfst (mk_lift_const_of_lift_def (qty_isom --> qty_ret)))) sel_names sel_rhs lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   381
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   382
    fun lift_isom_tac ctxt = Local_Defs.unfold_tac ctxt [id_apply] THEN HEADGOAL atac;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   383
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   384
    val (rep_isom_lift_def, lthy) = lift_def ld_no_notes (Binding.qualified true "Rep_isom" uTname, NoSyn)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   385
      (qty_isom --> qty) (HOLogic.id_const rty) lift_isom_tac [] lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   386
      |> apfst (inst_of_lift_def lthy (qty_isom --> qty));
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   387
    val (abs_isom, lthy) = lift_def ld_no_notes (Binding.qualified true "Abs_isom" uTname, NoSyn)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   388
      (qty --> qty_isom) (HOLogic.id_const rty) lift_isom_tac [] lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   389
      |> apfst (mk_lift_const_of_lift_def (qty --> qty_isom));
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   390
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   391
    fun mk_type_definition newT oldT RepC AbsC A =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   392
      let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   393
        val typedefC =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   394
          Const (@{const_name type_definition},
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   395
            (newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   396
      in typedefC $ RepC $ AbsC $ A end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   397
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   398
    val rep_isom = lift_const_of_lift_def rep_isom_lift_def
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   399
    val typedef_goal = mk_type_definition qty_isom qty rep_isom abs_isom (HOLogic.mk_UNIV qty) |>
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   400
      HOLogic.mk_Trueprop;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   401
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   402
      fun typ_isom_tac ctxt i =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   403
        EVERY' [ SELECT_GOAL (Local_Defs.unfold_tac ctxt @{thms type_definition_def}),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   404
          DETERM o Transfer.transfer_tac true ctxt, Raw_Simplifier.rewrite_goal_tac ctxt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   405
            (map safe_mk_meta_eq @{thms id_apply simp_thms Ball_def}),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   406
           rtac TrueI] i;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   407
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   408
    val (_, transfer_lthy) = Proof_Context.note_thmss "" [((Binding.empty, []),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   409
      [(@{thms right_total_UNIV_transfer},[Transfer.transfer_add]),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   410
       (@{thms Domain_eq_top}, [Transfer.transfer_domain_add]) ])] lthy;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   411
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   412
    val quot_thm_isom = Goal.prove_sorry transfer_lthy [] [] typedef_goal
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   413
      (fn {context = ctxt, prems = _} => typ_isom_tac ctxt 1)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   414
      |> Thm.close_derivation
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   415
      |> singleton (Variable.export transfer_lthy lthy)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   416
      |> (fn thm => @{thm UNIV_typedef_to_Quotient} OF [thm, @{thm reflexive}])
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   417
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   418
    val qty_isom_name = Tname qty_isom;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   419
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   420
    val quot_isom_rep =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   421
      let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   422
        val (quotients : Lifting_Term.quotients) = Symtab.insert (Lifting_Info.quotient_eq) (qty_isom_name,
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   423
          {quot_thm = quot_thm_isom, pcr_info = NONE}) Symtab.empty
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   424
        val id_actions = { constr = K I, lift = K I, comp_lift = K I }
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   425
      in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   426
        fn ctxt => fn (rty, qty) => Lifting_Term.prove_schematic_quot_thm id_actions quotients
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   427
          ctxt (rty, qty) () |> fst |> Lifting_Term.force_qty_type ctxt qty
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   428
          |> quot_thm_rep
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   429
      end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   430
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   431
    val x_lthy = lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   432
    val (x, lthy) = yield_singleton (mk_Frees "x") qty_isom lthy;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   433
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   434
    fun mk_ctr ctr ctr_Ts sels =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   435
      let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   436
        val sel_ret_Ts = map (dest_Const #> snd #> body_type) sels;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   437
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   438
        fun rep_isom lthy t (rty, qty) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   439
          let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   440
            val rep = quot_isom_rep lthy (rty, qty)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   441
          in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   442
            if is_Const rep andalso (rep |> dest_Const |> fst) = @{const_name id} then
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   443
              t else rep $ t
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   444
          end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   445
      in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   446
        @{fold 3} (fn sel => fn ctr_T => fn sel_ret_T => fn ctr =>
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   447
          ctr $ rep_isom lthy (sel $ x) (ctr_T, sel_ret_T)) sels ctr_Ts sel_ret_Ts ctr
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   448
      end;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   449
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   450
    (* stolen from Metis *)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   451
    exception BREAK_LIST
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   452
    fun break_list (x :: xs) = (x, xs)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   453
      | break_list _ = raise BREAK_LIST
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   454
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   455
    val (ctr, ctrs) = qty_ctrs |> rev |> break_list;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   456
    val (ctr_Ts, ctr_Tss) = qty_ctr_Tss |> rev |> break_list;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   457
    val (sel, rselss) = selss |> rev |> break_list;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   458
    val rdiss = rev diss |> tl;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   459
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   460
    val first_ctr = mk_ctr ctr ctr_Ts sel;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   461
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   462
    fun mk_If_ctr dis ctr ctr_Ts sel elsex = mk_If (dis$x) (mk_ctr ctr ctr_Ts sel) elsex;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   463
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   464
    val rhs = @{fold 4} mk_If_ctr rdiss ctrs ctr_Tss rselss first_ctr;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   465
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   466
    val rep_isom_code_goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (rep_isom$x, rhs));
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   467
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   468
    local
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   469
      val rep_isom_code_tac_rules = map safe_mk_meta_eq @{thms refl id_apply if_splits simp_thms}
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   470
    in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   471
      fun rep_isom_code_tac (ctr_sugar:Ctr_Sugar.ctr_sugar) ctxt i =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   472
        let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   473
          val exhaust = ctr_sugar |> #exhaust
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   474
          val cases = ctr_sugar |> #case_thms
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   475
          val map_ids = fp |> #fp_nesting_bnfs |> map BNF_Def.map_id0_of_bnf
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   476
          val simp_rules = map safe_mk_meta_eq (cases @ map_ids) @ rep_isom_code_tac_rules
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   477
        in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   478
          EVERY' [Transfer.gen_frees_tac [] ctxt, DETERM o (Transfer.transfer_tac true ctxt),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   479
            case_tac exhaust ctxt THEN_ALL_NEW EVERY' [hyp_subst_tac ctxt,
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   480
              Raw_Simplifier.rewrite_goal_tac ctxt simp_rules, rtac TrueI ]] i
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   481
        end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   482
    end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   483
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   484
    val rep_isom_code = Goal.prove_sorry lthy [] [] rep_isom_code_goal
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   485
      (fn {context = ctxt, prems = _} => rep_isom_code_tac ctr_sugar ctxt 1)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   486
      |> Thm.close_derivation
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   487
      |> singleton(Variable.export lthy x_lthy)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   488
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   489
    val lthy = x_lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   490
    val pointer = Lifting_Setup.pointer_of_bundle_binding lthy qty_isom_bundle
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   491
    fun code_dt phi context = code_dt_of lthy (rty, qty) |> the |>
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   492
      update_rep_isom rep_isom (transfer_rules_of_lift_def rep_isom_lift_def |> hd)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   493
       (bundle_name_of_bundle_binding qty_isom_bundle phi context) pointer;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   494
    val lthy =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   495
      lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   496
      |> snd o Local_Theory.note ((Binding.empty, @{attributes [code]}), [rep_isom_code])
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   497
      |> Local_Theory.declaration {syntax = false, pervasive = true}
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   498
        (fn phi => fn context => Data.map (Item_Net.update (morph_code_dt phi (code_dt phi context))) context)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   499
      |> Local_Theory.restore
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   500
      |> Lifting_Setup.lifting_forget pointer
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   501
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   502
    ((selss, diss, rep_isom_code), lthy)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   503
  end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   504
and constr qty (quot_thm, (lthy, rel_eq_onps)) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   505
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   506
    val quot_thm = Lifting_Term.force_qty_type lthy qty quot_thm
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   507
    val (rty, qty) = quot_thm_rty_qty quot_thm
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   508
    val rty_name = Tname rty;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   509
    val pred_data = Transfer.lookup_pred_data lthy rty_name
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   510
    val pred_data = if is_some pred_data then the pred_data
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   511
      else error ("code_dt: " ^ quote rty_name ^ " is not a datatype.")
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   512
    val rel_eq_onp = safe_mk_meta_eq (Transfer.rel_eq_onp pred_data);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   513
    val R_conv = Transfer.top_sweep_rewr_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]}
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   514
      then_conv Conv.rewr_conv rel_eq_onp
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   515
    val quot_thm = Conv.fconv_rule(HOLogic.Trueprop_conv (Quotient_R_conv R_conv)) quot_thm;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   516
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   517
    if is_none (code_dt_of lthy (rty, qty)) then
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   518
      let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   519
        val non_empty_pred = quot_thm RS @{thm type_definition_Quotient_not_empty}
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   520
        val pred = quot_thm_rel quot_thm |> dest_comb |> snd;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   521
        val (pred, lthy) = yield_singleton (Variable.import_terms true) pred lthy;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   522
        val TFrees = Term.add_tfreesT qty []
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   523
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   524
        fun non_empty_typedef_tac non_empty_pred ctxt i =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   525
          (SELECT_GOAL (Local_Defs.unfold_tac ctxt [mem_Collect_eq]) THEN' rtac non_empty_pred) i
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   526
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   527
        val uTname = unique_Tname (rty, qty)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   528
        val Tdef_set = HOLogic.mk_Collect ("x", rty, pred $ Free("x", rty));
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   529
        val ((_, tcode_dt), lthy) = conceal_naming_result (typedef (Binding.concealed uTname, TFrees, NoSyn)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   530
          Tdef_set NONE (fn lthy => non_empty_typedef_tac non_empty_pred lthy 1)) lthy;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   531
        val type_definition_thm = tcode_dt |> snd |> #type_definition;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   532
        val qty_isom = tcode_dt |> fst |> #abs_type;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   533
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   534
        val config = { notes = false}
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   535
        val (binding, lthy) = conceal_naming_result (Lifting_Setup.setup_by_typedef_thm
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   536
          config type_definition_thm) lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   537
        val lthy = Local_Theory.restore lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   538
        val (wit, wit_thm) = mk_witness quot_thm;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   539
        val code_dt = mk_code_dt rty qty wit wit_thm NONE;
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   540
        val lthy = lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   541
          |> update_code_dt code_dt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   542
          |> Local_Theory.restore
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   543
          |> mk_rep_isom binding (rty, qty, qty_isom) |> snd
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   544
      in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   545
        (quot_thm, (lthy, rel_eq_onp :: rel_eq_onps))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   546
      end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   547
    else
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   548
      (quot_thm, (lthy, rel_eq_onp :: rel_eq_onps))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   549
  end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   550
and lift_def_code_dt config var qty rhs tac par_thms lthy = gen_lift_def (add_lift_def_code_dt config)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   551
  var qty rhs tac par_thms lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   552
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   553
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   554
(** from parsed parameters to the config record **)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   555
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   556
fun map_config_code_dt f1 f2 { code_dt = code_dt, lift_config = lift_config } =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   557
  { code_dt = f1 code_dt, lift_config = f2 lift_config }
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   558
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   559
fun update_config_code_dt nval = map_config_code_dt (K nval) I
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   560
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   561
val config_flags = [("code_dt", update_config_code_dt true)]
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   562
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   563
fun evaluate_params params =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   564
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   565
    fun eval_param param config =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   566
      case AList.lookup (op =) config_flags param of
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   567
        SOME update => update config
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   568
        | NONE => error ("Unknown parameter: " ^ (quote param))
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   569
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   570
    fold eval_param params default_config_code_dt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   571
  end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   572
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   573
(**
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   574
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   575
  lift_definition command. It opens a proof of a corresponding respectfulness
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   576
  theorem in a user-friendly, readable form. Then add_lift_def_code_dt is called internally.
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   577
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   578
**)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   579
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   580
fun lift_def_cmd (params, raw_var, rhs_raw, par_xthms) lthy =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   581
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   582
    val config = evaluate_params params
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   583
    val ((binding, SOME qty, mx), lthy) = yield_singleton Proof_Context.read_vars raw_var lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   584
    val var = (binding, mx)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   585
    val rhs = (Syntax.check_term lthy o Syntax.parse_term lthy) rhs_raw
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   586
    val par_thms = Attrib.eval_thms lthy par_xthms
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   587
    val (goal, after_qed) = prepare_lift_def (add_lift_def_code_dt config) var qty rhs par_thms lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   588
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   589
    Proof.theorem NONE (snd oo after_qed) [map (rpair []) (the_list goal)] lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   590
  end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   591
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   592
fun quot_thm_err ctxt (rty, qty) pretty_msg =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   593
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   594
    val error_msg = cat_lines
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   595
       ["Lifting failed for the following types:",
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   596
        Pretty.string_of (Pretty.block
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   597
         [Pretty.str "Raw type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty]),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   598
        Pretty.string_of (Pretty.block
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   599
         [Pretty.str "Abstract type:", Pretty.brk 2, Syntax.pretty_typ ctxt qty]),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   600
        "",
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   601
        (Pretty.string_of (Pretty.block
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   602
         [Pretty.str "Reason:", Pretty.brk 2, pretty_msg]))]
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   603
  in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   604
    error error_msg
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   605
  end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   606
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   607
fun check_rty_err ctxt (rty_schematic, rty_forced) (raw_var, rhs_raw) =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   608
  let
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   609
    val (_, ctxt') = yield_singleton Proof_Context.read_vars raw_var ctxt
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   610
    val rhs = (Syntax.check_term ctxt' o Syntax.parse_term ctxt') rhs_raw
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   611
    val error_msg = cat_lines
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   612
       ["Lifting failed for the following term:",
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   613
        Pretty.string_of (Pretty.block
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   614
         [Pretty.str "Term:", Pretty.brk 2, Syntax.pretty_term ctxt rhs]),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   615
        Pretty.string_of (Pretty.block
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   616
         [Pretty.str "Type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty_schematic]),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   617
        "",
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   618
        (Pretty.string_of (Pretty.block
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   619
         [Pretty.str "Reason:",
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   620
          Pretty.brk 2,
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   621
          Pretty.str "The type of the term cannot be instantiated to",
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   622
          Pretty.brk 1,
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   623
          Pretty.quote (Syntax.pretty_typ ctxt rty_forced),
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   624
          Pretty.str "."]))]
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   625
    in
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   626
      error error_msg
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   627
    end
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   628
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   629
fun lift_def_cmd_with_err_handling (params, (raw_var, rhs_raw, par_xthms)) lthy =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   630
  (lift_def_cmd (params, raw_var, rhs_raw, par_xthms) lthy
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   631
    handle Lifting_Term.QUOT_THM (rty, qty, msg) => quot_thm_err lthy (rty, qty) msg)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   632
    handle Lifting_Term.CHECK_RTY (rty_schematic, rty_forced) =>
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   633
      check_rty_err lthy (rty_schematic, rty_forced) (raw_var, rhs_raw);
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   634
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   635
val parse_param = Parse.name
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   636
val parse_params = Scan.optional (Args.parens (Parse.list parse_param)) [];
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   637
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   638
(* parser and command *)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   639
val liftdef_parser =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   640
  parse_params --
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   641
  (((Parse.binding -- (@{keyword "::"} |-- (Parse.typ >> SOME) -- Parse.opt_mixfix') >> Parse.triple2)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   642
    --| @{keyword "is"} -- Parse.term --
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   643
      Scan.optional (@{keyword "parametric"} |-- Parse.!!! Parse.xthms1) []) >> Parse.triple1)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   644
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   645
val _ =
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   646
  Outer_Syntax.local_theory_to_proof @{command_keyword "lift_definition"}
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   647
    "definition for constants over the quotient type"
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   648
      (liftdef_parser >> lift_def_cmd_with_err_handling)
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   649
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   650
end