src/HOL/Tools/Lifting/lifting_setup.ML
author wenzelm
Fri, 25 Sep 2015 20:37:59 +0200
changeset 61268 abe08fb15a12
parent 60784 4f590c08fd5d
child 61814 1ca1142e1711
permissions -rw-r--r--
moved remaining display.ML to more_thm.ML;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Lifting/lifting_setup.ML
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     2
    Author:     Ondrej Kuncar
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     3
47352
e0bff2ae939f fix typos
huffman
parents: 47334
diff changeset
     4
Setting up the lifting infrastructure.
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     5
*)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     6
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     7
signature LIFTING_SETUP =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     8
sig
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     9
  exception SETUP_LIFTING_INFR of string
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    10
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    11
  type config = { notes: bool };
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    12
  val default_config: config;
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    13
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    14
  val setup_by_quotient: config -> thm -> thm option -> thm option -> local_theory -> 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    15
    binding * local_theory
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    16
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    17
  val setup_by_typedef_thm: config -> thm -> local_theory -> binding * local_theory
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
    18
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
    19
  val lifting_restore: Lifting_Info.quotient -> Context.generic -> Context.generic
60226
ec23f2a97ba4 publish lifting_forget and lifting_udpate interface
kuncar
parents: 60225
diff changeset
    20
ec23f2a97ba4 publish lifting_forget and lifting_udpate interface
kuncar
parents: 60225
diff changeset
    21
  val lifting_forget: string -> local_theory -> local_theory
ec23f2a97ba4 publish lifting_forget and lifting_udpate interface
kuncar
parents: 60225
diff changeset
    22
  val update_transfer_rules: string -> local_theory -> local_theory
ec23f2a97ba4 publish lifting_forget and lifting_udpate interface
kuncar
parents: 60225
diff changeset
    23
  val pointer_of_bundle_binding: Proof.context -> binding -> string
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
    24
end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    25
47334
4708384e759d fix typo in ML structure name
huffman
parents: 47308
diff changeset
    26
structure Lifting_Setup: LIFTING_SETUP =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    27
struct
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    28
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47623
diff changeset
    29
open Lifting_Util
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    30
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47623
diff changeset
    31
infix 0 MRSL
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    32
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    33
exception SETUP_LIFTING_INFR of string
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    34
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    35
(* Config *)
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    36
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    37
type config = { notes: bool };
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    38
val default_config = { notes = true };
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    39
60239
755e11e2e15d make SML/NJ more happy;
wenzelm
parents: 60231
diff changeset
    40
fun define_crel (config: config) rep_fun lthy =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    41
  let
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47352
diff changeset
    42
    val (qty, rty) = (dest_funT o fastype_of) rep_fun
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47352
diff changeset
    43
    val rep_fun_graph = (HOLogic.eq_const rty) $ Bound 1 $ (rep_fun $ Bound 0)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    44
    val def_term = Abs ("x", rty, Abs ("y", qty, rep_fun_graph))
50175
b27cf0646080 generate correct correspondence relation name
kuncar
parents: 47982
diff changeset
    45
    val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    46
    val crel_name = Binding.prefix_name "cr_" qty_name
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    47
    val (fixed_def_term, lthy) = yield_singleton (Variable.importT_terms) def_term lthy
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    48
    val ((_, (_ , def_thm)), lthy) = if #notes config then
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    49
        Local_Theory.define ((crel_name, NoSyn), ((Thm.def_binding crel_name, []), fixed_def_term)) lthy
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    50
      else 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    51
        Local_Theory.define ((Binding.concealed crel_name, NoSyn), ((Binding.empty, []), fixed_def_term)) lthy
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    52
  in  
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    53
    (def_thm, lthy)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    54
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    55
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    56
fun print_define_pcrel_warning msg = 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    57
  let
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    58
    val warning_msg = cat_lines 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    59
      ["Generation of a parametrized correspondence relation failed.",
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    60
      (Pretty.string_of (Pretty.block
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    61
         [Pretty.str "Reason:", Pretty.brk 2, msg]))]
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    62
  in
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    63
    warning warning_msg
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    64
  end
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    65
60239
755e11e2e15d make SML/NJ more happy;
wenzelm
parents: 60231
diff changeset
    66
fun define_pcrel (config: config) crel lthy =
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    67
  let
50288
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    68
    val (fixed_crel, lthy) = yield_singleton Variable.importT_terms crel lthy
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    69
    val [rty', qty] = (binder_types o fastype_of) fixed_crel
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    70
    val (param_rel, args) = Lifting_Term.generate_parametrized_relator lthy rty'
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    71
    val rty_raw = (domain_type o range_type o fastype_of) param_rel
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    72
    val thy = Proof_Context.theory_of lthy
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    73
    val tyenv_match = Sign.typ_match thy (rty_raw, rty') Vartab.empty
50288
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    74
    val param_rel_subst = Envir.subst_term (tyenv_match,Vartab.empty) param_rel
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    75
    val args_subst = map (Envir.subst_term (tyenv_match,Vartab.empty)) args
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    76
    val lthy = Variable.declare_names fixed_crel lthy
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    77
    val (instT, lthy) = Variable.importT_inst (param_rel_subst :: args_subst) lthy
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    78
    val args_fixed = (map (Term_Subst.instantiate (instT, []))) args_subst
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    79
    val param_rel_fixed = Term_Subst.instantiate (instT, []) param_rel_subst
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    80
    val rty = (domain_type o fastype_of) param_rel_fixed
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    81
    val relcomp_op = Const (@{const_name "relcompp"}, 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    82
          (rty --> rty' --> HOLogic.boolT) --> 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    83
          (rty' --> qty --> HOLogic.boolT) --> 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    84
          rty --> qty --> HOLogic.boolT)
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    85
    val qty_name = (fst o dest_Type) qty
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    86
    val pcrel_name = Binding.prefix_name "pcr_" ((Binding.name o Long_Name.base_name) qty_name)
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    87
    val relator_type = foldr1 (op -->) ((map type_of args_fixed) @ [rty, qty, HOLogic.boolT])
50288
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    88
    val lhs = Library.foldl (op $) ((Free (Binding.name_of pcrel_name, relator_type)), args_fixed)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    89
    val rhs = relcomp_op $ param_rel_fixed $ fixed_crel
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    90
    val definition_term = Logic.mk_equals (lhs, rhs)
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    91
    fun note_def lthy =
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    92
      Specification.definition ((SOME (pcrel_name, SOME relator_type, NoSyn)), 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    93
        ((Binding.empty, []), definition_term)) lthy |>> (snd #> snd);
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    94
    fun raw_def lthy =
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    95
      let
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    96
        val ((_, rhs), prove) = Local_Defs.derived_def lthy true definition_term;
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    97
        val ((_, (_, raw_th)), lthy) = lthy
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    98
          |> Local_Theory.define ((Binding.concealed pcrel_name, NoSyn), ((Binding.empty, []), rhs));
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
    99
        val th = prove lthy raw_th;
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   100
      in
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   101
        (th, lthy)
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   102
      end
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   103
    val (def_thm, lthy) = if #notes config then note_def lthy else raw_def lthy
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   104
  in
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   105
    (SOME def_thm, lthy)
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   106
  end
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   107
  handle Lifting_Term.PARAM_QUOT_THM (_, msg) => (print_define_pcrel_warning msg; (NONE, lthy))
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   108
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   109
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   110
local
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   111
  val eq_OO_meta = mk_meta_eq @{thm eq_OO} 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   112
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   113
  fun print_generate_pcr_cr_eq_error ctxt term = 
56257
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   114
    let
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   115
      val goal = Const (@{const_name HOL.eq}, dummyT) $ term $ Const (@{const_name HOL.eq}, dummyT)
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   116
      val error_msg = cat_lines 
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   117
        ["Generation of a pcr_cr_eq failed.",
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   118
        (Pretty.string_of (Pretty.block
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   119
           [Pretty.str "Reason: Cannot prove this: ", Pretty.brk 2, Syntax.pretty_term ctxt goal])),
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   120
         "Most probably a relator_eq rule for one of the involved types is missing."]
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   121
    in
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   122
      error error_msg
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   123
    end
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   124
in
60239
755e11e2e15d make SML/NJ more happy;
wenzelm
parents: 60231
diff changeset
   125
  fun define_pcr_cr_eq (config: config) lthy pcr_rel_def =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   126
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   127
      val lhs = (Thm.term_of o Thm.lhs_of) pcr_rel_def
60239
755e11e2e15d make SML/NJ more happy;
wenzelm
parents: 60231
diff changeset
   128
      val qty_name =
755e11e2e15d make SML/NJ more happy;
wenzelm
parents: 60231
diff changeset
   129
        (Binding.name o Long_Name.base_name o fst o dest_Type o
755e11e2e15d make SML/NJ more happy;
wenzelm
parents: 60231
diff changeset
   130
          List.last o binder_types o fastype_of) lhs
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   131
      val args = (snd o strip_comb) lhs
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   132
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   133
      fun make_inst var ctxt = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   134
        let 
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   135
          val typ = snd (relation_types (#2 (dest_Var var)))
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   136
          val sort = Type.sort_of_atyp typ
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   137
          val (fresh_var, ctxt') = yield_singleton Variable.invent_types sort ctxt
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   138
          val inst = (#1 (dest_Var var), Thm.cterm_of ctxt' (HOLogic.eq_const (TFree fresh_var)))
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   139
        in (inst, ctxt') end
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   140
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   141
      val orig_lthy = lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   142
      val (args_inst, lthy) = fold_map make_inst args lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   143
      val pcr_cr_eq = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   144
        pcr_rel_def
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   145
        |> infer_instantiate lthy args_inst
52883
0a7c97c76f46 expand equalities in the transfer relation in transfer_prover if the relation doesn't follow the functional structure
kuncar
parents: 51994
diff changeset
   146
        |> Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv 
0a7c97c76f46 expand equalities in the transfer relation in transfer_prover if the relation doesn't follow the functional structure
kuncar
parents: 51994
diff changeset
   147
          (Transfer.bottom_rewr_conv (Transfer.get_relator_eq lthy))))
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   148
  in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   149
    case (Thm.term_of o Thm.rhs_of) pcr_cr_eq of
56257
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   150
      Const (@{const_name "relcompp"}, _) $ Const (@{const_name HOL.eq}, _) $ _ =>
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   151
        let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   152
          val thm = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   153
            pcr_cr_eq
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   154
            |> Conv.fconv_rule (Conv.arg_conv (Conv.rewr_conv eq_OO_meta))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   155
            |> mk_HOL_eq
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   156
            |> singleton (Variable.export lthy orig_lthy)
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   157
          val lthy = (#notes config ? (Local_Theory.note 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   158
              ((Binding.qualified true "pcr_cr_eq" qty_name, []), [thm]) #> snd)) lthy
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   159
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   160
          (thm, lthy)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   161
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   162
      | Const (@{const_name "relcompp"}, _) $ t $ _ => print_generate_pcr_cr_eq_error lthy t
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   163
      | _ => error "generate_pcr_cr_eq: implementation error"
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   164
  end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   165
end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   166
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   167
fun define_code_constr quot_thm lthy =
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   168
  let
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47943
diff changeset
   169
    val abs = quot_thm_abs quot_thm
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   170
  in
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   171
    if is_Const abs then
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   172
      let
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   173
        val (fixed_abs, lthy') = yield_singleton Variable.importT_terms abs lthy
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   174
      in  
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   175
         Local_Theory.background_theory (Code.add_datatype [dest_Const fixed_abs]) lthy'
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   176
      end
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   177
    else
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   178
      lthy
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   179
  end
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   180
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   181
fun define_abs_type quot_thm lthy =
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   182
  if Lifting_Def.can_generate_code_cert quot_thm then
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   183
    let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   184
      val abs_type_thm = quot_thm RS @{thm Quotient_abs_rep}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   185
      val add_abstype_attribute = 
59458
9de8ac92cafa abstract code equation may also be default
haftmann
parents: 59084
diff changeset
   186
          Thm.declaration_attribute (fn thm => Context.mapping (Code.add_abstype_default thm) I)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   187
        val add_abstype_attrib = Attrib.internal (K add_abstype_attribute)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   188
    in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   189
      lthy
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   190
        |> (snd oo Local_Theory.note) ((Binding.empty, [add_abstype_attrib]), [abs_type_thm])
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   191
    end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   192
  else
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   193
    lthy
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   194
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   195
local
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   196
  exception QUOT_ERROR of Pretty.T list
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   197
in
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   198
fun quot_thm_sanity_check ctxt quot_thm =
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   199
  let
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   200
    val _ = 
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   201
      if (Thm.nprems_of quot_thm > 0) then   
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   202
          raise QUOT_ERROR [Pretty.block
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   203
            [Pretty.str "The Quotient theorem has extra assumptions:",
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   204
             Pretty.brk 1,
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60784
diff changeset
   205
             Thm.pretty_thm ctxt quot_thm]]
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   206
      else ()
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   207
    val _ = quot_thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_Quotient
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   208
    handle TERM _ => raise QUOT_ERROR
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   209
          [Pretty.block
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   210
            [Pretty.str "The Quotient theorem is not of the right form:",
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   211
             Pretty.brk 1,
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60784
diff changeset
   212
             Thm.pretty_thm ctxt quot_thm]]
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   213
    val ((_, [quot_thm_fixed]), ctxt') = Variable.importT [quot_thm] ctxt 
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47943
diff changeset
   214
    val (rty, qty) = quot_thm_rty_qty quot_thm_fixed
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   215
    val rty_tfreesT = Term.add_tfree_namesT rty []
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   216
    val qty_tfreesT = Term.add_tfree_namesT qty []
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   217
    val extra_rty_tfrees =
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   218
      case subtract (op =) qty_tfreesT rty_tfreesT of
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   219
        [] => []
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   220
      | extras => [Pretty.block ([Pretty.str "Extra variables in the raw type:",
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   221
                                 Pretty.brk 1] @ 
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   222
                                 ((Pretty.commas o map (Pretty.str o quote)) extras) @
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   223
                                 [Pretty.str "."])]
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   224
    val not_type_constr = 
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   225
      case qty of
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   226
         Type _ => []
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   227
         | _ => [Pretty.block [Pretty.str "The quotient type ",
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   228
                                Pretty.quote (Syntax.pretty_typ ctxt' qty),
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   229
                                Pretty.brk 1,
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   230
                                Pretty.str "is not a type constructor."]]
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   231
    val errs = extra_rty_tfrees @ not_type_constr
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   232
  in
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   233
    if null errs then () else raise QUOT_ERROR errs
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   234
  end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   235
  handle QUOT_ERROR errs => error (cat_lines (["Sanity check of the quotient theorem failed:"] 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   236
                                            @ (map (Pretty.string_of o Pretty.item o single) errs)))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   237
end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   238
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   239
fun lifting_bundle qty_full_name qinfo lthy = 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   240
  let
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   241
    fun qualify suffix defname = Binding.qualified true suffix defname
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   242
    val binding =  qty_full_name |> Long_Name.base_name |> Binding.name |> qualify "lifting"
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   243
    val morphed_binding = Morphism.binding (Local_Theory.target_morphism lthy) binding
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   244
    val bundle_name = Name_Space.full_name (Name_Space.naming_of 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   245
      (Context.Theory (Proof_Context.theory_of lthy))) morphed_binding
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   246
    fun phi_qinfo phi = Lifting_Info.transform_quotient phi qinfo
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   247
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   248
    val thy = Proof_Context.theory_of lthy
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   249
    val dummy_thm = Thm.transfer thy Drule.dummy_thm
59084
f982f3072d79 more robust bundle_name: avoid assumptions about identifier, keywords etc.;
wenzelm
parents: 59083
diff changeset
   250
    val pointer =
f982f3072d79 more robust bundle_name: avoid assumptions about identifier, keywords etc.;
wenzelm
parents: 59083
diff changeset
   251
      Token.explode (Thy_Header.get_keywords thy) Position.none (cartouche bundle_name)
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   252
    val restore_lifting_att = 
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57918
diff changeset
   253
      ([dummy_thm], [Token.src ("Lifting.lifting_restore_internal", Position.none) pointer])
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   254
  in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   255
    lthy 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   256
      |> Local_Theory.declaration {syntax = false, pervasive = true}
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   257
           (fn phi => Lifting_Info.init_restore_data bundle_name (phi_qinfo phi))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   258
      |> Bundle.bundle ((binding, [restore_lifting_att])) []
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   259
      |> pair binding
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   260
  end
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   261
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   262
fun setup_lifting_infr config quot_thm opt_reflp_thm lthy =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   263
  let
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   264
    val _ = quot_thm_sanity_check lthy quot_thm
57663
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 56519
diff changeset
   265
    val (_, qty) = quot_thm_rty_qty quot_thm
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   266
    val (pcrel_def, lthy) = define_pcrel config (quot_thm_crel quot_thm) lthy
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   267
    (**)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   268
    val pcrel_def = Option.map (Morphism.thm (Local_Theory.target_morphism lthy)) pcrel_def
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   269
    (**)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   270
    val (pcr_cr_eq, lthy) = case pcrel_def of
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   271
      SOME pcrel_def => apfst SOME (define_pcr_cr_eq config lthy pcrel_def)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   272
      | NONE => (NONE, lthy)
53219
ca237b9e4542 use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents: 52883
diff changeset
   273
    val pcr_info = case pcrel_def of
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   274
      SOME pcrel_def => SOME { pcrel_def = pcrel_def, pcr_cr_eq = the pcr_cr_eq }
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   275
      | NONE => NONE
53219
ca237b9e4542 use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents: 52883
diff changeset
   276
    val quotients = { quot_thm = quot_thm, pcr_info = pcr_info }
57663
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 56519
diff changeset
   277
    val qty_full_name = (fst o dest_Type) qty
53219
ca237b9e4542 use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents: 52883
diff changeset
   278
    fun quot_info phi = Lifting_Info.transform_quotient phi quotients
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51956
diff changeset
   279
    val reflexivity_rule_attr = Attrib.internal (K Lifting_Info.add_reflexivity_rule_attribute)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   280
    val lthy = case opt_reflp_thm of
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   281
      SOME reflp_thm => lthy
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51956
diff changeset
   282
        |> (snd oo Local_Theory.note) ((Binding.empty, [reflexivity_rule_attr]),
55563
a64d49f49ca3 implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
kuncar
parents: 55487
diff changeset
   283
              [reflp_thm RS @{thm reflp_ge_eq}])
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   284
        |> define_code_constr quot_thm
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   285
      | NONE => lthy
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   286
        |> define_abs_type quot_thm
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   287
  in
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   288
    lthy
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   289
      |> Local_Theory.declaration {syntax = false, pervasive = true}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   290
        (fn phi => Lifting_Info.update_quotients qty_full_name (quot_info phi))
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   291
      |> lifting_bundle qty_full_name quotients
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   292
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   293
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   294
local
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   295
  fun importT_inst_exclude exclude ts ctxt =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   296
    let
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   297
      val tvars = rev (subtract op= exclude (fold Term.add_tvars ts []))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   298
      val (tfrees, ctxt') = Variable.invent_types (map #2 tvars) ctxt
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   299
    in (tvars ~~ map TFree tfrees, ctxt') end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   300
  
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   301
  fun import_inst_exclude exclude ts ctxt =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   302
    let
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   303
      val excludeT = fold (Term.add_tvarsT o snd) exclude []
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   304
      val (instT, ctxt') = importT_inst_exclude excludeT ts ctxt
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   305
      val vars = map (apsnd (Term_Subst.instantiateT instT)) 
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   306
        (rev (subtract op= exclude (fold Term.add_vars ts [])))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   307
      val (xs, ctxt'') = Variable.variant_fixes (map (#1 o #1) vars) ctxt'
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   308
      val inst = vars ~~ map Free (xs ~~ map #2 vars)
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   309
    in ((instT, inst), ctxt'') end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   310
  
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   311
  fun import_terms_exclude exclude ts ctxt =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   312
    let val (inst, ctxt') = import_inst_exclude exclude ts ctxt
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   313
    in (map (Term_Subst.instantiate inst) ts, ctxt') end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   314
in
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   315
  fun reduce_goal not_fix goal tac ctxt =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   316
    let
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   317
      val (fixed_goal, ctxt') = yield_singleton (import_terms_exclude not_fix) goal ctxt
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   318
      val init_goal = Goal.init (Thm.cterm_of ctxt' fixed_goal)
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   319
    in
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   320
      (singleton (Variable.export ctxt' ctxt) o Goal.conclude) (the (SINGLE tac init_goal))
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   321
    end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   322
end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   323
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   324
local 
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   325
  val OO_rules = @{thms left_total_OO left_unique_OO right_total_OO right_unique_OO bi_total_OO
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   326
    bi_unique_OO}
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   327
in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   328
  fun parametrize_class_constraint ctxt pcr_def constraint =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   329
    let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   330
      fun generate_transfer_rule pcr_def constraint goal ctxt =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   331
        let
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   332
          val (fixed_goal, ctxt') = yield_singleton (Variable.import_terms true) goal ctxt
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   333
          val init_goal = Goal.init (Thm.cterm_of ctxt' fixed_goal)
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   334
          val rules = Transfer.get_transfer_raw ctxt'
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   335
          val rules = constraint :: OO_rules @ rules
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59487
diff changeset
   336
          val tac =
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   337
            K (Local_Defs.unfold_tac ctxt' [pcr_def]) THEN' REPEAT_ALL_NEW (resolve_tac ctxt' rules)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   338
        in
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   339
          (singleton (Variable.export ctxt' ctxt) o Goal.conclude) (the (SINGLE (tac 1) init_goal))
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   340
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   341
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   342
      fun make_goal pcr_def constr =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   343
        let 
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   344
          val pred_name =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   345
            (fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o Thm.prop_of) constr
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   346
          val arg = (fst o Logic.dest_equals o Thm.prop_of) pcr_def
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   347
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   348
          HOLogic.mk_Trueprop ((Const (pred_name, (fastype_of arg) --> HOLogic.boolT)) $ arg)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   349
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   350
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   351
      val check_assms =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   352
        let 
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   353
          val right_names = ["right_total", "right_unique", "left_total", "left_unique", "bi_total",
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   354
            "bi_unique"]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   355
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   356
          fun is_right_name name = member op= right_names (Long_Name.base_name name)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   357
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   358
          fun is_trivial_assm (Const (name, _) $ Var (_, _)) = is_right_name name
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   359
            | is_trivial_assm (Const (name, _) $ Free (_, _)) = is_right_name name
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   360
            | is_trivial_assm _ = false
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   361
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   362
          fn thm => 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   363
            let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   364
              val prems = map HOLogic.dest_Trueprop (Thm.prems_of thm)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   365
              val thm_name =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   366
                (Long_Name.base_name o fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o Thm.concl_of) thm
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   367
              val non_trivial_assms = filter_out is_trivial_assm prems
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   368
            in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   369
              if null non_trivial_assms then ()
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   370
              else
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   371
                let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   372
                  val pretty_msg = Pretty.block ([Pretty.str "Non-trivial assumptions in ",
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   373
                    Pretty.str thm_name,
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   374
                    Pretty.str " transfer rule found:",
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   375
                    Pretty.brk 1] @ 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   376
                    ((Pretty.commas o map (Syntax.pretty_term ctxt)) non_trivial_assms) @
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   377
                                       [Pretty.str "."])
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   378
                in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   379
                  warning (Pretty.str_of pretty_msg)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   380
                end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   381
            end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   382
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   383
  
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   384
      val goal = make_goal pcr_def constraint
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   385
      val thm = generate_transfer_rule pcr_def constraint goal ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   386
      val _ = check_assms thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   387
    in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   388
      thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   389
    end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   390
end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   391
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   392
local
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   393
  val id_unfold = (Conv.rewr_conv (mk_meta_eq @{thm id_def}))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   394
in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   395
  fun generate_parametric_id lthy rty id_transfer_rule =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   396
    let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   397
      (* it doesn't raise an exception because it would have already raised it in define_pcrel *)
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   398
      val (quot_thm, _, ctxt') = Lifting_Term.prove_param_quot_thm lthy rty
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   399
      val parametrized_relator =
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   400
        singleton (Variable.export_terms ctxt' lthy) (quot_thm_crel quot_thm)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   401
      val id_transfer = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   402
         @{thm id_transfer}
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   403
        |> Thm.incr_indexes (Term.maxidx_of_term parametrized_relator + 1)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   404
        |> Conv.fconv_rule(HOLogic.Trueprop_conv (Conv.arg_conv id_unfold then_conv Conv.arg1_conv id_unfold))
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   405
      val var = hd (Term.add_vars (Thm.prop_of id_transfer) [])
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   406
      val inst = [(#1 var, Thm.cterm_of lthy parametrized_relator)]
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   407
      val id_par_thm = infer_instantiate lthy inst id_transfer
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   408
    in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   409
      Lifting_Def.generate_parametric_transfer_rule lthy id_transfer_rule id_par_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   410
    end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   411
    handle Lifting_Term.MERGE_TRANSFER_REL msg => 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   412
      let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   413
        val error_msg = cat_lines 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   414
          ["Generation of a parametric transfer rule for the abs. or the rep. function failed.",
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   415
          "A non-parametric version will be used.",
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   416
          (Pretty.string_of (Pretty.block
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   417
             [Pretty.str "Reason:", Pretty.brk 2, msg]))]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   418
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   419
        (warning error_msg; id_transfer_rule)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   420
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   421
end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   422
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   423
local
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   424
  fun rewrite_first_Domainp_arg rewr_thm thm = Conv.fconv_rule (Conv.concl_conv ~1 (HOLogic.Trueprop_conv 
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   425
      (Conv.arg1_conv (Conv.arg_conv (Conv.rewr_conv rewr_thm))))) thm
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   426
  
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   427
  fun fold_Domainp_pcrel pcrel_def thm =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   428
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   429
      val ct =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   430
        thm |> Thm.cprop_of |> Drule.strip_imp_concl
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   431
        |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_arg
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   432
      val pcrel_def = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) pcrel_def
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   433
      val thm = Thm.instantiate (Thm.match (ct, Thm.rhs_of pcrel_def)) thm
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   434
        handle Pattern.MATCH => raise CTERM ("fold_Domainp_pcrel", [ct, Thm.rhs_of pcrel_def])
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   435
    in
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   436
      rewrite_first_Domainp_arg (Thm.symmetric pcrel_def) thm
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   437
    end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   438
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   439
  fun reduce_Domainp ctxt rules thm =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   440
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   441
      val goal = thm |> Thm.prems_of |> hd
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   442
      val var = goal |> HOLogic.dest_Trueprop |> dest_comb |> snd |> dest_Var 
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59487
diff changeset
   443
      val reduced_assm =
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59487
diff changeset
   444
        reduce_goal [var] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   445
    in
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   446
      reduced_assm RS thm
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   447
    end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   448
in
53219
ca237b9e4542 use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents: 52883
diff changeset
   449
  fun parametrize_domain dom_thm (pcr_info : Lifting_Info.pcr) ctxt =
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   450
    let
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   451
      fun reduce_first_assm ctxt rules thm =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   452
        let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   453
          val goal = thm |> Thm.prems_of |> hd
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59487
diff changeset
   454
          val reduced_assm =
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59487
diff changeset
   455
            reduce_goal [] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   456
        in
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   457
          reduced_assm RS thm
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   458
        end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   459
53219
ca237b9e4542 use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents: 52883
diff changeset
   460
      val pcr_cr_met_eq = #pcr_cr_eq pcr_info RS @{thm eq_reflection}
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   461
      val pcr_Domainp_eq = rewrite_first_Domainp_arg (Thm.symmetric pcr_cr_met_eq) dom_thm
53219
ca237b9e4542 use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents: 52883
diff changeset
   462
      val pcrel_def = #pcrel_def pcr_info
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   463
      val pcr_Domainp_par_left_total = 
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   464
        (dom_thm RS @{thm pcr_Domainp_par_left_total})
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   465
          |> fold_Domainp_pcrel pcrel_def
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   466
          |> reduce_first_assm ctxt (Lifting_Info.get_reflexivity_rules ctxt)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   467
      val pcr_Domainp_par = 
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   468
        (dom_thm RS @{thm pcr_Domainp_par})      
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   469
          |> fold_Domainp_pcrel pcrel_def
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   470
          |> reduce_Domainp ctxt (Transfer.get_relator_domain ctxt)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   471
      val pcr_Domainp = 
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   472
        (dom_thm RS @{thm pcr_Domainp})
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   473
          |> fold_Domainp_pcrel pcrel_def
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   474
      val thms =
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   475
        [("domain",                 [pcr_Domainp], @{attributes [transfer_domain_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   476
         ("domain_par",             [pcr_Domainp_par], @{attributes [transfer_domain_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   477
         ("domain_par_left_total",  [pcr_Domainp_par_left_total], @{attributes [transfer_domain_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   478
         ("domain_eq",              [pcr_Domainp_eq], @{attributes [transfer_domain_rule]})]
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   479
    in
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   480
      thms
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   481
    end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   482
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   483
  fun parametrize_total_domain left_total pcrel_def ctxt =
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   484
    let
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   485
      val thm =
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   486
        (left_total RS @{thm pcr_Domainp_total})
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   487
          |> fold_Domainp_pcrel pcrel_def 
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   488
          |> reduce_Domainp ctxt (Transfer.get_relator_domain ctxt)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   489
    in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   490
      [("domain", [thm], @{attributes [transfer_domain_rule]})]
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   491
    end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   492
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   493
end
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   494
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   495
fun get_pcrel_info ctxt qty_full_name =  
53219
ca237b9e4542 use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents: 52883
diff changeset
   496
  #pcr_info (the (Lifting_Info.lookup_quotients ctxt qty_full_name))
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   497
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   498
fun get_Domainp_thm quot_thm =
56519
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   499
   the (get_first (try(curry op RS quot_thm)) [@{thm eq_onp_to_Domainp}, @{thm Quotient_to_Domainp}])
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   500
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   501
fun notes names thms = 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   502
  let
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   503
    val notes =
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   504
        if names then map (fn (name, thms, attrs) => ((name, []), [(thms, attrs)])) thms
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   505
        else map_filter (fn (_, thms, attrs) => if null attrs then NONE 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   506
          else SOME ((Binding.empty, []), [(thms, attrs)])) thms
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   507
  in
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   508
    Local_Theory.notes notes #> snd
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   509
  end
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   510
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   511
fun map_thms map_name map_thm thms = 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   512
  map (fn (name, thms, attr) => (map_name name, map map_thm thms, attr)) thms
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   513
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   514
(*
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   515
  Sets up the Lifting package by a quotient theorem.
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   516
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   517
  quot_thm - a quotient theorem (Quotient R Abs Rep T)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   518
  opt_reflp_thm - a theorem saying that a relation from quot_thm is reflexive
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   519
    (in the form "reflp R")
54335
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 54333
diff changeset
   520
  opt_par_thm - a parametricity theorem for R
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   521
*)
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   522
60239
755e11e2e15d make SML/NJ more happy;
wenzelm
parents: 60231
diff changeset
   523
fun setup_by_quotient (config: config) quot_thm opt_reflp_thm opt_par_thm lthy =
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   524
  let
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   525
    (**)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   526
    val quot_thm = Morphism.thm (Local_Theory.target_morphism lthy) quot_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   527
    (**)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   528
    val (rty, qty) = quot_thm_rty_qty quot_thm
47575
b90cd7016d4f generate abs_induct rules for quotient types
huffman
parents: 47566
diff changeset
   529
    val induct_attr = Attrib.internal (K (Induct.induct_type (fst (dest_Type qty))))
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   530
    val qty_full_name = (fst o dest_Type) qty
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   531
    val qty_name = (Binding.name o Long_Name.base_name) qty_full_name
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   532
    fun qualify suffix = Binding.qualified true suffix qty_name
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   533
    val notes1 = case opt_reflp_thm of
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   534
      SOME reflp_thm =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   535
        let 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   536
          val thms =
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   537
            [("abs_induct",     @{thms Quotient_total_abs_induct}, [induct_attr]),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   538
             ("abs_eq_iff",     @{thms Quotient_total_abs_eq_iff}, []           )]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   539
        in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   540
          map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   541
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   542
      | NONE =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   543
        let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   544
          val thms = 
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   545
            [("abs_induct",     @{thms Quotient_abs_induct},       [induct_attr])]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   546
        in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   547
          map_thms qualify (fn thm => quot_thm RS thm) thms
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   548
        end
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   549
    val dom_thm = get_Domainp_thm quot_thm
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   550
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   551
    fun setup_transfer_rules_nonpar notes =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   552
      let
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   553
        val notes1 =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   554
          case opt_reflp_thm of
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   555
            SOME reflp_thm =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   556
              let 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   557
                val thms =
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   558
                  [("id_abs_transfer",@{thms Quotient_id_abs_transfer}, @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   559
                   ("left_total",     @{thms Quotient_left_total},      @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   560
                   ("bi_total",       @{thms Quotient_bi_total},        @{attributes [transfer_rule]})]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   561
              in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   562
                map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   563
              end
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   564
            | NONE => map_thms qualify I [("domain", [dom_thm], @{attributes [transfer_domain_rule]})]
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   565
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   566
        val notes2 = map_thms qualify (fn thm => quot_thm RS thm)
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   567
          [("rel_eq_transfer", @{thms Quotient_rel_eq_transfer}, @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   568
           ("right_unique",    @{thms Quotient_right_unique},    @{attributes [transfer_rule]}), 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   569
           ("right_total",     @{thms Quotient_right_total},     @{attributes [transfer_rule]})]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   570
      in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   571
         notes2 @ notes1 @ notes
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   572
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   573
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   574
    fun generate_parametric_rel_eq lthy transfer_rule opt_param_thm =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   575
      option_fold transfer_rule (Lifting_Def.generate_parametric_transfer_rule lthy transfer_rule) opt_param_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   576
      handle Lifting_Term.MERGE_TRANSFER_REL msg => 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   577
        let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   578
          val error_msg = cat_lines 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   579
            ["Generation of a parametric transfer rule for the quotient relation failed.",
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   580
            (Pretty.string_of (Pretty.block
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   581
               [Pretty.str "Reason:", Pretty.brk 2, msg]))]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   582
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   583
          error error_msg
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   584
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   585
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   586
    fun setup_transfer_rules_par lthy notes =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   587
      let
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   588
        val pcrel_info = (the (get_pcrel_info lthy qty_full_name))
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   589
        val pcrel_def = #pcrel_def pcrel_info
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   590
        val notes1 =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   591
          case opt_reflp_thm of
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   592
            SOME reflp_thm =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   593
              let
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   594
                val left_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_left_total})
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   595
                val bi_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total})
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   596
                val domain_thms = parametrize_total_domain left_total pcrel_def lthy
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   597
                val id_abs_transfer = generate_parametric_id lthy rty
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   598
                  (Lifting_Term.parametrize_transfer_rule lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   599
                    ([quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer}))
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   600
                val left_total = parametrize_class_constraint lthy pcrel_def left_total
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   601
                val bi_total = parametrize_class_constraint lthy pcrel_def bi_total
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   602
                val thms = 
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   603
                  [("id_abs_transfer", [id_abs_transfer], @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   604
                   ("left_total",      [left_total],      @{attributes [transfer_rule]}),  
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   605
                   ("bi_total",        [bi_total],        @{attributes [transfer_rule]})]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   606
              in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   607
                map_thms qualify I thms @ map_thms qualify I domain_thms
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   608
              end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   609
            | NONE =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   610
              let
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   611
                val thms = parametrize_domain dom_thm pcrel_info lthy
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   612
              in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   613
                map_thms qualify I thms
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   614
              end
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   615
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   616
        val rel_eq_transfer = generate_parametric_rel_eq lthy 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   617
          (Lifting_Term.parametrize_transfer_rule lthy (quot_thm RS @{thm Quotient_rel_eq_transfer}))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   618
            opt_par_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   619
        val right_unique = parametrize_class_constraint lthy pcrel_def 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   620
            (quot_thm RS @{thm Quotient_right_unique})
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   621
        val right_total = parametrize_class_constraint lthy pcrel_def 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   622
            (quot_thm RS @{thm Quotient_right_total})
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   623
        val notes2 = map_thms qualify I
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   624
          [("rel_eq_transfer", [rel_eq_transfer], @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   625
           ("right_unique",    [right_unique],    @{attributes [transfer_rule]}), 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   626
           ("right_total",     [right_total],     @{attributes [transfer_rule]})]      
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   627
      in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   628
        notes2 @ notes1 @ notes
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   629
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   630
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   631
    fun setup_rules lthy = 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   632
      let
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   633
        val thms =  if is_some (get_pcrel_info lthy qty_full_name) 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   634
          then setup_transfer_rules_par lthy notes1 else setup_transfer_rules_nonpar notes1
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   635
      in
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   636
        notes (#notes config) thms lthy
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   637
      end
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   638
  in
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   639
    lthy
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   640
      |> setup_lifting_infr config quot_thm opt_reflp_thm
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   641
      ||> setup_rules
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   642
  end
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   643
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   644
(*
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   645
  Sets up the Lifting package by a typedef theorem.
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   646
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   647
  gen_code - flag if an abstract type given by typedef_thm should be registred 
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   648
    as an abstract type in the code generator
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   649
  typedef_thm - a typedef theorem (type_definition Rep Abs S)
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   650
*)
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   651
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   652
fun setup_by_typedef_thm config typedef_thm lthy =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   653
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   654
    val (_ $ rep_fun $ _ $ typedef_set) = (HOLogic.dest_Trueprop o Thm.prop_of) typedef_thm
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   655
    val (T_def, lthy) = define_crel config rep_fun lthy
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   656
    (**)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   657
    val T_def = Morphism.thm (Local_Theory.target_morphism lthy) T_def
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   658
    (**)    
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   659
    val quot_thm = case typedef_set of
56257
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   660
      Const (@{const_name top}, _) => 
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   661
        [typedef_thm, T_def] MRSL @{thm UNIV_typedef_to_Quotient}
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   662
      | Const (@{const_name "Collect"}, _) $ Abs (_, _, _) => 
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   663
        [typedef_thm, T_def] MRSL @{thm open_typedef_to_Quotient}
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   664
      | _ => 
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   665
        [typedef_thm, T_def] MRSL @{thm typedef_to_Quotient}
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   666
    val (rty, qty) = quot_thm_rty_qty quot_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   667
    val qty_full_name = (fst o dest_Type) qty
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   668
    val qty_name = (Binding.name o Long_Name.base_name) qty_full_name
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   669
    fun qualify suffix = Binding.qualified true suffix qty_name
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   670
    val opt_reflp_thm = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   671
      case typedef_set of
56257
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   672
        Const (@{const_name top}, _) => 
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   673
          SOME ((typedef_thm RS @{thm UNIV_typedef_to_equivp}) RS @{thm equivp_reflp2})
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   674
        | _ =>  NONE
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   675
    val dom_thm = get_Domainp_thm quot_thm
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   676
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   677
    fun setup_transfer_rules_nonpar notes =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   678
      let
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   679
        val notes1 =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   680
          case opt_reflp_thm of
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   681
            SOME reflp_thm =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   682
              let 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   683
                val thms =
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   684
                  [("id_abs_transfer",@{thms Quotient_id_abs_transfer}, @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   685
                   ("left_total",     @{thms Quotient_left_total},      @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   686
                   ("bi_total",       @{thms Quotient_bi_total},        @{attributes [transfer_rule]})]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   687
              in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   688
                map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   689
              end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   690
            | NONE =>
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   691
              map_thms qualify I [("domain", [dom_thm], @{attributes [transfer_domain_rule]})]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   692
        val thms = 
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   693
          [("rep_transfer", @{thms typedef_rep_transfer}, @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   694
           ("left_unique",  @{thms typedef_left_unique},  @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   695
           ("right_unique", @{thms typedef_right_unique}, @{attributes [transfer_rule]}), 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   696
           ("right_total",  @{thms typedef_right_total},  @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   697
           ("bi_unique",    @{thms typedef_bi_unique},    @{attributes [transfer_rule]})]
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   698
      in                                               
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   699
        map_thms qualify (fn thm => [typedef_thm, T_def] MRSL thm) thms @ notes1 @ notes
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   700
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   701
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   702
    fun setup_transfer_rules_par lthy notes =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   703
      let
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   704
        val pcrel_info = (the (get_pcrel_info lthy qty_full_name))
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   705
        val pcrel_def = #pcrel_def pcrel_info
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   706
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   707
        val notes1 =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   708
          case opt_reflp_thm of
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   709
            SOME reflp_thm =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   710
              let
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   711
                val left_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_left_total})
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   712
                val bi_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total})
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   713
                val domain_thms = parametrize_total_domain left_total pcrel_def lthy
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56257
diff changeset
   714
                val left_total = parametrize_class_constraint lthy pcrel_def left_total
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   715
                val bi_total = parametrize_class_constraint lthy pcrel_def bi_total
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   716
                val id_abs_transfer = generate_parametric_id lthy rty
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   717
                  (Lifting_Term.parametrize_transfer_rule lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   718
                    ([quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer}))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   719
                val thms = 
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   720
                  [("left_total",     [left_total],      @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   721
                   ("bi_total",       [bi_total],        @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   722
                   ("id_abs_transfer",[id_abs_transfer], @{attributes [transfer_rule]})]              
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   723
              in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   724
                map_thms qualify I thms @ map_thms qualify I domain_thms
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   725
              end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   726
            | NONE =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   727
              let
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   728
                val thms = parametrize_domain dom_thm pcrel_info lthy
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   729
              in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   730
                map_thms qualify I thms
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   731
              end
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51374
diff changeset
   732
              
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   733
        val notes2 = map_thms qualify (fn thm => generate_parametric_id lthy rty 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   734
            (Lifting_Term.parametrize_transfer_rule lthy ([typedef_thm, T_def] MRSL thm)))
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   735
          [("rep_transfer", @{thms typedef_rep_transfer}, @{attributes [transfer_rule]})];
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   736
        val notes3 =
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   737
          map_thms qualify
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   738
          (fn thm => parametrize_class_constraint lthy pcrel_def ([typedef_thm, T_def] MRSL thm))
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   739
          [("left_unique",  @{thms typedef_left_unique}, @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   740
           ("right_unique", @{thms typedef_right_unique},@{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   741
           ("bi_unique",    @{thms typedef_bi_unique},   @{attributes [transfer_rule]}),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   742
           ("right_total",  @{thms typedef_right_total}, @{attributes [transfer_rule]})]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   743
      in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   744
        notes3 @ notes2 @ notes1 @ notes
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   745
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   746
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   747
    val notes1 = [(Binding.prefix_name "Quotient_" qty_name, [quot_thm], [])]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   748
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   749
    fun setup_rules lthy = 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   750
      let
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   751
        val thms =  if is_some (get_pcrel_info lthy qty_full_name) 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   752
          then setup_transfer_rules_par lthy notes1 else setup_transfer_rules_nonpar notes1
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   753
      in
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   754
        notes (#notes config) thms lthy
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   755
      end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   756
  in
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   757
    lthy
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   758
      |> setup_lifting_infr config quot_thm opt_reflp_thm
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   759
      ||> setup_rules
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   760
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   761
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   762
fun setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm lthy =
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   763
  let 
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   764
    val input_thm = singleton (Attrib.eval_thms lthy) xthm
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   765
    val input_term = (HOLogic.dest_Trueprop o Thm.prop_of) input_thm
47566
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   766
      handle TERM _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported."
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   767
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   768
    fun sanity_check_reflp_thm reflp_thm = 
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   769
      let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   770
        val reflp_tm = (HOLogic.dest_Trueprop o Thm.prop_of) reflp_thm
47566
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   771
          handle TERM _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"."
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   772
      in
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   773
        case reflp_tm of
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   774
          Const (@{const_name reflp}, _) $ _ => ()
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   775
          | _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"."
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   776
      end
55487
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   777
      
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   778
    fun check_qty qty = if not (is_Type qty) 
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   779
          then error "The abstract type must be a type constructor."
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   780
          else ()
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   781
   
47566
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   782
    fun setup_quotient () = 
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   783
      let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   784
        val opt_reflp_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_reflp_xthm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   785
        val _ = if is_some opt_reflp_thm then sanity_check_reflp_thm (the opt_reflp_thm) else ()
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   786
        val opt_par_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_par_xthm
55487
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   787
        val _ = check_qty (snd (quot_thm_rty_qty input_thm))
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   788
      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: 60226
diff changeset
   789
        setup_by_quotient default_config input_thm opt_reflp_thm opt_par_thm lthy |> snd
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   790
      end
47566
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   791
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   792
    fun setup_typedef () = 
55487
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   793
      let
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   794
        val qty = (range_type o fastype_of o hd o get_args 2) input_term
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   795
        val _ = check_qty qty
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   796
      in
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   797
        case opt_reflp_xthm of
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   798
          SOME _ => error "The reflexivity theorem cannot be specified if the type_definition theorem is used."
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   799
          | NONE => (
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   800
            case opt_par_xthm of
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   801
              SOME _ => error "The parametricity theorem cannot be specified if the type_definition theorem is used."
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: 60226
diff changeset
   802
              | NONE => setup_by_typedef_thm default_config input_thm lthy |> snd
55487
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   803
          )
6380313b8ed5 abstract type must be a type constructor; check it
kuncar
parents: 54335
diff changeset
   804
      end
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   805
  in
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   806
    case input_term of
47566
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   807
      (Const (@{const_name Quotient}, _) $ _ $ _ $ _ $ _) => setup_quotient ()
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   808
      | (Const (@{const_name type_definition}, _) $ _ $ _ $ _) => setup_typedef ()
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   809
      | _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported."
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   810
  end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   811
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   812
val _ = 
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59630
diff changeset
   813
  Outer_Syntax.local_theory @{command_keyword setup_lifting}
50214
67fb9a168d10 tuned command descriptions;
wenzelm
parents: 50176
diff changeset
   814
    "setup lifting infrastructure" 
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   815
      (Parse.xthm -- Scan.option Parse.xthm 
58028
e4250d370657 tuned signature -- define some elementary operations earlier;
wenzelm
parents: 58011
diff changeset
   816
      -- Scan.option (@{keyword "parametric"} |-- Parse.!!! Parse.xthm) >> 
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   817
        (fn ((xthm, opt_reflp_xthm), opt_par_xthm) => 
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59458
diff changeset
   818
          setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm))
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   819
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   820
(* restoring lifting infrastructure *)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   821
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   822
local
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   823
  exception PCR_ERROR of Pretty.T list
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   824
in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   825
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   826
fun lifting_restore_sanity_check ctxt (qinfo:Lifting_Info.quotient) =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   827
  let
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   828
    val quot_thm = (#quot_thm qinfo)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   829
    val _ = quot_thm_sanity_check ctxt quot_thm
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   830
    val pcr_info_err =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   831
      (case #pcr_info qinfo of
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   832
        SOME pcr => 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   833
          let
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   834
            val pcrel_def = #pcrel_def pcr
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   835
            val pcr_cr_eq = #pcr_cr_eq pcr
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   836
            val (def_lhs, _) = Logic.dest_equals (Thm.prop_of pcrel_def)
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   837
              handle TERM _ => raise PCR_ERROR [Pretty.block 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   838
                    [Pretty.str "The pcr definiton theorem is not a plain meta equation:",
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   839
                    Pretty.brk 1,
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60784
diff changeset
   840
                    Thm.pretty_thm ctxt pcrel_def]]
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   841
            val pcr_const_def = head_of def_lhs
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   842
            val (eq_lhs, eq_rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of pcr_cr_eq))
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   843
              handle TERM _ => raise PCR_ERROR [Pretty.block 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   844
                    [Pretty.str "The pcr_cr equation theorem is not a plain equation:",
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   845
                    Pretty.brk 1,
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60784
diff changeset
   846
                    Thm.pretty_thm ctxt pcr_cr_eq]]
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   847
            val (pcr_const_eq, eqs) = strip_comb eq_lhs
56257
589fafcc7cb6 more antiquotations;
wenzelm
parents: 56035
diff changeset
   848
            fun is_eq (Const (@{const_name HOL.eq}, _)) = true
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   849
              | is_eq _ = false
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   850
            fun eq_Const (Const (name1, _)) (Const (name2, _)) = (name1 = name2)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   851
              | eq_Const _ _ = false
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   852
            val all_eqs = if not (forall is_eq eqs) then 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   853
              [Pretty.block
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   854
                    [Pretty.str "Arguments of the lhs of the pcr_cr equation theorem are not only equalities:",
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   855
                    Pretty.brk 1,
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60784
diff changeset
   856
                    Thm.pretty_thm ctxt pcr_cr_eq]]
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   857
              else []
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   858
            val pcr_consts_not_equal = if not (eq_Const pcr_const_def pcr_const_eq) then
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   859
              [Pretty.block
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   860
                    [Pretty.str "Parametrized correspondence relation constants in pcr_def and pcr_cr_eq are not equal:",
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   861
                    Pretty.brk 1,
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   862
                    Syntax.pretty_term ctxt pcr_const_def,
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   863
                    Pretty.brk 1,
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   864
                    Pretty.str "vs.",
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   865
                    Pretty.brk 1,
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   866
                    Syntax.pretty_term ctxt pcr_const_eq]]
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   867
              else []
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   868
            val crel = quot_thm_crel quot_thm
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   869
            val cr_consts_not_equal = if not (eq_Const crel eq_rhs) then
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   870
              [Pretty.block
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   871
                    [Pretty.str "Correspondence relation constants in the Quotient theorem and pcr_cr_eq are not equal:",
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   872
                    Pretty.brk 1,
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   873
                    Syntax.pretty_term ctxt crel,
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   874
                    Pretty.brk 1,
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   875
                    Pretty.str "vs.",
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   876
                    Pretty.brk 1,
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   877
                    Syntax.pretty_term ctxt eq_rhs]]
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   878
              else []
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   879
          in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   880
            all_eqs @ pcr_consts_not_equal @ cr_consts_not_equal
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   881
          end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   882
        | NONE => [])
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   883
    val errs = pcr_info_err
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   884
  in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   885
    if null errs then () else raise PCR_ERROR errs
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   886
  end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   887
  handle PCR_ERROR errs => error (cat_lines (["Sanity check failed:"] 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   888
                                            @ (map (Pretty.string_of o Pretty.item o single) errs)))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   889
end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   890
54335
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 54333
diff changeset
   891
(*
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 54333
diff changeset
   892
  Registers the data in qinfo in the Lifting infrastructure.
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 54333
diff changeset
   893
*)
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 54333
diff changeset
   894
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   895
fun lifting_restore qinfo ctxt =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   896
  let
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   897
    val _ = lifting_restore_sanity_check (Context.proof_of ctxt) qinfo
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   898
    val (_, qty) = quot_thm_rty_qty (#quot_thm qinfo)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   899
    val qty_full_name = (fst o dest_Type) qty
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   900
    val stored_qinfo = Lifting_Info.lookup_quotients (Context.proof_of ctxt) qty_full_name
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   901
  in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   902
    if is_some (stored_qinfo) andalso not (Lifting_Info.quotient_eq (qinfo, (the stored_qinfo)))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   903
      then error (Pretty.string_of 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   904
        (Pretty.block
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   905
          [Pretty.str "Lifting is already setup for the type",
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   906
           Pretty.brk 1,
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   907
           Pretty.quote (Syntax.pretty_typ (Context.proof_of ctxt) qty)]))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   908
      else Lifting_Info.update_quotients qty_full_name qinfo ctxt
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   909
  end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   910
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   911
val parse_opt_pcr =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   912
  Scan.optional (Attrib.thm -- Attrib.thm >> 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   913
    (fn (pcrel_def, pcr_cr_eq) => SOME {pcrel_def = pcrel_def, pcr_cr_eq = pcr_cr_eq})) NONE
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   914
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   915
val lifting_restore_attribute_setup =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   916
  Attrib.setup @{binding lifting_restore}
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   917
    ((Attrib.thm -- parse_opt_pcr) >>
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   918
      (fn (quot_thm, opt_pcr) =>
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   919
        let val qinfo = { quot_thm = quot_thm, pcr_info = opt_pcr}
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   920
        in Thm.declaration_attribute (K (lifting_restore qinfo)) end))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   921
    "restoring lifting infrastructure"
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   922
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   923
val _ = Theory.setup lifting_restore_attribute_setup 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   924
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   925
fun lifting_restore_internal bundle_name ctxt = 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   926
  let 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   927
    val restore_info = Lifting_Info.lookup_restore_data (Context.proof_of ctxt) bundle_name
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   928
  in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   929
    case restore_info of
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   930
      SOME restore_info =>
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   931
        ctxt 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   932
        |> lifting_restore (#quotient restore_info)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   933
        |> fold_rev Transfer.transfer_raw_add (Item_Net.content (#transfer_rules restore_info))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   934
      | NONE => ctxt
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   935
  end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   936
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   937
val lifting_restore_internal_attribute_setup =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   938
  Attrib.setup @{binding lifting_restore_internal}
59084
f982f3072d79 more robust bundle_name: avoid assumptions about identifier, keywords etc.;
wenzelm
parents: 59083
diff changeset
   939
    (Scan.lift Parse.cartouche >>
f982f3072d79 more robust bundle_name: avoid assumptions about identifier, keywords etc.;
wenzelm
parents: 59083
diff changeset
   940
      (fn name => Thm.declaration_attribute (K (lifting_restore_internal name))))
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   941
    "restoring lifting infrastructure; internal attribute; not meant to be used directly by regular users"
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   942
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   943
val _ = Theory.setup lifting_restore_internal_attribute_setup 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   944
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   945
(* lifting_forget *)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   946
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   947
val monotonicity_names = [@{const_name right_unique}, @{const_name left_unique}, @{const_name right_total},
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   948
  @{const_name left_total}, @{const_name bi_unique}, @{const_name bi_total}]
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   949
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   950
fun fold_transfer_rel f (Const (@{const_name "Transfer.Rel"}, _) $ rel $ _ $ _) = f rel
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   951
  | fold_transfer_rel f (Const (@{const_name "HOL.eq"}, _) $ 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   952
    (Const (@{const_name Domainp}, _) $ rel) $ _) = f rel
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   953
  | fold_transfer_rel f (Const (name, _) $ rel) = 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   954
    if member op= monotonicity_names name then f rel else f @{term undefined}
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   955
  | fold_transfer_rel f _ = f @{term undefined}
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   956
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   957
fun filter_transfer_rules_by_rel transfer_rel transfer_rules =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   958
  let
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   959
    val transfer_rel_name = transfer_rel |> dest_Const |> fst;
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   960
    fun has_transfer_rel thm = 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   961
      let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   962
        val concl = thm |> Thm.concl_of |> HOLogic.dest_Trueprop
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   963
      in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   964
        member op= (fold_transfer_rel (fn tm => Term.add_const_names tm []) concl) transfer_rel_name
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   965
      end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   966
      handle TERM _ => false
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   967
  in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   968
    filter has_transfer_rel transfer_rules
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   969
  end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   970
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   971
type restore_data = {quotient : Lifting_Info.quotient, transfer_rules: thm Item_Net.T}
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   972
53754
124bb918f45f make SML/NJ happy
kuncar
parents: 53651
diff changeset
   973
fun get_transfer_rel (qinfo : Lifting_Info.quotient) =
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   974
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   975
    fun get_pcrel pcr_def = pcr_def |> Thm.concl_of |> Logic.dest_equals |> fst |> head_of
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   976
  in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   977
    if is_some (#pcr_info qinfo) 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   978
      then get_pcrel (#pcrel_def (the (#pcr_info qinfo)))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   979
      else quot_thm_crel (#quot_thm qinfo)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   980
  end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   981
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   982
fun pointer_of_bundle_name bundle_name ctxt =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   983
  let
56026
893fe12639bc tuned signature -- prefer Name_Space.get with its builtin error;
wenzelm
parents: 55563
diff changeset
   984
    val bundle = Bundle.get_bundle_cmd ctxt bundle_name
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   985
  in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   986
    case bundle of
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   987
      [(_, [arg_src])] => 
56035
745f568837f1 proper Args.syntax for slightly odd bundle trickery;
wenzelm
parents: 56029
diff changeset
   988
        let
59084
f982f3072d79 more robust bundle_name: avoid assumptions about identifier, keywords etc.;
wenzelm
parents: 59083
diff changeset
   989
          val (name, _) = Token.syntax (Scan.lift Parse.cartouche) arg_src ctxt
56035
745f568837f1 proper Args.syntax for slightly odd bundle trickery;
wenzelm
parents: 56029
diff changeset
   990
            handle ERROR _ => error "The provided bundle is not a lifting bundle."
745f568837f1 proper Args.syntax for slightly odd bundle trickery;
wenzelm
parents: 56029
diff changeset
   991
        in name end
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   992
      | _ => error "The provided bundle is not a lifting bundle."
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   993
  end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   994
60226
ec23f2a97ba4 publish lifting_forget and lifting_udpate interface
kuncar
parents: 60225
diff changeset
   995
fun pointer_of_bundle_binding ctxt binding = Name_Space.full_name (Name_Space.naming_of 
ec23f2a97ba4 publish lifting_forget and lifting_udpate interface
kuncar
parents: 60225
diff changeset
   996
      (Context.Theory (Proof_Context.theory_of ctxt))) binding
ec23f2a97ba4 publish lifting_forget and lifting_udpate interface
kuncar
parents: 60225
diff changeset
   997
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   998
fun lifting_forget pointer lthy =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
   999
  let
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1000
    fun get_transfer_rules_to_delete qinfo ctxt =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1001
      let
53754
124bb918f45f make SML/NJ happy
kuncar
parents: 53651
diff changeset
  1002
        val transfer_rel = get_transfer_rel qinfo
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1003
      in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1004
         filter_transfer_rules_by_rel transfer_rel (Transfer.get_transfer_raw ctxt)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1005
      end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1006
  in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1007
    case Lifting_Info.lookup_restore_data lthy pointer of
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1008
      SOME restore_info =>
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1009
        let
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1010
          val qinfo = #quotient restore_info
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1011
          val quot_thm = #quot_thm qinfo
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1012
          val transfer_rules = get_transfer_rules_to_delete qinfo lthy
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1013
        in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1014
          Local_Theory.declaration {syntax = false, pervasive = true}
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1015
            (K (fold (Transfer.transfer_raw_del) transfer_rules #> Lifting_Info.delete_quotients quot_thm))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1016
            lthy
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1017
        end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1018
      | NONE => error "The lifting bundle refers to non-existent restore data."
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1019
    end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1020
    
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1021
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1022
fun lifting_forget_cmd bundle_name lthy = 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1023
  lifting_forget (pointer_of_bundle_name bundle_name lthy) lthy
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1024
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1025
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1026
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59630
diff changeset
  1027
  Outer_Syntax.local_theory @{command_keyword lifting_forget} 
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1028
    "unsetup Lifting and Transfer for the given lifting bundle"
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1029
    (Parse.position Parse.xname >> (lifting_forget_cmd))
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1030
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1031
(* lifting_update *)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1032
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1033
fun update_transfer_rules pointer lthy =
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1034
  let
53754
124bb918f45f make SML/NJ happy
kuncar
parents: 53651
diff changeset
  1035
    fun new_transfer_rules ({ quotient = qinfo, ... }:Lifting_Info.restore_data) lthy =
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1036
      let
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1037
        val transfer_rel = get_transfer_rel qinfo
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1038
        val transfer_rules = filter_transfer_rules_by_rel transfer_rel (Transfer.get_transfer_raw lthy)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1039
      in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1040
        fn phi => fold_rev 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1041
          (Item_Net.update o Morphism.thm phi) transfer_rules Thm.full_rules
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1042
      end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1043
  in
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1044
    case Lifting_Info.lookup_restore_data lthy pointer of
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1045
      SOME refresh_data => 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1046
        Local_Theory.declaration {syntax = false, pervasive = true}
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1047
          (fn phi => Lifting_Info.add_transfer_rules_in_restore_data pointer 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1048
            (new_transfer_rules refresh_data lthy phi)) lthy
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1049
      | NONE => error "The lifting bundle refers to non-existent restore data."
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1050
  end
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1051
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1052
fun lifting_update_cmd bundle_name lthy = 
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1053
  update_transfer_rules (pointer_of_bundle_name bundle_name lthy) lthy
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1054
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1055
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59630
diff changeset
  1056
  Outer_Syntax.local_theory @{command_keyword lifting_update}
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1057
    "add newly introduced transfer rules to a bundle storing the state of Lifting and Transfer"
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1058
    (Parse.position Parse.xname >> lifting_update_cmd)
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1059
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53219
diff changeset
  1060
end