src/HOL/Tools/Lifting/lifting_def.ML
author kuncar
Thu, 10 Apr 2014 17:48:18 +0200
changeset 57866 f4ba736040fa
parent 57861 c1048f5bbb45
child 57882 8267d1ff646f
permissions -rw-r--r--
setup for Transfer and Lifting from BNF; tuned thm names
     1 (*  Title:      HOL/Tools/Lifting/lifting_def.ML
     2     Author:     Ondrej Kuncar
     3 
     4 Definitions for constants on quotient types.
     5 *)
     6 
     7 signature LIFTING_DEF =
     8 sig
     9   val generate_parametric_transfer_rule:
    10     Proof.context -> thm -> thm -> thm
    11 
    12   val add_lift_def:
    13     (binding * mixfix) -> typ -> term -> thm -> thm list -> local_theory -> local_theory
    14 
    15   val lift_def_cmd:
    16     (binding * string option * mixfix) * string * (Facts.ref * Args.src list) list -> local_theory -> Proof.state
    17 
    18   val can_generate_code_cert: thm -> bool
    19 end
    20 
    21 structure Lifting_Def: LIFTING_DEF =
    22 struct
    23 
    24 open Lifting_Util
    25 
    26 infix 0 MRSL
    27 
    28 (* Reflexivity prover *)
    29 
    30 fun mono_eq_prover ctxt prop =
    31   let
    32     val refl_rules = Lifting_Info.get_reflexivity_rules ctxt
    33     val transfer_rules = Transfer.get_transfer_raw ctxt
    34     
    35     fun main_tac (t, i) =
    36       case HOLogic.dest_Trueprop t of 
    37         Const (@{const_name "less_eq"}, _) $ _ $ _ => REPEAT_ALL_NEW (resolve_tac refl_rules) i
    38         |  _ => REPEAT_ALL_NEW (resolve_tac transfer_rules) i
    39   in
    40     SOME (Goal.prove ctxt [] [] prop (K (SUBGOAL main_tac 1)))
    41       handle ERROR _ => NONE
    42   end
    43 
    44 fun try_prove_reflexivity ctxt prop =
    45   let
    46     val thy = Proof_Context.theory_of ctxt
    47     val cprop = cterm_of thy prop
    48     val rule = @{thm ge_eq_refl}
    49     val concl_pat = Drule.strip_imp_concl (cprop_of rule)
    50     val insts = Thm.first_order_match (concl_pat, cprop)
    51     val rule = Drule.instantiate_normalize insts rule
    52     val prop = hd (prems_of rule)
    53   in
    54     case mono_eq_prover ctxt prop of
    55       SOME thm => SOME (thm RS rule)
    56       | NONE => NONE
    57   end
    58 
    59 (* 
    60   Generates a parametrized transfer rule.
    61   transfer_rule - of the form T t f
    62   parametric_transfer_rule - of the form par_R t' t
    63   
    64   Result: par_T t' f, after substituing op= for relations in par_R that relate
    65     a type constructor to the same type constructor, it is a merge of (par_R' OO T) t' f
    66     using Lifting_Term.merge_transfer_relations
    67 *)
    68 
    69 fun generate_parametric_transfer_rule ctxt transfer_rule parametric_transfer_rule =
    70   let
    71     fun preprocess ctxt thm =
    72       let
    73         val tm = (strip_args 2 o HOLogic.dest_Trueprop o concl_of) thm;
    74         val param_rel = (snd o dest_comb o fst o dest_comb) tm;
    75         val thy = Proof_Context.theory_of ctxt;
    76         val free_vars = Term.add_vars param_rel [];
    77         
    78         fun make_subst (var as (_, typ)) subst = 
    79           let
    80             val [rty, rty'] = binder_types typ
    81           in
    82             if (Term.is_TVar rty andalso is_Type rty') then
    83               (Var var, HOLogic.eq_const rty')::subst
    84             else
    85               subst
    86           end;
    87         
    88         val subst = fold make_subst free_vars [];
    89         val csubst = map (pairself (cterm_of thy)) subst;
    90         val inst_thm = Drule.cterm_instantiate csubst thm;
    91       in
    92         Conv.fconv_rule 
    93           ((Conv.concl_conv (nprems_of inst_thm) o HOLogic.Trueprop_conv o Conv.fun2_conv o Conv.arg1_conv)
    94             (Raw_Simplifier.rewrite ctxt false (Transfer.get_sym_relator_eq ctxt))) inst_thm
    95       end
    96 
    97     fun inst_relcomppI thy ant1 ant2 =
    98       let
    99         val t1 = (HOLogic.dest_Trueprop o concl_of) ant1
   100         val t2 = (HOLogic.dest_Trueprop o prop_of) ant2
   101         val fun1 = cterm_of thy (strip_args 2 t1)
   102         val args1 = map (cterm_of thy) (get_args 2 t1)
   103         val fun2 = cterm_of thy (strip_args 2 t2)
   104         val args2 = map (cterm_of thy) (get_args 1 t2)
   105         val relcomppI = Drule.incr_indexes2 ant1 ant2 @{thm relcomppI}
   106         val vars = (rev (Term.add_vars (prop_of relcomppI) []))
   107         val subst = map (apfst ((cterm_of thy) o Var)) (vars ~~ ([fun1] @ args1 @ [fun2] @ args2))
   108       in
   109         Drule.cterm_instantiate subst relcomppI
   110       end
   111 
   112     fun zip_transfer_rules ctxt thm =
   113       let
   114         val thy = Proof_Context.theory_of ctxt
   115         fun mk_POS ty = Const (@{const_name POS}, ty --> ty --> HOLogic.boolT)
   116         val rel = (Thm.dest_fun2 o Thm.dest_arg o cprop_of) thm
   117         val typ = (typ_of o ctyp_of_term) rel
   118         val POS_const = cterm_of thy (mk_POS typ)
   119         val var = cterm_of thy (Var (("X", #maxidx (rep_cterm (rel)) + 1), typ))
   120         val goal = Thm.apply (cterm_of thy HOLogic.Trueprop) (Thm.apply (Thm.apply POS_const rel) var)
   121       in
   122         [Lifting_Term.merge_transfer_relations ctxt goal, thm] MRSL @{thm POS_apply}
   123       end
   124      
   125     val thm = (inst_relcomppI (Proof_Context.theory_of ctxt) parametric_transfer_rule transfer_rule) 
   126                 OF [parametric_transfer_rule, transfer_rule]
   127     val preprocessed_thm = preprocess ctxt thm
   128     val orig_ctxt = ctxt
   129     val (fixed_thm, ctxt) = yield_singleton (apfst snd oo Variable.import true) preprocessed_thm ctxt
   130     val assms = cprems_of fixed_thm
   131     val add_transfer_rule = Thm.attribute_declaration Transfer.transfer_add
   132     val (prems, ctxt) = fold_map Thm.assume_hyps assms ctxt
   133     val ctxt = Context.proof_map (fold add_transfer_rule prems) ctxt
   134     val zipped_thm =
   135       fixed_thm
   136       |> undisch_all
   137       |> zip_transfer_rules ctxt
   138       |> implies_intr_list assms
   139       |> singleton (Variable.export ctxt orig_ctxt)
   140   in
   141     zipped_thm
   142   end
   143 
   144 fun print_generate_transfer_info msg = 
   145   let
   146     val error_msg = cat_lines 
   147       ["Generation of a parametric transfer rule failed.",
   148       (Pretty.string_of (Pretty.block
   149          [Pretty.str "Reason:", Pretty.brk 2, msg]))]
   150   in
   151     error error_msg
   152   end
   153 
   154 fun map_ter _ x [] = x
   155     | map_ter f _ xs = map f xs
   156 
   157 fun generate_transfer_rules lthy quot_thm rsp_thm def_thm par_thms =
   158   let
   159     val transfer_rule =
   160       ([quot_thm, rsp_thm, def_thm] MRSL @{thm Quotient_to_transfer})
   161       |> Lifting_Term.parametrize_transfer_rule lthy
   162   in
   163     (map_ter (generate_parametric_transfer_rule lthy transfer_rule) [transfer_rule] par_thms
   164     handle Lifting_Term.MERGE_TRANSFER_REL msg => (print_generate_transfer_info msg; [transfer_rule]))
   165   end
   166 
   167 (* Generation of the code certificate from the rsp theorem *)
   168 
   169 fun get_body_types (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_types (U, V)
   170   | get_body_types (U, V)  = (U, V)
   171 
   172 fun get_binder_types (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types (U, W)
   173   | get_binder_types _ = []
   174 
   175 fun get_binder_types_by_rel (Const (@{const_name "rel_fun"}, _) $ _ $ S) (Type ("fun", [T, U]), Type ("fun", [V, W])) = 
   176     (T, V) :: get_binder_types_by_rel S (U, W)
   177   | get_binder_types_by_rel _ _ = []
   178 
   179 fun get_body_type_by_rel (Const (@{const_name "rel_fun"}, _) $ _ $ S) (Type ("fun", [_, U]), Type ("fun", [_, V])) = 
   180     get_body_type_by_rel S (U, V)
   181   | get_body_type_by_rel _ (U, V)  = (U, V)
   182 
   183 fun force_rty_type ctxt rty rhs = 
   184   let
   185     val thy = Proof_Context.theory_of ctxt
   186     val rhs_schematic = singleton (Variable.polymorphic ctxt) rhs
   187     val rty_schematic = fastype_of rhs_schematic
   188     val match = Sign.typ_match thy (rty_schematic, rty) Vartab.empty
   189   in
   190     Envir.subst_term_types match rhs_schematic
   191   end
   192 
   193 fun unabs_def ctxt def = 
   194   let
   195     val (_, rhs) = Thm.dest_equals (cprop_of def)
   196     fun dest_abs (Abs (var_name, T, _)) = (var_name, T)
   197       | dest_abs tm = raise TERM("get_abs_var",[tm])
   198     val (var_name, T) = dest_abs (term_of rhs)
   199     val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt
   200     val thy = Proof_Context.theory_of ctxt'
   201     val refl_thm = Thm.reflexive (cterm_of thy (Free (hd new_var_names, T)))
   202   in
   203     Thm.combination def refl_thm |>
   204     singleton (Proof_Context.export ctxt' ctxt)
   205   end
   206 
   207 fun unabs_all_def ctxt def = 
   208   let
   209     val (_, rhs) = Thm.dest_equals (cprop_of def)
   210     val xs = strip_abs_vars (term_of rhs)
   211   in  
   212     fold (K (unabs_def ctxt)) xs def
   213   end
   214 
   215 val map_fun_unfolded = 
   216   @{thm map_fun_def[abs_def]} |>
   217   unabs_def @{context} |>
   218   unabs_def @{context} |>
   219   Local_Defs.unfold @{context} [@{thm comp_def}]
   220 
   221 fun unfold_fun_maps ctm =
   222   let
   223     fun unfold_conv ctm =
   224       case (Thm.term_of ctm) of
   225         Const (@{const_name "map_fun"}, _) $ _ $ _ => 
   226           (Conv.arg_conv unfold_conv then_conv Conv.rewr_conv map_fun_unfolded) ctm
   227         | _ => Conv.all_conv ctm
   228   in
   229     (Conv.fun_conv unfold_conv) ctm
   230   end
   231 
   232 fun unfold_fun_maps_beta ctm =
   233   let val try_beta_conv = Conv.try_conv (Thm.beta_conversion false)
   234   in 
   235     (unfold_fun_maps then_conv try_beta_conv) ctm 
   236   end
   237 
   238 fun prove_rel ctxt rsp_thm (rty, qty) =
   239   let
   240     val ty_args = get_binder_types (rty, qty)
   241     fun disch_arg args_ty thm = 
   242       let
   243         val quot_thm = Lifting_Term.prove_quot_thm ctxt args_ty
   244       in
   245         [quot_thm, thm] MRSL @{thm apply_rsp''}
   246       end
   247   in
   248     fold disch_arg ty_args rsp_thm
   249   end
   250 
   251 exception CODE_CERT_GEN of string
   252 
   253 fun simplify_code_eq ctxt def_thm = 
   254   Local_Defs.unfold ctxt [@{thm o_apply}, @{thm map_fun_def}, @{thm id_apply}] def_thm
   255 
   256 (*
   257   quot_thm - quotient theorem (Quotient R Abs Rep T).
   258   returns: whether the Lifting package is capable to generate code for the abstract type
   259     represented by quot_thm
   260 *)
   261 
   262 fun can_generate_code_cert quot_thm  =
   263   case quot_thm_rel quot_thm of
   264     Const (@{const_name HOL.eq}, _) => true
   265     | Const (@{const_name eq_onp}, _) $ _  => true
   266     | _ => false
   267 
   268 fun generate_rep_eq ctxt def_thm rsp_thm (rty, qty) =
   269   let
   270     val unfolded_def = Conv.fconv_rule (Conv.arg_conv unfold_fun_maps_beta) def_thm
   271     val unabs_def = unabs_all_def ctxt unfolded_def
   272   in  
   273     if body_type rty = body_type qty then 
   274       SOME (simplify_code_eq ctxt (unabs_def RS @{thm meta_eq_to_obj_eq}))
   275     else 
   276       let
   277         val thy = Proof_Context.theory_of ctxt
   278         val quot_thm = Lifting_Term.prove_quot_thm ctxt (get_body_types (rty, qty))
   279         val rel_fun = prove_rel ctxt rsp_thm (rty, qty)
   280         val rep_abs_thm = [quot_thm, rel_fun] MRSL @{thm Quotient_rep_abs_eq}
   281       in
   282         case mono_eq_prover ctxt (hd(prems_of rep_abs_thm)) of
   283           SOME mono_eq_thm =>
   284             let
   285               val rep_abs_eq = mono_eq_thm RS rep_abs_thm
   286               val rep = (cterm_of thy o quot_thm_rep) quot_thm
   287               val rep_refl = Thm.reflexive rep RS @{thm meta_eq_to_obj_eq}
   288               val repped_eq = [rep_refl, unabs_def RS @{thm meta_eq_to_obj_eq}] MRSL @{thm cong}
   289               val code_cert = [repped_eq, rep_abs_eq] MRSL trans
   290             in
   291               SOME (simplify_code_eq ctxt code_cert)
   292             end
   293           | NONE => NONE
   294       end
   295   end
   296 
   297 fun generate_abs_eq ctxt def_thm rsp_thm quot_thm =
   298   let
   299     val abs_eq_with_assms =
   300       let
   301         val (rty, qty) = quot_thm_rty_qty quot_thm
   302         val rel = quot_thm_rel quot_thm
   303         val ty_args = get_binder_types_by_rel rel (rty, qty)
   304         val body_type = get_body_type_by_rel rel (rty, qty)
   305         val quot_ret_thm = Lifting_Term.prove_quot_thm ctxt body_type
   306         
   307         val rep_abs_folded_unmapped_thm = 
   308           let
   309             val rep_id = [quot_thm, def_thm] MRSL @{thm Quotient_Rep_eq}
   310             val ctm = Thm.dest_equals_lhs (cprop_of rep_id)
   311             val unfolded_maps_eq = unfold_fun_maps ctm
   312             val t1 = [quot_thm, def_thm, rsp_thm] MRSL @{thm Quotient_rep_abs_fold_unmap}
   313             val prems_pat = (hd o Drule.cprems_of) t1
   314             val insts = Thm.first_order_match (prems_pat, cprop_of unfolded_maps_eq)
   315           in
   316             unfolded_maps_eq RS (Drule.instantiate_normalize insts t1)
   317           end
   318       in
   319         rep_abs_folded_unmapped_thm
   320         |> fold (fn _ => fn thm => thm RS @{thm rel_funD2}) ty_args
   321         |> (fn x => x RS (@{thm Quotient_rel_abs2} OF [quot_ret_thm]))
   322       end
   323     
   324     val prems = prems_of abs_eq_with_assms
   325     val indexed_prems = map_index (apfst (fn x => x + 1)) prems
   326     val indexed_assms = map (apsnd (try_prove_reflexivity ctxt)) indexed_prems
   327     val proved_assms = map (apsnd the) (filter (is_some o snd) indexed_assms)
   328     val abs_eq = fold_rev (fn (i, assms) => fn thm => assms RSN (i, thm)) proved_assms abs_eq_with_assms
   329   in
   330     simplify_code_eq ctxt abs_eq
   331   end
   332 
   333 
   334 fun register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy =
   335   let
   336     fun no_abstr (t $ u) = no_abstr t andalso no_abstr u
   337       | no_abstr (Abs (_, _, t)) = no_abstr t
   338       | no_abstr (Const (name, _)) = not (Code.is_abstr thy name)
   339       | no_abstr _ = true
   340     fun is_valid_eq eqn = can (Code.assert_eqn thy) (mk_meta_eq eqn, true) 
   341       andalso no_abstr (prop_of eqn)
   342     fun is_valid_abs_eq abs_eq = can (Code.assert_abs_eqn thy NONE) (mk_meta_eq abs_eq)
   343 
   344   in
   345     if is_valid_eq abs_eq_thm then
   346       Code.add_default_eqn abs_eq_thm thy
   347     else
   348       let
   349         val (rty_body, qty_body) = get_body_types (rty, qty)
   350       in
   351         if rty_body = qty_body then
   352          Code.add_default_eqn (the opt_rep_eq_thm) thy
   353         else
   354           if is_some opt_rep_eq_thm andalso is_valid_abs_eq (the opt_rep_eq_thm)
   355           then
   356             Code.add_abs_eqn (the opt_rep_eq_thm) thy
   357           else
   358             thy
   359       end
   360   end
   361 
   362 local
   363   fun encode_code_eq thy abs_eq opt_rep_eq (rty, qty) = 
   364     let
   365       fun mk_type typ = typ |> Logic.mk_type |> cterm_of thy |> Drule.mk_term
   366     in
   367       Conjunction.intr_balanced [abs_eq, (the_default TrueI opt_rep_eq), mk_type rty, mk_type qty]
   368     end
   369   
   370   fun decode_code_eq thm =
   371     let
   372       val [abs_eq, rep_eq, rty, qty] = Conjunction.elim_balanced 4 thm
   373       val opt_rep_eq = if Thm.eq_thm_prop (rep_eq, TrueI) then NONE else SOME rep_eq
   374       fun dest_type typ = typ |> Drule.dest_term |> term_of |> Logic.dest_type
   375     in
   376       (abs_eq, opt_rep_eq, (dest_type rty, dest_type qty)) 
   377     end
   378   
   379   fun register_encoded_code_eq thm thy =
   380     let
   381       val (abs_eq_thm, opt_rep_eq_thm, (rty, qty)) = decode_code_eq thm
   382     in
   383       register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy
   384     end
   385   
   386   val register_code_eq_attribute = Thm.declaration_attribute
   387     (fn thm => Context.mapping (register_encoded_code_eq thm) I)
   388   val register_code_eq_attrib = Attrib.internal (K register_code_eq_attribute)
   389 in
   390 
   391 fun register_code_eq abs_eq_thm opt_rep_eq_thm (rty, qty) lthy =
   392   let
   393     val thy = Proof_Context.theory_of lthy
   394     val encoded_code_eq = encode_code_eq thy abs_eq_thm opt_rep_eq_thm (rty, qty)
   395   in
   396     (snd oo Local_Theory.note) ((Binding.empty, [register_code_eq_attrib]), 
   397       [encoded_code_eq]) lthy
   398   end
   399 end
   400             
   401 (*
   402   Defines an operation on an abstract type in terms of a corresponding operation 
   403     on a representation type.
   404 
   405   var - a binding and a mixfix of the new constant being defined
   406   qty - an abstract type of the new constant
   407   rhs - a term representing the new constant on the raw level
   408   rsp_thm - a respectfulness theorem in the internal tagged form (like '(R ===> R ===> R) f f'),
   409     i.e. "(Lifting_Term.equiv_relation (fastype_of rhs, qty)) $ rhs $ rhs"
   410   par_thms - a parametricity theorem for rhs
   411 *)
   412 
   413 fun add_lift_def var qty rhs rsp_thm par_thms lthy =
   414   let
   415     val rty = fastype_of rhs
   416     val quot_thm = Lifting_Term.prove_quot_thm lthy (rty, qty)
   417     val absrep_trm =  quot_thm_abs quot_thm
   418     val rty_forced = (domain_type o fastype_of) absrep_trm
   419     val forced_rhs = force_rty_type lthy rty_forced rhs
   420     val lhs = Free (Binding.name_of (#1 var), qty)
   421     val prop = Logic.mk_equals (lhs, absrep_trm $ forced_rhs)
   422     val (_, prop') = Local_Defs.cert_def lthy prop
   423     val (_, newrhs) = Local_Defs.abs_def prop'
   424 
   425     val ((_, (_ , def_thm)), lthy') = 
   426       Local_Theory.define (var, ((Thm.def_binding (#1 var), []), newrhs)) lthy
   427 
   428     val transfer_rules = generate_transfer_rules lthy' quot_thm rsp_thm def_thm par_thms
   429 
   430     val abs_eq_thm = generate_abs_eq lthy' def_thm rsp_thm quot_thm
   431     val opt_rep_eq_thm = generate_rep_eq lthy' def_thm rsp_thm (rty_forced, qty)
   432 
   433     fun qualify defname suffix = Binding.qualified true suffix defname
   434 
   435     val lhs_name = (#1 var)
   436     val rsp_thm_name = qualify lhs_name "rsp"
   437     val abs_eq_thm_name = qualify lhs_name "abs_eq"
   438     val rep_eq_thm_name = qualify lhs_name "rep_eq"
   439     val transfer_rule_name = qualify lhs_name "transfer"
   440     val transfer_attr = Attrib.internal (K Transfer.transfer_add)
   441   in
   442     lthy'
   443       |> (snd oo Local_Theory.note) ((rsp_thm_name, []), [rsp_thm])
   444       |> (snd oo Local_Theory.note) ((transfer_rule_name, [transfer_attr]), transfer_rules)
   445       |> (snd oo Local_Theory.note) ((abs_eq_thm_name, []), [abs_eq_thm])
   446       |> (case opt_rep_eq_thm of 
   447             SOME rep_eq_thm => (snd oo Local_Theory.note) ((rep_eq_thm_name, []), [rep_eq_thm])
   448             | NONE => I)
   449       |> register_code_eq abs_eq_thm opt_rep_eq_thm (rty_forced, qty)
   450   end
   451 
   452 local
   453   val eq_onp_assms_tac_fixed_rules = map (Transfer.prep_transfer_domain_thm @{context})
   454     [@{thm pcr_Domainp_total}, @{thm pcr_Domainp_par_left_total}, @{thm pcr_Domainp_par}, 
   455       @{thm pcr_Domainp}]
   456 in
   457 fun mk_readable_rsp_thm_eq tm lthy =
   458   let
   459     val ctm = cterm_of (Proof_Context.theory_of lthy) tm
   460     
   461     (* This is not very cheap way of getting the rules but we have only few active
   462       liftings in the current setting *)
   463     fun get_cr_pcr_eqs ctxt =
   464       let
   465         fun collect (data : Lifting_Info.quotient) l =
   466           if is_some (#pcr_info data) 
   467           then ((Thm.symmetric o safe_mk_meta_eq o #pcr_cr_eq o the o #pcr_info) data :: l) 
   468           else l
   469         val table = Lifting_Info.get_quotients ctxt
   470       in
   471         Symtab.fold (fn (_, data) => fn l => collect data l) table []
   472       end
   473 
   474     fun assms_rewr_conv tactic rule ct =
   475       let
   476         fun prove_extra_assms thm =
   477           let
   478             val assms = cprems_of thm
   479             fun finish thm = if Thm.no_prems thm then SOME (Goal.conclude thm) else NONE
   480             fun prove ctm = Option.mapPartial finish (SINGLE tactic (Goal.init ctm))
   481           in
   482             map_interrupt prove assms
   483           end
   484     
   485         fun cconl_of thm = Drule.strip_imp_concl (cprop_of thm)
   486         fun lhs_of thm = fst (Thm.dest_equals (cconl_of thm))
   487         fun rhs_of thm = snd (Thm.dest_equals (cconl_of thm))
   488         val rule1 = Thm.incr_indexes (#maxidx (Thm.rep_cterm ct) + 1) rule;
   489         val lhs = lhs_of rule1;
   490         val rule2 = Thm.rename_boundvars (Thm.term_of lhs) (Thm.term_of ct) rule1;
   491         val rule3 =
   492           Thm.instantiate (Thm.match (lhs, ct)) rule2
   493             handle Pattern.MATCH => raise CTERM ("assms_rewr_conv", [lhs, ct]);
   494         val proved_assms = prove_extra_assms rule3
   495       in
   496         case proved_assms of
   497           SOME proved_assms =>
   498             let
   499               val rule3 = proved_assms MRSL rule3
   500               val rule4 =
   501                 if lhs_of rule3 aconvc ct then rule3
   502                 else
   503                   let val ceq = Thm.dest_fun2 (Thm.cprop_of rule3)
   504                   in rule3 COMP Thm.trivial (Thm.mk_binop ceq ct (rhs_of rule3)) end
   505             in Thm.transitive rule4 (Thm.beta_conversion true (rhs_of rule4)) end
   506           | NONE => Conv.no_conv ct
   507       end
   508 
   509     fun assms_rewrs_conv tactic rules = Conv.first_conv (map (assms_rewr_conv tactic) rules)
   510 
   511     fun simp_arrows_conv ctm =
   512       let
   513         val unfold_conv = Conv.rewrs_conv 
   514           [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, 
   515             @{thm rel_fun_eq_onp_rel[THEN eq_reflection]},
   516             @{thm rel_fun_eq[THEN eq_reflection]},
   517             @{thm rel_fun_eq_rel[THEN eq_reflection]}, 
   518             @{thm rel_fun_def[THEN eq_reflection]}]
   519         fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
   520         val eq_onp_assms_tac_rules = @{thm left_unique_OO} :: 
   521             eq_onp_assms_tac_fixed_rules @ (Transfer.get_transfer_raw lthy)
   522         val eq_onp_assms_tac = (TRY o REPEAT_ALL_NEW (resolve_tac eq_onp_assms_tac_rules) 
   523           THEN_ALL_NEW (DETERM o Transfer.eq_tac lthy)) 1
   524         val relator_eq_onp_conv = Conv.bottom_conv
   525           (K (Conv.try_conv (assms_rewrs_conv eq_onp_assms_tac
   526             (Lifting_Info.get_relator_eq_onp_rules lthy)))) lthy
   527         val relator_eq_conv = Conv.bottom_conv
   528           (K (Conv.try_conv (Conv.rewrs_conv (Transfer.get_relator_eq lthy)))) lthy
   529       in
   530         case (Thm.term_of ctm) of
   531           Const (@{const_name "rel_fun"}, _) $ _ $ _ => 
   532             (binop_conv2 simp_arrows_conv simp_arrows_conv then_conv unfold_conv) ctm
   533           | _ => (relator_eq_onp_conv then_conv relator_eq_conv) ctm
   534       end
   535     
   536     val unfold_ret_val_invs = Conv.bottom_conv 
   537       (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) lthy
   538     val cr_to_pcr_conv = Raw_Simplifier.rewrite lthy false (get_cr_pcr_eqs lthy)
   539     val unfold_inv_conv = 
   540       Conv.top_sweep_conv (K (Conv.rewr_conv @{thm eq_onp_def[THEN eq_reflection]})) lthy
   541     val simp_conv = HOLogic.Trueprop_conv (Conv.fun2_conv 
   542       (cr_to_pcr_conv then_conv simp_arrows_conv))
   543     val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]}
   544     val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) lthy
   545     val beta_conv = Thm.beta_conversion true
   546     val eq_thm = 
   547       (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs
   548          then_conv unfold_inv_conv) ctm
   549   in
   550     Object_Logic.rulify lthy (eq_thm RS Drule.equal_elim_rule2)
   551   end
   552 end
   553 
   554 fun rename_to_tnames ctxt term =
   555   let
   556     fun all_typs (Const (@{const_name Pure.all}, _) $ Abs (_, T, t)) = T :: all_typs t
   557       | all_typs _ = []
   558 
   559     fun rename (Const (@{const_name Pure.all}, T1) $ Abs (_, T2, t)) (new_name :: names) = 
   560         (Const (@{const_name Pure.all}, T1) $ Abs (new_name, T2, rename t names)) 
   561       | rename t _ = t
   562 
   563     val (fixed_def_t, _) = yield_singleton (Variable.importT_terms) term ctxt
   564     val new_names = Datatype_Prop.make_tnames (all_typs fixed_def_t)
   565   in
   566     rename term new_names
   567   end
   568 
   569 (*
   570 
   571   lifting_definition command. It opens a proof of a corresponding respectfulness 
   572   theorem in a user-friendly, readable form. Then add_lift_def is called internally.
   573 
   574 *)
   575 
   576 fun lift_def_cmd (raw_var, rhs_raw, par_xthms) lthy =
   577   let
   578     val ((binding, SOME qty, mx), lthy) = yield_singleton Proof_Context.read_vars raw_var lthy 
   579     val rhs = (Syntax.check_term lthy o Syntax.parse_term lthy) rhs_raw
   580     val rsp_rel = Lifting_Term.equiv_relation lthy (fastype_of rhs, qty)
   581     val rty_forced = (domain_type o fastype_of) rsp_rel;
   582     val forced_rhs = force_rty_type lthy rty_forced rhs;
   583     val internal_rsp_tm = HOLogic.mk_Trueprop (rsp_rel $ forced_rhs $ forced_rhs)
   584     val opt_proven_rsp_thm = try_prove_reflexivity lthy internal_rsp_tm
   585     val par_thms = Attrib.eval_thms lthy par_xthms
   586     
   587     fun after_qed internal_rsp_thm lthy = 
   588       add_lift_def (binding, mx) qty rhs internal_rsp_thm par_thms lthy
   589 
   590   in
   591     case opt_proven_rsp_thm of
   592       SOME thm => Proof.theorem NONE (K (after_qed thm)) [] lthy
   593       | NONE =>  
   594         let
   595           val readable_rsp_thm_eq = mk_readable_rsp_thm_eq internal_rsp_tm lthy
   596           val (readable_rsp_tm, _) = Logic.dest_implies (prop_of readable_rsp_thm_eq)
   597           val readable_rsp_tm_tnames = rename_to_tnames lthy readable_rsp_tm
   598       
   599           fun after_qed' thm_list lthy = 
   600             let
   601               val internal_rsp_thm = Goal.prove lthy [] [] internal_rsp_tm 
   602                   (fn {context = ctxt, ...} =>
   603                     rtac readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac ctxt (hd thm_list) 1)
   604             in
   605               after_qed internal_rsp_thm lthy
   606             end
   607         in
   608           Proof.theorem NONE after_qed' [[(readable_rsp_tm_tnames,[])]] lthy
   609         end 
   610   end
   611 
   612 fun quot_thm_err ctxt (rty, qty) pretty_msg =
   613   let
   614     val error_msg = cat_lines
   615        ["Lifting failed for the following types:",
   616         Pretty.string_of (Pretty.block
   617          [Pretty.str "Raw type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty]),
   618         Pretty.string_of (Pretty.block
   619          [Pretty.str "Abstract type:", Pretty.brk 2, Syntax.pretty_typ ctxt qty]),
   620         "",
   621         (Pretty.string_of (Pretty.block
   622          [Pretty.str "Reason:", Pretty.brk 2, pretty_msg]))]
   623   in
   624     error error_msg
   625   end
   626 
   627 fun check_rty_err ctxt (rty_schematic, rty_forced) (raw_var, rhs_raw) =
   628   let
   629     val (_, ctxt') = yield_singleton Proof_Context.read_vars raw_var ctxt 
   630     val rhs = (Syntax.check_term ctxt' o Syntax.parse_term ctxt') rhs_raw
   631     val error_msg = cat_lines
   632        ["Lifting failed for the following term:",
   633         Pretty.string_of (Pretty.block
   634          [Pretty.str "Term:", Pretty.brk 2, Syntax.pretty_term ctxt rhs]),
   635         Pretty.string_of (Pretty.block
   636          [Pretty.str "Type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty_schematic]),
   637         "",
   638         (Pretty.string_of (Pretty.block
   639          [Pretty.str "Reason:", 
   640           Pretty.brk 2, 
   641           Pretty.str "The type of the term cannot be instantiated to",
   642           Pretty.brk 1,
   643           Pretty.quote (Syntax.pretty_typ ctxt rty_forced),
   644           Pretty.str "."]))]
   645     in
   646       error error_msg
   647     end
   648 
   649 fun lift_def_cmd_with_err_handling (raw_var, rhs_raw, par_xthms) lthy =
   650   (lift_def_cmd (raw_var, rhs_raw, par_xthms) lthy
   651     handle Lifting_Term.QUOT_THM (rty, qty, msg) => quot_thm_err lthy (rty, qty) msg)
   652     handle Lifting_Term.CHECK_RTY (rty_schematic, rty_forced) => 
   653       check_rty_err lthy (rty_schematic, rty_forced) (raw_var, rhs_raw)
   654 
   655 (* parser and command *)
   656 val liftdef_parser =
   657   (((Parse.binding -- (@{keyword "::"} |-- (Parse.typ >> SOME) -- Parse.opt_mixfix')) >> Parse.triple2)
   658     --| @{keyword "is"} -- Parse.term -- 
   659       Scan.optional (@{keyword "parametric"} |-- Parse.!!! Parse_Spec.xthms1) []) >> Parse.triple1
   660 val _ =
   661   Outer_Syntax.local_theory_to_proof @{command_spec "lift_definition"}
   662     "definition for constants over the quotient type"
   663       (liftdef_parser >> lift_def_cmd_with_err_handling)
   664 
   665 
   666 end (* structure *)