src/HOL/Tools/Lifting/lifting_setup.ML
author wenzelm
Thu, 04 Apr 2013 18:20:00 +0200
changeset 51618 a3577cd80c41
parent 51374 84d01fd733cf
child 51956 a4d81cdebf8b
permissions -rw-r--r--
tuned signature -- avoid intrusion of slightly odd Swing structures into pure Markup_Tree;
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
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    11
  val setup_by_quotient: bool -> thm -> thm option -> thm option -> local_theory -> local_theory
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    12
47566
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
    13
  val setup_by_typedef_thm: bool -> thm -> local_theory -> local_theory
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    14
end;
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    15
47334
4708384e759d fix typo in ML structure name
huffman
parents: 47308
diff changeset
    16
structure Lifting_Setup: LIFTING_SETUP =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    17
struct
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    18
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47623
diff changeset
    19
open Lifting_Util
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    20
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47623
diff changeset
    21
infix 0 MRSL
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    22
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    23
exception SETUP_LIFTING_INFR of string
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    24
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    25
fun define_crel rep_fun lthy =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    26
  let
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47352
diff changeset
    27
    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
    28
    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
    29
    val def_term = Abs ("x", rty, Abs ("y", qty, rep_fun_graph))
50175
b27cf0646080 generate correct correspondence relation name
kuncar
parents: 47982
diff changeset
    30
    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
    31
    val crel_name = Binding.prefix_name "cr_" qty_name
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    32
    val (fixed_def_term, lthy') = yield_singleton (Variable.importT_terms) def_term lthy
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    33
    val ((_, (_ , def_thm)), lthy'') =
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    34
      Local_Theory.define ((crel_name, NoSyn), ((Thm.def_binding crel_name, []), fixed_def_term)) lthy'
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    35
  in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    36
    (def_thm, lthy'')
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    37
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    38
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    39
fun print_define_pcrel_warning msg = 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    40
  let
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    41
    val warning_msg = cat_lines 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    42
      ["Generation of a parametrized correspondence relation failed.",
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    43
      (Pretty.string_of (Pretty.block
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    44
         [Pretty.str "Reason:", Pretty.brk 2, msg]))]
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    45
  in
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    46
    warning warning_msg
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    47
  end
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    48
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    49
fun define_pcrel crel lthy =
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    50
  let
50288
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    51
    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
    52
    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
    53
    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
    54
    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
    55
    val thy = Proof_Context.theory_of lthy
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    56
    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
    57
    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
    58
    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
    59
    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
    60
    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
    61
    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
    62
    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
    63
    val rty = (domain_type o fastype_of) param_rel_fixed
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    64
    val relcomp_op = Const (@{const_name "relcompp"}, 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    65
          (rty --> rty' --> HOLogic.boolT) --> 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    66
          (rty' --> qty --> HOLogic.boolT) --> 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    67
          rty --> qty --> HOLogic.boolT)
50288
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    68
    val relator_type = foldr1 (op -->) ((map type_of args_fixed) @ [rty, qty, HOLogic.boolT])
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    69
    val qty_name = (fst o dest_Type) qty
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    70
    val pcrel_name = Binding.prefix_name "pcr_" ((Binding.name o Long_Name.base_name) qty_name)
50288
986598b0efd1 parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents: 50231
diff changeset
    71
    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
    72
    val rhs = relcomp_op $ param_rel_fixed $ fixed_crel
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    73
    val definition_term = Logic.mk_equals (lhs, rhs)
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    74
    val ((_, (_, def_thm)), lthy) = Specification.definition ((SOME (pcrel_name, SOME relator_type, NoSyn)), 
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    75
      ((Binding.empty, []), definition_term)) lthy
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    76
  in
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    77
    (SOME def_thm, lthy)
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    78
  end
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
    79
  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
    80
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    81
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    82
local
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    83
  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
    84
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    85
  fun print_generate_pcr_cr_eq_error ctxt term = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    86
  let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    87
    val goal = (Const ("HOL.eq", dummyT)) $ term $ Const ("HOL.eq", dummyT)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    88
    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
    89
      ["Generation of a pcr_cr_eq failed.",
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    90
      (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
    91
         [Pretty.str "Reason: Cannot prove this: ", Pretty.brk 2, Syntax.pretty_term ctxt goal])),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    92
       "Most probably a relator_eq rule for one of the involved types is missing."]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    93
  in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    94
    error error_msg
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    95
  end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    96
in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    97
  fun define_pcr_cr_eq lthy pcr_rel_def =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    98
    let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
    99
      val lhs = (term_of o Thm.lhs_of) pcr_rel_def
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   100
      val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type o List.last o binder_types o fastype_of) lhs
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   101
      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
   102
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   103
      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
   104
        let 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   105
          val typ = (snd o relation_types o snd o dest_Var) var
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   106
          val sort = Type.sort_of_atyp typ
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   107
          val (fresh_var, ctxt) = yield_singleton Variable.invent_types sort ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   108
          val thy = Proof_Context.theory_of ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   109
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   110
          ((cterm_of thy var, cterm_of thy (HOLogic.eq_const (TFree fresh_var))), ctxt)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   111
        end
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
      val orig_lthy = lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   114
      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
   115
      val pcr_cr_eq = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   116
        pcr_rel_def
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   117
        |> Drule.cterm_instantiate args_inst    
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   118
        |> Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (bottom_rewr_conv (Transfer.get_relator_eq lthy))))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   119
  in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   120
    case (term_of o Thm.rhs_of) pcr_cr_eq of
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   121
      Const (@{const_name "relcompp"}, _) $ Const ("HOL.eq", _) $ _ => 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   122
        let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   123
          val thm = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   124
            pcr_cr_eq
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   125
            |> 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
   126
            |> mk_HOL_eq
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   127
            |> singleton (Variable.export lthy orig_lthy)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   128
          val ((_, [thm]), lthy) = Local_Theory.note ((Binding.qualified true "pcr_cr_eq" qty_name, []), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   129
            [thm]) lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   130
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   131
          (thm, lthy)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   132
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   133
      | 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
   134
      | _ => 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
   135
  end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   136
end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   137
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
   138
fun define_code_constr gen_code quot_thm 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
   139
  let
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47943
diff changeset
   140
    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
   141
  in
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   142
    if gen_code andalso 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
   143
      let
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   144
        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
   145
      in  
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   146
         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
   147
      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
   148
    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
   149
      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
   150
  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
   151
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
   152
fun define_abs_type gen_code quot_thm 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
   153
  if gen_code andalso Lifting_Def.can_generate_code_cert quot_thm then
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   154
    let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   155
      val abs_type_thm = quot_thm RS @{thm Quotient_abs_rep}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   156
      val add_abstype_attribute = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   157
          Thm.declaration_attribute (fn thm => Context.mapping (Code.add_abstype thm) I)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   158
        val add_abstype_attrib = Attrib.internal (K add_abstype_attribute)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   159
    in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   160
      lthy
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   161
        |> (snd oo Local_Theory.note) ((Binding.empty, [add_abstype_attrib]), [abs_type_thm])
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   162
    end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   163
  else
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   164
    lthy
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   165
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   166
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
   167
  let
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   168
    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
   169
    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
   170
    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
   171
    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
   172
    val extra_rty_tfrees =
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   173
      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
   174
        [] => []
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   175
      | 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
   176
                                 Pretty.brk 1] @ 
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   177
                                 ((Pretty.commas o map (Pretty.str o quote)) extras) @
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   178
                                 [Pretty.str "."])]
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   179
    val not_type_constr = 
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   180
      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
   181
         Type _ => []
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   182
         | _ => [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
   183
                                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
   184
                                Pretty.brk 1,
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   185
                                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
   186
    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
   187
  in
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   188
    if null errs then () else error (cat_lines (["Sanity check of the quotient theorem failed:",""] 
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   189
                                                @ (map Pretty.string_of errs)))
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   190
  end
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   191
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   192
fun setup_lifting_infr gen_code quot_thm opt_reflp_thm lthy =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   193
  let
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47361
diff changeset
   194
    val _ = quot_thm_sanity_check lthy quot_thm
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47943
diff changeset
   195
    val (_, qtyp) = quot_thm_rty_qty quot_thm
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   196
    val (pcrel_def, lthy) = define_pcrel (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
   197
    (**)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   198
    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
   199
    (**)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   200
    val (pcr_cr_eq, lthy) = case pcrel_def of
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   201
      SOME pcrel_def => apfst SOME (define_pcr_cr_eq lthy pcrel_def)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   202
      | NONE => (NONE, lthy)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   203
    val pcrel_info = case pcrel_def of
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   204
      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
   205
      | NONE => NONE
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   206
    val quotients = { quot_thm = quot_thm, pcrel_info = pcrel_info }
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   207
    val qty_full_name = (fst o dest_Type) qtyp  
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   208
    fun quot_info phi = Lifting_Info.transform_quotients phi quotients
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   209
    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
   210
      SOME reflp_thm => lthy
47982
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47951
diff changeset
   211
        |> (snd oo Local_Theory.note) ((Binding.empty, [Lifting_Info.add_reflexivity_rule_attrib]),
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
   212
              [reflp_thm])
47982
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47951
diff changeset
   213
        |> (snd oo Local_Theory.note) ((Binding.empty, [Lifting_Info.add_reflexivity_rule_attrib]),
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47951
diff changeset
   214
              [[quot_thm, reflp_thm] MRSL @{thm Quotient_to_left_total}])
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
   215
        |> define_code_constr gen_code quot_thm
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
   216
      | NONE => 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
   217
        |> define_abs_type gen_code quot_thm
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   218
  in
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   219
    lthy
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   220
      |> Local_Theory.declaration {syntax = false, pervasive = true}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   221
        (fn phi => Lifting_Info.update_quotients qty_full_name (quot_info phi))
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   222
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   223
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   224
local 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   225
  val OO_rules = [@{thm bi_total_OO}, @{thm bi_unique_OO}, @{thm right_total_OO}, @{thm right_unique_OO}]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   226
in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   227
  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
   228
    let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   229
      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
   230
        let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   231
          val thy = Proof_Context.theory_of ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   232
          val orig_ctxt = ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   233
          val (fixed_goal, ctxt) = yield_singleton (Variable.import_terms true) goal ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   234
          val init_goal = Goal.init (cterm_of thy fixed_goal)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   235
          val rules = Transfer.get_transfer_raw ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   236
          val rules = constraint :: OO_rules @ rules
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   237
          val tac = K (Local_Defs.unfold_tac ctxt [pcr_def]) THEN' REPEAT_ALL_NEW (resolve_tac rules)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   238
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   239
          (singleton (Variable.export ctxt orig_ctxt) o Goal.conclude) (the (SINGLE (tac 1) init_goal))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   240
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   241
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   242
      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
   243
        let 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   244
          val pred_name = (fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o prop_of) constr
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   245
          val arg = (fst o Logic.dest_equals o prop_of) pcr_def
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   246
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   247
          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
   248
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   249
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   250
      val check_assms =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   251
        let 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   252
          val right_names = ["bi_total", "bi_unique", "right_total", "right_unique"]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   253
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   254
          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
   255
      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   256
          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
   257
            | 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
   258
            | is_trivial_assm _ = false
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   259
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   260
          fn thm => 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   261
            let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   262
              val prems = map HOLogic.dest_Trueprop (prems_of thm)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   263
              val thm_name = (Long_Name.base_name o fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o concl_of) thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   264
              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
   265
            in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   266
              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
   267
              else
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   268
                let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   269
                  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
   270
                    Pretty.str thm_name,
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   271
                    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
   272
                    Pretty.brk 1] @ 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   273
                    ((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
   274
                                       [Pretty.str "."])
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   275
                in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   276
                  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
   277
                end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   278
            end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   279
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   280
  
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   281
      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
   282
      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
   283
      val _ = check_assms thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   284
    in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   285
      thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   286
    end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   287
end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   288
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   289
local
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   290
  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
   291
in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   292
  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
   293
    let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   294
      val orig_lthy = lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   295
      (* it doesn't raise an exception because it would have already raised it in define_pcrel *)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   296
      val (quot_thm, _, lthy) = Lifting_Term.prove_param_quot_thm lthy rty
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   297
      val parametrized_relator = singleton (Variable.export_terms lthy orig_lthy) (quot_thm_crel quot_thm);
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   298
      val lthy = orig_lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   299
      val id_transfer = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   300
         @{thm id_transfer}
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   301
        |> 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
   302
        |> Conv.fconv_rule(HOLogic.Trueprop_conv (Conv.arg_conv id_unfold then_conv Conv.arg1_conv id_unfold))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   303
      val var = Var (hd (Term.add_vars (prop_of id_transfer) []));
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   304
      val thy = Proof_Context.theory_of lthy;
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   305
      val inst = [(cterm_of thy var, cterm_of thy parametrized_relator)]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   306
      val id_par_thm = Drule.cterm_instantiate inst id_transfer;
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   307
    in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   308
      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
   309
    end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   310
    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
   311
      let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   312
        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
   313
          ["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
   314
          "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
   315
          (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
   316
             [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
   317
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   318
        (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
   319
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   320
end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   321
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   322
fun parametrize_quantifier lthy quant_transfer_rule =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   323
  Lifting_Term.parametrize_transfer_rule lthy quant_transfer_rule
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   324
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   325
fun get_pcrel_info ctxt qty_full_name =  
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   326
  #pcrel_info (the (Lifting_Info.lookup_quotients ctxt qty_full_name))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   327
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   328
(*
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   329
  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
   330
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
   331
  gen_code - flag if an abstract type given by quot_thm should be registred 
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   332
    as an abstract type in the code generator
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   333
  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
   334
  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
   335
    (in the form "reflp R")
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   336
*)
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   337
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   338
fun setup_by_quotient gen_code 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
   339
  let
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   340
    (**)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   341
    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
   342
    (**)
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   343
    val transfer_attr = Attrib.internal (K Transfer.transfer_add)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   344
    val (rty, qty) = quot_thm_rty_qty quot_thm
47575
b90cd7016d4f generate abs_induct rules for quotient types
huffman
parents: 47566
diff changeset
   345
    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
   346
    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
   347
    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
   348
    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
   349
    val lthy = 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
   350
      SOME reflp_thm =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   351
        let 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   352
          val thms =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   353
            [("abs_induct",     @{thm Quotient_total_abs_induct}, [induct_attr]),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   354
             ("abs_eq_iff",     @{thm Quotient_total_abs_eq_iff}, []           )]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   355
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   356
          lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   357
            |> fold (fn (name, thm, attr) => (snd oo Local_Theory.note) ((qualify name, attr), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   358
              [[quot_thm, reflp_thm] MRSL thm])) thms
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   359
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   360
      | NONE =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   361
        let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   362
          val thms = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   363
            [("abs_induct",     @{thm Quotient_abs_induct},       [induct_attr])]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   364
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   365
          fold (fn (name, thm, attr) => (snd oo Local_Theory.note) ((qualify name, attr), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   366
            [quot_thm RS thm])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   367
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   368
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   369
    fun setup_transfer_rules_nonpar lthy =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   370
      let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   371
        val lthy =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   372
          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
   373
            SOME reflp_thm =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   374
              let 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   375
                val thms =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   376
                  [("id_abs_transfer",@{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
   377
                   ("bi_total",       @{thm Quotient_bi_total}       )]
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
                fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   380
                    [[quot_thm, reflp_thm] MRSL thm])) thms lthy
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
            | NONE =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   383
              let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   384
                val thms = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   385
                  [("All_transfer",   @{thm Quotient_All_transfer}   ),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   386
                   ("Ex_transfer",    @{thm Quotient_Ex_transfer}    ),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   387
                   ("forall_transfer",@{thm Quotient_forall_transfer})]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   388
              in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   389
                fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   390
                  [quot_thm RS thm])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   391
              end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   392
        val thms = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   393
          [("rel_eq_transfer", @{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
   394
           ("right_unique",    @{thm Quotient_right_unique}   ), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   395
           ("right_total",     @{thm Quotient_right_total}    )]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   396
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   397
        fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   398
          [quot_thm RS thm])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   399
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   400
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   401
    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
   402
      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
   403
      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
   404
        let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   405
          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
   406
            ["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
   407
            (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
   408
               [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
   409
        in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   410
          error error_msg
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   411
        end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   412
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   413
    fun setup_transfer_rules_par lthy =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   414
      let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   415
        val pcrel_def = #pcrel_def (the (get_pcrel_info lthy qty_full_name))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   416
        val lthy =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   417
          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
   418
            SOME reflp_thm =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   419
              let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   420
                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
   421
                  (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
   422
                    ([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
   423
                val bi_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
   424
                    ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total})
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   425
                val thms = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   426
                  [("id_abs_transfer",id_abs_transfer),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   427
                   ("bi_total",       bi_total       )]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   428
              in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   429
                fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   430
                    [thm])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   431
              end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   432
            | NONE =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   433
              let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   434
                val thms = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   435
                  [("All_transfer",   @{thm Quotient_All_transfer}   ),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   436
                   ("Ex_transfer",    @{thm Quotient_Ex_transfer}    ),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   437
                   ("forall_transfer",@{thm Quotient_forall_transfer})]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   438
              in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   439
                fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   440
                  [parametrize_quantifier lthy (quot_thm RS thm)])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   441
              end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   442
        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
   443
          (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
   444
            opt_par_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   445
        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
   446
            (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
   447
        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
   448
            (quot_thm RS @{thm Quotient_right_total})
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   449
        val thms = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   450
          [("rel_eq_transfer", rel_eq_transfer),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   451
           ("right_unique",    right_unique   ), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   452
           ("right_total",     right_total    )]      
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   453
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   454
        fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   455
          [thm])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   456
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   457
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   458
    fun setup_transfer_rules lthy = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   459
      if is_some (get_pcrel_info lthy qty_full_name) then setup_transfer_rules_par lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   460
                                                     else setup_transfer_rules_nonpar lthy
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   461
  in
50227
01d545993e8c generate a parameterized correspondence relation
kuncar
parents: 50176
diff changeset
   462
    lthy
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   463
      |> setup_lifting_infr gen_code quot_thm opt_reflp_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   464
      |> setup_transfer_rules
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   465
  end
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   466
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   467
(*
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   468
  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
   469
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
   470
  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
   471
    as an abstract type in the code generator
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   472
  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
   473
*)
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47779
diff changeset
   474
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
   475
fun setup_by_typedef_thm gen_code typedef_thm lthy =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   476
  let
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   477
    val transfer_attr = Attrib.internal (K Transfer.transfer_add)
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   478
    val (_ $ rep_fun $ _ $ typedef_set) = (HOLogic.dest_Trueprop o prop_of) typedef_thm
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   479
    val (T_def, lthy) = define_crel rep_fun lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   480
    (**)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   481
    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
   482
    (**)    
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   483
    val quot_thm = case typedef_set of
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   484
      Const ("Orderings.top_class.top", _) => 
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   485
        [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
   486
      | Const (@{const_name "Collect"}, _) $ Abs (_, _, _) => 
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   487
        [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
   488
      | _ => 
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47535
diff changeset
   489
        [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
   490
    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
   491
    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
   492
    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
   493
    fun qualify suffix = Binding.qualified true suffix qty_name
47623
01e4fdf9d748 setup_lifting preprocesses forall_transfer rule by unfolding mem_Collect_eq
huffman
parents: 47608
diff changeset
   494
    val simplify = Raw_Simplifier.rewrite_rule [mk_meta_eq @{thm mem_Collect_eq}]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   495
    val opt_reflp_thm = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   496
      case typedef_set of
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   497
        Const ("Orderings.top_class.top", _) => 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   498
          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
   499
        | _ =>  NONE
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   500
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   501
    fun setup_transfer_rules_nonpar lthy =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   502
      let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   503
        val lthy =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   504
          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
   505
            SOME reflp_thm =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   506
              let 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   507
                val thms =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   508
                  [("id_abs_transfer",@{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
   509
                   ("bi_total",       @{thm Quotient_bi_total}       )]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   510
              in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   511
                fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   512
                    [[quot_thm, reflp_thm] MRSL thm])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   513
              end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   514
            | NONE =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   515
              let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   516
                val thms = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   517
                  [("All_transfer",   @{thm typedef_All_transfer}   ),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   518
                   ("Ex_transfer",    @{thm typedef_Ex_transfer}    )]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   519
              in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   520
                lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   521
                |> (snd oo Local_Theory.note) ((qualify "forall_transfer", [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   522
                  [simplify ([typedef_thm, T_def] MRSL @{thm typedef_forall_transfer})])
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   523
                |> fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   524
                  [[typedef_thm, T_def] MRSL thm])) thms
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   525
              end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   526
        val thms = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   527
          [("rep_transfer", @{thm typedef_rep_transfer}),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   528
           ("bi_unique",    @{thm typedef_bi_unique}   ),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   529
           ("right_unique", @{thm typedef_right_unique}), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   530
           ("right_total",  @{thm typedef_right_total} )]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   531
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   532
        fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   533
          [[typedef_thm, T_def] MRSL thm])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   534
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   535
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   536
    fun setup_transfer_rules_par lthy =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   537
      let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   538
        val pcrel_def = #pcrel_def (the (get_pcrel_info lthy qty_full_name))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   539
        val lthy =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   540
          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
   541
            SOME reflp_thm =>
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   542
              let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   543
                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
   544
                  (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
   545
                    ([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
   546
                val bi_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
   547
                    ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total})
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   548
                val thms = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   549
                  [("id_abs_transfer",id_abs_transfer),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   550
                   ("bi_total",       bi_total       )]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   551
              in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   552
                fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   553
                    [thm])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   554
              end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   555
            | NONE =>
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 = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   558
                  ("forall_transfer", simplify ([typedef_thm, T_def] MRSL @{thm typedef_forall_transfer})) 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   559
                  ::
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   560
                  (map_snd (fn thm => [typedef_thm, T_def] MRSL thm)                
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   561
                  [("All_transfer", @{thm typedef_All_transfer}),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   562
                   ("Ex_transfer",  @{thm typedef_Ex_transfer} )])
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   563
              in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   564
                fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   565
                  [parametrize_quantifier lthy thm])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   566
              end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   567
        val thms = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   568
          ("rep_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
   569
            (Lifting_Term.parametrize_transfer_rule lthy ([typedef_thm, T_def] MRSL @{thm typedef_rep_transfer})))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   570
          ::
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   571
          (map_snd (fn thm => parametrize_class_constraint lthy pcrel_def ([typedef_thm, T_def] MRSL thm))
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   572
          [("bi_unique",    @{thm typedef_bi_unique} ),
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   573
           ("right_unique", @{thm typedef_right_unique}), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   574
           ("right_total",  @{thm typedef_right_total} )])
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   575
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   576
        fold (fn (name, thm) => (snd oo Local_Theory.note) ((qualify name, [transfer_attr]), 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   577
          [thm])) thms lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   578
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   579
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   580
    fun setup_transfer_rules lthy = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   581
      if is_some (get_pcrel_info lthy qty_full_name) then setup_transfer_rules_par lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   582
                                                     else setup_transfer_rules_nonpar lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   583
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   584
  in
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   585
    lthy
47943
c09326cedb41 note Quotient theorem for typedefs in setup_lifting
kuncar
parents: 47937
diff changeset
   586
      |> (snd oo Local_Theory.note) ((Binding.prefix_name "Quotient_" qty_name, []), 
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   587
            [quot_thm])
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   588
      |> setup_lifting_infr gen_code quot_thm opt_reflp_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   589
      |> setup_transfer_rules
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   590
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   591
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   592
fun setup_lifting_cmd gen_code 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
   593
  let 
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   594
    val input_thm = singleton (Attrib.eval_thms lthy) xthm
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   595
    val input_term = (HOLogic.dest_Trueprop o prop_of) input_thm
47566
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   596
      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
   597
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   598
    fun sanity_check_reflp_thm reflp_thm = 
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   599
      let
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   600
        val reflp_tm = (HOLogic.dest_Trueprop o prop_of) reflp_thm
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   601
          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
   602
      in
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   603
        case reflp_tm of
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   604
          Const (@{const_name reflp}, _) $ _ => ()
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   605
          | _ => 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
   606
      end
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   607
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   608
    fun setup_quotient () = 
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   609
      let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   610
        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
   611
        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
   612
        val opt_par_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_par_xthm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   613
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   614
        setup_by_quotient gen_code input_thm opt_reflp_thm opt_par_thm lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   615
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   616
      
47566
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   617
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   618
    fun setup_typedef () = 
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   619
      case opt_reflp_xthm of
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   620
        SOME _ => error "The reflexivity theorem cannot be specified if the type_definition theorem is used."
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
   621
        | NONE => setup_by_typedef_thm gen_code input_thm lthy
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   622
  in
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47379
diff changeset
   623
    case input_term of
47566
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   624
      (Const (@{const_name Quotient}, _) $ _ $ _ $ _ $ _) => setup_quotient ()
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   625
      | (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
   626
      | _ => 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
   627
  end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   628
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
   629
val opt_gen_code =
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
   630
  Scan.optional (@{keyword "("} |-- Parse.!!! ((Parse.reserved "no_code" >> K false) --| @{keyword ")"})) true
47566
c201a1fe0a81 setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents: 47545
diff changeset
   631
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   632
val _ = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   633
  Outer_Syntax.local_theory @{command_spec "setup_lifting"}
50214
67fb9a168d10 tuned command descriptions;
wenzelm
parents: 50176
diff changeset
   634
    "setup lifting infrastructure" 
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   635
      (opt_gen_code -- Parse_Spec.xthm -- Scan.option Parse_Spec.xthm 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   636
      -- Scan.option (@{keyword "parametric"} |-- Parse.!!! Parse_Spec.xthm) >> 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   637
        (fn (((gen_code, xthm), opt_reflp_xthm), opt_par_xthm) => 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50288
diff changeset
   638
          setup_lifting_cmd gen_code xthm opt_reflp_xthm opt_par_xthm))
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   639
end;