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