src/HOL/Tools/Quotient/quotient_def.ML
author kuncar
Thu, 10 Apr 2014 17:48:18 +0200
changeset 57866 f4ba736040fa
parent 57861 c1048f5bbb45
child 59180 85ec71012df8
permissions -rw-r--r--
setup for Transfer and Lifting from BNF; tuned thm names
     1 (*  Title:      HOL/Tools/Quotient/quotient_def.ML
     2     Author:     Cezary Kaliszyk and Christian Urban
     3 
     4 Definitions for constants on quotient types.
     5 *)
     6 
     7 signature QUOTIENT_DEF =
     8 sig
     9   val add_quotient_def:
    10     ((binding * mixfix) * Attrib.binding) * (term * term) -> thm ->
    11     local_theory -> Quotient_Info.quotconsts * local_theory
    12 
    13   val quotient_def:
    14     (binding * typ option * mixfix) option * (Attrib.binding * (term * term)) ->
    15     local_theory -> Proof.state
    16 
    17   val quotient_def_cmd:
    18     (binding * string option * mixfix) option * (Attrib.binding * (string * string)) ->
    19     local_theory -> Proof.state
    20 
    21 end;
    22 
    23 structure Quotient_Def: QUOTIENT_DEF =
    24 struct
    25 
    26 (** Interface and Syntax Setup **)
    27 
    28 (* Generation of the code certificate from the rsp theorem *)
    29 
    30 open Lifting_Util
    31 
    32 infix 0 MRSL
    33 
    34 (* The ML-interface for a quotient definition takes
    35    as argument:
    36 
    37     - an optional binding and mixfix annotation
    38     - attributes
    39     - the new constant as term
    40     - the rhs of the definition as term
    41     - respectfulness theorem for the rhs
    42 
    43    It stores the qconst_info in the quotconsts data slot.
    44 
    45    Restriction: At the moment the left- and right-hand
    46    side of the definition must be a constant.
    47 *)
    48 fun error_msg bind str =
    49   let
    50     val name = Binding.name_of bind
    51     val pos = Position.here (Binding.pos_of bind)
    52   in
    53     error ("Head of quotient_definition " ^
    54       quote str ^ " differs from declaration " ^ name ^ pos)
    55   end
    56 
    57 fun add_quotient_def ((var, (name, atts)), (lhs, rhs)) rsp_thm lthy =
    58   let
    59     val rty = fastype_of rhs
    60     val qty = fastype_of lhs
    61     val absrep_trm = 
    62       Quotient_Term.absrep_fun lthy Quotient_Term.AbsF (rty, qty) $ rhs
    63     val prop = Syntax.check_term lthy (Logic.mk_equals (lhs, absrep_trm))
    64     val (_, prop') = Local_Defs.cert_def lthy prop
    65     val (_, newrhs) = Local_Defs.abs_def prop'
    66 
    67     val ((trm, (_ , def_thm)), lthy') =
    68       Local_Theory.define (var, ((Thm.def_binding_optional (#1 var) name, atts), newrhs)) lthy
    69 
    70     (* data storage *)
    71     val qconst_data = {qconst = trm, rconst = rhs, def = def_thm}
    72      
    73     fun qualify defname suffix = Binding.name suffix
    74       |> Binding.qualify true defname
    75 
    76     val lhs_name = Binding.name_of (#1 var)
    77     val rsp_thm_name = qualify lhs_name "rsp"
    78     
    79     val lthy'' = lthy'
    80       |> Local_Theory.declaration {syntax = false, pervasive = true}
    81         (fn phi =>
    82           (case Quotient_Info.transform_quotconsts phi qconst_data of
    83             qcinfo as {qconst = Const (c, _), ...} =>
    84               Quotient_Info.update_quotconsts c qcinfo
    85           | _ => I))
    86       |> (snd oo Local_Theory.note) 
    87         ((rsp_thm_name, [Attrib.internal (K Quotient_Info.rsp_rules_add)]),
    88         [rsp_thm])
    89   in
    90     (qconst_data, lthy'')
    91   end
    92 
    93 fun mk_readable_rsp_thm_eq tm lthy =
    94   let
    95     val ctm = cterm_of (Proof_Context.theory_of lthy) tm
    96     
    97     fun norm_fun_eq ctm = 
    98       let
    99         fun abs_conv2 cv = Conv.abs_conv (K (Conv.abs_conv (K cv) lthy)) lthy
   100         fun erase_quants ctm' =
   101           case (Thm.term_of ctm') of
   102             Const (@{const_name HOL.eq}, _) $ _ $ _ => Conv.all_conv ctm'
   103             | _ => (Conv.binder_conv (K erase_quants) lthy then_conv 
   104               Conv.rewr_conv @{thm fun_eq_iff[symmetric, THEN eq_reflection]}) ctm'
   105       in
   106         (abs_conv2 erase_quants then_conv Thm.eta_conversion) ctm
   107       end
   108 
   109     fun simp_arrows_conv ctm =
   110       let
   111         val unfold_conv = Conv.rewrs_conv 
   112           [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, @{thm rel_fun_eq_rel[THEN eq_reflection]}, 
   113             @{thm rel_fun_def[THEN eq_reflection]}]
   114         val left_conv = simp_arrows_conv then_conv Conv.try_conv norm_fun_eq
   115         fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
   116       in
   117         case (Thm.term_of ctm) of
   118           Const (@{const_name rel_fun}, _) $ _ $ _ => 
   119             (binop_conv2 left_conv simp_arrows_conv then_conv unfold_conv) ctm
   120           | _ => Conv.all_conv ctm
   121       end
   122 
   123     val unfold_ret_val_invs = Conv.bottom_conv 
   124       (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) lthy 
   125     val simp_conv = Conv.arg_conv (Conv.fun2_conv simp_arrows_conv)
   126     val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]}
   127     val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) lthy
   128     val beta_conv = Thm.beta_conversion true
   129     val eq_thm = 
   130       (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs) ctm
   131   in
   132     Object_Logic.rulify lthy (eq_thm RS Drule.equal_elim_rule2)
   133   end
   134 
   135 
   136 
   137 fun gen_quotient_def prep_vars prep_term (raw_var, (attr, (lhs_raw, rhs_raw))) lthy =
   138   let
   139     val (vars, ctxt) = prep_vars (the_list raw_var) lthy
   140     val T_opt = (case vars of [(_, SOME T, _)] => SOME T | _ => NONE)
   141     val lhs = prep_term T_opt ctxt lhs_raw
   142     val rhs = prep_term NONE ctxt rhs_raw
   143 
   144     val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined."
   145     val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction"
   146     val _ = if is_Const rhs then () else warning "The definiens is not a constant"
   147 
   148     val var =
   149       (case vars of 
   150         [] => (Binding.name lhs_str, NoSyn)
   151       | [(binding, _, mx)] =>
   152           if Variable.check_name binding = lhs_str then (binding, mx)
   153           else error_msg binding lhs_str
   154       | _ => raise Match)
   155     
   156     fun try_to_prove_refl thm = 
   157       let
   158         val lhs_eq =
   159           thm
   160           |> prop_of
   161           |> Logic.dest_implies
   162           |> fst
   163           |> strip_all_body
   164           |> try HOLogic.dest_Trueprop
   165       in
   166         case lhs_eq of
   167           SOME (Const (@{const_name HOL.eq}, _) $ _ $ _) => SOME (@{thm refl} RS thm)
   168           | SOME _ => (case body_type (fastype_of lhs) of
   169             Type (typ_name, _) =>
   170               try (fn () =>
   171                 #equiv_thm (the (Quotient_Info.lookup_quotients lthy typ_name)) 
   172                   RS @{thm Equiv_Relations.equivp_reflp} RS thm) ()
   173             | _ => NONE
   174             )
   175           | _ => NONE
   176       end
   177 
   178     val rsp_rel = Quotient_Term.equiv_relation lthy (fastype_of rhs, lhs_ty)
   179     val internal_rsp_tm = HOLogic.mk_Trueprop (Syntax.check_term lthy (rsp_rel $ rhs $ rhs))
   180     val readable_rsp_thm_eq = mk_readable_rsp_thm_eq internal_rsp_tm lthy
   181     val maybe_proven_rsp_thm = try_to_prove_refl readable_rsp_thm_eq
   182     val (readable_rsp_tm, _) = Logic.dest_implies (prop_of readable_rsp_thm_eq)
   183   
   184     fun after_qed thm_list lthy = 
   185       let
   186         val internal_rsp_thm =
   187           case thm_list of
   188             [] => the maybe_proven_rsp_thm
   189           | [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm 
   190             (fn _ => rtac readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac ctxt [thm] 1)
   191       in
   192         snd (add_quotient_def ((var, attr), (lhs, rhs)) internal_rsp_thm lthy)
   193       end
   194 
   195   in
   196     case maybe_proven_rsp_thm of
   197       SOME _ => Proof.theorem NONE after_qed [] lthy
   198       | NONE =>  Proof.theorem NONE after_qed [[(readable_rsp_tm,[])]] lthy
   199   end
   200 
   201 fun check_term' cnstr ctxt =
   202   Syntax.check_term ctxt o (case cnstr of SOME T => Type.constraint T | _ => I)
   203 
   204 fun read_term' cnstr ctxt =
   205   check_term' cnstr ctxt o Syntax.parse_term ctxt
   206 
   207 val quotient_def = gen_quotient_def Proof_Context.cert_vars check_term'
   208 val quotient_def_cmd = gen_quotient_def Proof_Context.read_vars read_term'
   209 
   210 
   211 (* parser and command *)
   212 val quotdef_parser =
   213   Scan.option Parse_Spec.constdecl --
   214     Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term --| @{keyword "is"} -- Parse.term))
   215 
   216 val _ =
   217   Outer_Syntax.local_theory_to_proof @{command_spec "quotient_definition"}
   218     "definition for constants over the quotient type"
   219       (quotdef_parser >> quotient_def_cmd)
   220 
   221 
   222 end; (* structure *)