src/HOL/Tools/Lifting/lifting_def_code_dt.ML
author kuncar
Sat, 02 May 2015 13:58:06 +0200
changeset 60235 3cab6f891c2f
parent 60234 17ff843807cb
child 60236 865741686926
permissions -rw-r--r--
reorder some steps in the construction to support mutual datatypes
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
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
    49
                                                                       
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
    50
open Ctr_Sugar_Util BNF_Util BNF_FP_Util BNF_FP_Def_Sugar Lifting_Def Lifting_Util
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
    51
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
    52
infix 0 MRSL
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
    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
(** 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
    55
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
(* 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
    57
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
    58
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
    59
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
    60
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
    61
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
    62
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
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
    64
  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
    65
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
    66
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
    67
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
    68
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
    69
fun rep_isom_data_of_code_dt (CODE_DT code_dt) = #rep_isom_data code_dt;
60232
29ac1c6a1fbb equivalence in code_dt data structure must respect both rty and qty
kuncar
parents: 60231
diff changeset
    70
fun ty_alpha_equiv (T, U) = Type.raw_instance (T, U) andalso Type.raw_instance (U, T);
29ac1c6a1fbb equivalence in code_dt data structure must respect both rty and qty
kuncar
parents: 60231
diff changeset
    71
fun code_dt_eq c = (ty_alpha_equiv o apply2 rty_of_code_dt) c 
29ac1c6a1fbb equivalence in code_dt data structure must respect both rty and qty
kuncar
parents: 60231
diff changeset
    72
  andalso (ty_alpha_equiv o apply2 qty_of_code_dt) c;
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
    73
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
    74
  |> 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
    75
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
(* 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
    77
fun is_code_dt_of_type (rty, qty) code_dt = code_dt |> `rty_of_code_dt ||> qty_of_code_dt
60232
29ac1c6a1fbb equivalence in code_dt data structure must respect both rty and qty
kuncar
parents: 60231
diff changeset
    78
  |> HOLogic.mk_prodT |> curry ty_alpha_equiv (HOLogic.mk_prodT (rty, qty));
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
    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_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
    81
  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
    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 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
    84
  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
    85
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
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
    87
  (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
    88
  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
    89
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
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
    91
  (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
    92
  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
    93
    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
    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 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
    96
  (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
    97
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
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
    99
  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
   100
    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
   101
    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
   102
    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
   103
  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
   104
    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
   105
  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
   106
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
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
   108
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
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
   110
(
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
  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
   112
  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
   113
  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
   114
  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
   115
);
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
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
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
   118
  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
   119
    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
   120
    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
   121
  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
   122
    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
   123
    |> 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
   124
  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
   125
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
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
   127
  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
   128
    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
   129
    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
   130
  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
   131
    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
   132
  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
   133
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
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
   135
  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
   136
    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
   137
    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
   138
  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
   139
    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
   140
  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
   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
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
   143
    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
   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
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
   146
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
   147
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
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
   149
  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
   150
    (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
   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_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
   153
  |> 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
   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_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
   156
  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
   157
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
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
   159
  (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
   160
    (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
   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
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
(** 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
   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 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
   166
  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
   167
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
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
   169
  | 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
   170
  | 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
   171
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
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
   173
  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
   174
    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
   175
  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
   176
    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
   177
  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
   178
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
(** 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
   180
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
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
   182
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
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
   184
  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
   185
    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
   186
    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
   187
  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
   188
    (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
   189
  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
   190
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
(** 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
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
   194
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
   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
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
(** 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
   198
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
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
   200
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
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
   202
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
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
   204
  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
   205
    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
   206
    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
   207
  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
   208
    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
   209
      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
   210
        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
   211
          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
   212
        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
   213
      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
   214
        (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
   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
    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
   217
      (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
   218
  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
   219
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
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
   221
  (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
   222
    (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
   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 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
   225
  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
   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_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
   228
 (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
   229
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
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
   231
  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
   232
    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
   233
      { 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
   234
  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
   235
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
   236
  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
   237
    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
   238
    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
   239
      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
   240
        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
   241
          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
   242
        | _ => 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
   243
    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
   244
      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
   245
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   246
    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
   247
  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
   248
    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
   249
      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
   250
      (* 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
   251
        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
   252
    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
   253
    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
   254
      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
   255
        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
   256
        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
   257
        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
   258
        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
   259
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
        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
   261
        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
   262
      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
   263
        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
   264
        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
   265
          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
   266
            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
   267
            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
   268
            val rep_isom_data = code_dt |> rep_isom_data_of_code_dt |> the
60235
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   269
            val pointer = pointer_of_rep_isom_data rep_isom_data
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   270
            val quot_active = 
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   271
              Lifting_Info.lookup_restore_data lthy pointer |> the |> #quotient |> #quot_thm
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   272
              |> Lifting_Info.lookup_quot_thm_quotients lthy |> is_some
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
   273
            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
   274
            val rep_isom = mk_rep_isom_of_code_dt qty_ret code_dt |> the
60235
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   275
            val lthy = if quot_active then lthy else Bundle.includes [qty_code_dt_bundle_name] lthy
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
   276
            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
   277
            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
   278
            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
   279
            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
   280
            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
   281
            val rsp = rsp_thm_of_lift_def lift_def
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   282
            val rel_eq_onps_conv = HOLogic.Trueprop_conv (Conv.fun2_conv (ret_rel_conv (R_conv rel_eq_onps)))
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
   283
            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
   284
            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
   285
            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
   286
              (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
   287
              |> 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
   288
            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
   289
            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
   290
            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
   291
            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
   292
            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
   293
            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
   294
            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
   295
            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
   296
            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
   297
              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
   298
                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
   299
            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
   300
            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
   301
              [([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
   302
            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
   303
              (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
   304
              |> 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
   305
              |> 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
   306
            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
   307
            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
   308
              |> 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
   309
              |> snd
60235
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   310
              (* if processing a mutual datatype (there is a cycle!) the corresponding quotient 
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   311
                 will be needed later and will be forgotten later *)
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   312
              |> (if quot_active then I else Lifting_Setup.lifting_forget pointer)
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
   313
          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
   314
            (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
   315
          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
   316
       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
   317
    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
   318
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
   319
  let
60235
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   320
    (* logical definition of qty qty_isom isomorphism *) 
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   321
    val uTname = unique_Tname (rty, qty)
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   322
    fun eq_onp_to_top_tac ctxt = SELECT_GOAL (Local_Defs.unfold_tac ctxt 
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   323
      (@{thm eq_onp_top_eq_eq[symmetric]} :: Lifting_Info.get_relator_eq_onp_rules ctxt))
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   324
    fun lift_isom_tac ctxt = HEADGOAL (eq_onp_to_top_tac ctxt
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   325
      THEN' (rtac @{thm id_transfer}));
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   326
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   327
    val (rep_isom_lift_def, lthy) = lift_def ld_no_notes (Binding.qualified true "Rep_isom" uTname, NoSyn)
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   328
      (qty_isom --> qty) (HOLogic.id_const rty) lift_isom_tac [] lthy
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   329
      |> apfst (inst_of_lift_def lthy (qty_isom --> qty));
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   330
    val (abs_isom, lthy) = lift_def ld_no_notes (Binding.qualified true "Abs_isom" uTname, NoSyn)
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   331
      (qty --> qty_isom) (HOLogic.id_const rty) lift_isom_tac [] lthy
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   332
      |> apfst (mk_lift_const_of_lift_def (qty --> qty_isom));
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   333
    val rep_isom = lift_const_of_lift_def rep_isom_lift_def
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   334
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   335
    val pointer = Lifting_Setup.pointer_of_bundle_binding lthy qty_isom_bundle
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   336
    fun code_dt phi context = code_dt_of lthy (rty, qty) |> the |>
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   337
      update_rep_isom rep_isom (transfer_rules_of_lift_def rep_isom_lift_def |> hd)
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   338
       (bundle_name_of_bundle_binding qty_isom_bundle phi context) pointer;
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   339
    val lthy = lthy  
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   340
      |> Local_Theory.declaration {syntax = false, pervasive = true}
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   341
        (fn phi => fn context => Data.map (Item_Net.update (morph_code_dt phi (code_dt phi context))) context)
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   342
      |> Local_Theory.restore
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   343
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   344
    (* in order to make the qty qty_isom isomorphism executable we have to define discriminators 
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   345
      and selectors for qty_isom *)
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
   346
    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
   347
    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
   348
    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
   349
    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
   350
      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
   351
    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
   352
    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
   353
    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
   354
    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
   355
    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
   356
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
    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
   358
    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
   359
    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
   360
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
    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
   362
        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
   363
          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
   364
      | 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
   365
60233
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   366
    val sel_retTs = map2 (map2 (sel_retT oo pair)) ctr_Tss qty_ctr_Tss
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   367
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   368
    fun lazy_prove_code_dt (rty, qty) lthy =
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   369
      if is_none (code_dt_of lthy (rty, qty)) then prove_code_dt (rty, qty) lthy |> fst else lthy;
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   370
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   371
    val lthy = fold2 (fold2 (lazy_prove_code_dt oo pair)) ctr_Tss sel_retTs lthy
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   372
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   373
    val sel_argss = @{map 4} (fn k => fn xs => @{map 2} (fn x => fn qty_ret => 
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   374
      (k, qty_ret, (xs, x)))) ks xss xss sel_retTs;
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
   375
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
    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
   377
    val dis_casex = Ctr_Sugar.mk_case typs HOLogic.boolT (#casex ctr_sugar);
60233
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   378
    fun mk_sel_case_args lthy ctr_Tss ks (k, qty_ret, (xs, x)) =
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
   379
      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
   380
        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
   381
        fun gen_undef_wit Ts wits =
60233
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   382
          case code_dt_of lthy (T, qty_ret) of
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
   383
            SOME code_dt =>
60233
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   384
              (fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_witness_of_code_dt qty_ret code_dt),
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
   385
                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
   386
            | 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
   387
      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
   388
        @{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
   389
          (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
   390
      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
   391
    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
   392
      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
   393
      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
   394
    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
   395
      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
   396
    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
   397
    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
   398
    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
   399
    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
   400
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
    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
   402
      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
   403
      |> 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
   404
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   405
    val unfold_lift_sel_rsp = @{lemma "(\<And>x. P1 x \<Longrightarrow> P2 (f x)) \<Longrightarrow> (rel_fun (eq_onp P1) (eq_onp P2)) f f"
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   406
      by (simp add: eq_onp_same_args rel_fun_eq_onp_rel)}
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   407
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
   408
    fun lift_sel_tac exhaust_rule dt_rules wits ctxt i =
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   409
        (Method.insert_tac wits THEN' 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   410
         eq_onp_to_top_tac ctxt THEN' (* normalize *)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   411
         rtac unfold_lift_sel_rsp THEN'
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   412
         case_tac exhaust_rule ctxt THEN_ALL_NEW (
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   413
        EVERY' [hyp_subst_tac ctxt, (* does not kill wits because = was rewritten to eq_onp top *)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   414
        Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq dt_rules), 
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
   415
        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
   416
    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
   417
    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
   418
    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
   419
      ("sel" ^ string_of_int k ^ string_of_int k') uTname) (1 upto length xs)) (ks ~~ ctr_Tss);
60233
89d500d7e3eb better precomputing
kuncar
parents: 60232
diff changeset
   420
    val (selss, lthy) = @{fold_map 2} (@{fold_map 2} (fn b => fn (qty_ret, wits, rhs) => fn lthy =>
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   421
        lift_def_code_dt { code_dt = true, lift_config = ld_no_notes }
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
   422
        (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
   423
      |> 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
   424
60235
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   425
    (* now we can execute the qty qty_isom isomorphism *)
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
   426
    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
   427
      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
   428
        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
   429
          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
   430
            (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
   431
      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
   432
    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
   433
      HOLogic.mk_Trueprop;
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   434
    fun typ_isom_tac ctxt i =
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   435
      EVERY' [ SELECT_GOAL (Local_Defs.unfold_tac ctxt @{thms type_definition_def}),
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   436
        DETERM o Transfer.transfer_tac true ctxt,
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   437
          SELECT_GOAL (Local_Defs.unfold_tac ctxt @{thms eq_onp_top_eq_eq}) (* normalize *), 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   438
          Raw_Simplifier.rewrite_goal_tac ctxt 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   439
          (map safe_mk_meta_eq @{thms id_apply simp_thms Ball_def}),
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   440
         rtac TrueI] i;
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
   441
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
    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
   443
      [(@{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
   444
       (@{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
   445
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
    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
   447
      (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
   448
      |> 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
   449
      |> 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
   450
      |> (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
   451
    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
   452
    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
   453
      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
   454
        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
   455
          {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
   456
        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
   457
      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
   458
        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
   459
          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
   460
          |> 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
   461
      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
   462
    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
   463
    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
   464
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
    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
   466
      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
   467
        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
   468
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
        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
   470
          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
   471
            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
   472
          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
   473
            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
   474
              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
   475
          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
   476
      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
   477
        @{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
   478
          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
   479
      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
   480
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
    (* 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
   482
    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
   483
    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
   484
      | 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
   485
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
    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
   487
    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
   488
    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
   489
    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
   490
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
    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
   492
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
    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
   494
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
    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
   496
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
    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
   498
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
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
      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
   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
      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
   503
        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
   504
          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
   505
          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
   506
          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
   507
          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
   508
        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
   509
          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
   510
            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
   511
              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
   512
        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
   513
    end
60234
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   514
    
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   515
    (* stolen from bnf_fp_n2m.ML *)
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   516
    fun force_typ ctxt T =
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   517
      Term.map_types Type_Infer.paramify_vars
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   518
      #> Type.constraint T
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   519
      #> singleton (Type_Infer_Context.infer_types ctxt);
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   520
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   521
    (* The following tests that types in rty have corresponding arities imposed by constraints of
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   522
       the datatype fp. Otherwise rep_isom_code_tac could fail (especially transfer in it) is such
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   523
       a way that it is not easy to infer the problem with sorts.
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   524
    *)
17ff843807cb more readable error message if some types do not correspond to sort constraints of the datatype
kuncar
parents: 60233
diff changeset
   525
    val _ = yield_singleton (mk_Frees "x") (#T fp) lthy |> fst |> force_typ lthy qty
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
   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 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
   528
      (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
   529
      |> 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
   530
      |> 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
   531
    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
   532
    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
   533
      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
   534
      |> 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
   535
      |> 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
   536
  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
   537
    ((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
   538
  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
   539
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
   540
  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
   541
    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
   542
    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
   543
    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
   544
    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
   545
    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
   546
      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
   547
    val rel_eq_onp = safe_mk_meta_eq (Transfer.rel_eq_onp pred_data);
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   548
    val rel_eq_onps = insert Thm.eq_thm rel_eq_onp rel_eq_onps
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
   549
    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
   550
      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
   551
    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
   552
  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
   553
    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
   554
      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
   555
        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
   556
        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
   557
        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
   558
        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
   559
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
        fun non_empty_typedef_tac non_empty_pred ctxt i =
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   561
          (Method.insert_tac [non_empty_pred] THEN' 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   562
            SELECT_GOAL (Local_Defs.unfold_tac ctxt [mem_Collect_eq]) THEN' atac) i
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
   563
        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
   564
        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
   565
        val ((_, tcode_dt), lthy) = conceal_naming_result (typedef (Binding.concealed uTname, TFrees, NoSyn)
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   566
          Tdef_set NONE (fn lthy => HEADGOAL (non_empty_typedef_tac non_empty_pred lthy))) lthy;
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
   567
        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
   568
        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
   569
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
        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
   571
        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
   572
          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
   573
        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
   574
        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
   575
        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
   576
        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
   577
          |> 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
   578
          |> 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
   579
          |> 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
   580
      in
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   581
        (quot_thm, (lthy, rel_eq_onps))
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
   582
      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
   583
    else
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   584
      (quot_thm, (lthy, rel_eq_onps))
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
   585
  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
   586
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
   587
  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
   588
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
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
(** 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
   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 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
   593
  { 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
   594
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
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
   596
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
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
   598
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
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
   600
  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
   601
    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
   602
      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
   603
        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
   604
        | 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
   605
  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
   606
    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
   607
  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
   608
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
(**
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
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
  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
   612
  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
   613
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
**)
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
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   616
local
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   617
  val eq_onp_assms_tac_fixed_rules = map (Transfer.prep_transfer_domain_thm @{context})
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   618
    [@{thm pcr_Domainp_total}, @{thm pcr_Domainp_par_left_total}, @{thm pcr_Domainp_par}, 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   619
      @{thm pcr_Domainp}]
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   620
in
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   621
fun mk_readable_rsp_thm_eq tm lthy =
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   622
  let
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   623
    val ctm = Thm.cterm_of lthy tm
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   624
    
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   625
    fun assms_rewr_conv tactic rule ct =
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   626
      let
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   627
        fun prove_extra_assms thm =
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   628
          let
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   629
            val assms = cprems_of thm
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   630
            fun finish thm = if Thm.no_prems thm then SOME (Goal.conclude thm) else NONE
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   631
            fun prove ctm = Option.mapPartial finish (SINGLE tactic (Goal.init ctm))
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   632
          in
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   633
            map_interrupt prove assms
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   634
          end
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   635
    
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   636
        fun cconl_of thm = Drule.strip_imp_concl (Thm.cprop_of thm)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   637
        fun lhs_of thm = fst (Thm.dest_equals (cconl_of thm))
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   638
        fun rhs_of thm = snd (Thm.dest_equals (cconl_of thm))
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   639
        val rule1 = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) rule;
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   640
        val lhs = lhs_of rule1;
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   641
        val rule2 = Thm.rename_boundvars (Thm.term_of lhs) (Thm.term_of ct) rule1;
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   642
        val rule3 =
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   643
          Thm.instantiate (Thm.match (lhs, ct)) rule2
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   644
            handle Pattern.MATCH => raise CTERM ("assms_rewr_conv", [lhs, ct]);
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   645
        val proved_assms = prove_extra_assms rule3
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   646
      in
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   647
        case proved_assms of
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   648
          SOME proved_assms =>
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   649
            let
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   650
              val rule3 = proved_assms MRSL rule3
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   651
              val rule4 =
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   652
                if lhs_of rule3 aconvc ct then rule3
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   653
                else
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   654
                  let val ceq = Thm.dest_fun2 (Thm.cprop_of rule3)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   655
                  in rule3 COMP Thm.trivial (Thm.mk_binop ceq ct (rhs_of rule3)) end
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   656
            in Thm.transitive rule4 (Thm.beta_conversion true (rhs_of rule4)) end
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   657
          | NONE => Conv.no_conv ct
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   658
      end
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   659
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   660
    fun assms_rewrs_conv tactic rules = Conv.first_conv (map (assms_rewr_conv tactic) rules)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   661
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   662
    fun simp_arrows_conv ctm =
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   663
      let
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   664
        val unfold_conv = Conv.rewrs_conv 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   665
          [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   666
            @{thm rel_fun_eq_onp_rel[THEN eq_reflection]},
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   667
            @{thm rel_fun_eq[THEN eq_reflection]},
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   668
            @{thm rel_fun_eq_rel[THEN eq_reflection]}, 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   669
            @{thm rel_fun_def[THEN eq_reflection]}]
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   670
        fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   671
        val eq_onp_assms_tac_rules = @{thm left_unique_OO} :: 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   672
            eq_onp_assms_tac_fixed_rules @ (Transfer.get_transfer_raw lthy)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   673
        val intro_top_rule = @{thm eq_onp_top_eq_eq[symmetric, THEN eq_reflection]}
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   674
        val kill_tops = Transfer.top_sweep_rewr_conv [@{thm eq_onp_top_eq_eq[THEN eq_reflection]}]
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   675
        val eq_onp_assms_tac = (CONVERSION kill_tops THEN' 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   676
          TRY o REPEAT_ALL_NEW (resolve_tac lthy eq_onp_assms_tac_rules) 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   677
          THEN_ALL_NEW (DETERM o Transfer.eq_tac lthy)) 1
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   678
        val relator_eq_onp_conv = Conv.bottom_conv
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   679
          (K (Conv.try_conv (assms_rewrs_conv eq_onp_assms_tac
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   680
            (intro_top_rule :: Lifting_Info.get_relator_eq_onp_rules lthy)))) lthy
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   681
          then_conv kill_tops
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   682
        val relator_eq_conv = Conv.bottom_conv
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   683
          (K (Conv.try_conv (Conv.rewrs_conv (Transfer.get_relator_eq lthy)))) lthy
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   684
      in
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   685
        case (Thm.term_of ctm) of
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   686
          Const (@{const_name "rel_fun"}, _) $ _ $ _ => 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   687
            (binop_conv2 simp_arrows_conv simp_arrows_conv then_conv unfold_conv) ctm
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   688
          | _ => (relator_eq_onp_conv then_conv relator_eq_conv) ctm
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   689
      end
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   690
    
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   691
    val unfold_ret_val_invs = Conv.bottom_conv 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   692
      (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) lthy
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   693
    val unfold_inv_conv = 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   694
      Conv.top_sweep_conv (K (Conv.rewr_conv @{thm eq_onp_def[THEN eq_reflection]})) lthy
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   695
    val simp_conv = HOLogic.Trueprop_conv (Conv.fun2_conv simp_arrows_conv)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   696
    val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]}
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   697
    val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) lthy
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   698
    val beta_conv = Thm.beta_conversion true
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   699
    val eq_thm = 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   700
      (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   701
         then_conv unfold_inv_conv) ctm
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   702
  in
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   703
    Object_Logic.rulify lthy (eq_thm RS Drule.equal_elim_rule2)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   704
  end
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   705
end
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   706
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   707
fun rename_to_tnames ctxt term =
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   708
  let
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   709
    fun all_typs (Const (@{const_name Pure.all}, _) $ Abs (_, T, t)) = T :: all_typs t
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   710
      | all_typs _ = []
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   711
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   712
    fun rename (Const (@{const_name Pure.all}, T1) $ Abs (_, T2, t)) (new_name :: names) = 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   713
        (Const (@{const_name Pure.all}, T1) $ Abs (new_name, T2, rename t names)) 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   714
      | rename t _ = t
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   715
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   716
    val (fixed_def_t, _) = yield_singleton (Variable.importT_terms) term ctxt
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   717
    val new_names = Old_Datatype_Prop.make_tnames (all_typs fixed_def_t)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   718
  in
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   719
    rename term new_names
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   720
  end
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   721
60235
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   722
fun quot_thm_err ctxt (rty, qty) pretty_msg =
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   723
  let
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   724
    val error_msg = cat_lines
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   725
       ["Lifting failed for the following types:",
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   726
        Pretty.string_of (Pretty.block
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   727
         [Pretty.str "Raw type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty]),
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   728
        Pretty.string_of (Pretty.block
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   729
         [Pretty.str "Abstract type:", Pretty.brk 2, Syntax.pretty_typ ctxt qty]),
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   730
        "",
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   731
        (Pretty.string_of (Pretty.block
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   732
         [Pretty.str "Reason:", Pretty.brk 2, pretty_msg]))]
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   733
  in
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   734
    error error_msg
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   735
  end
3cab6f891c2f reorder some steps in the construction to support mutual datatypes
kuncar
parents: 60234
diff changeset
   736
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
   737
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
   738
  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
   739
    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
   740
    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
   741
    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
   742
    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
   743
    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
   744
    val (goal, after_qed) = prepare_lift_def (add_lift_def_code_dt config) var qty rhs par_thms lthy
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   745
    val (goal, after_qed) =
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   746
      case goal of
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   747
        NONE => (goal, K (after_qed Drule.dummy_thm))
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   748
        | SOME prsp_tm =>
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   749
          let
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   750
            val readable_rsp_thm_eq = mk_readable_rsp_thm_eq prsp_tm lthy
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   751
            val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   752
            val readable_rsp_tm_tnames = rename_to_tnames lthy readable_rsp_tm
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   753
        
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   754
            fun after_qed' [[thm]] lthy = 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   755
              let
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   756
                val internal_rsp_thm = Goal.prove lthy [] [] prsp_tm 
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   757
                    (fn {context = ctxt, ...} =>
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   758
                      rtac readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac ctxt [thm] 1)
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   759
              in
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   760
                after_qed internal_rsp_thm lthy
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   761
              end
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   762
          in
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   763
            (SOME readable_rsp_tm_tnames, after_qed')
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   764
          end 
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
   765
  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
   766
    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
   767
  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
   768
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   769
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
   770
  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
   771
    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
   772
    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
   773
    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
   774
       ["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
   775
        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
   776
         [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
   777
        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
   778
         [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
   779
        "",
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   780
        (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
   781
         [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
   782
          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
   783
          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
   784
          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
   785
          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
   786
          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
   787
    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
   788
      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
   789
    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
   790
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   791
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
   792
  (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
   793
    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
   794
    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
   795
      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
   796
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   797
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
   798
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
   799
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   800
(* 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
   801
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
   802
  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
   803
  (((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
   804
    --| @{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
   805
      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
   806
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   807
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
   808
  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
   809
    "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
   810
      (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
   811
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
diff changeset
   812
end