src/HOL/Tools/Quotient/quotient_def.ML
author kuncar
Fri, 23 Mar 2012 14:03:58 +0100
changeset 47091 d5cd13aca90b
parent 46961 5c6955f487e5
child 47096 3ea48c19673e
permissions -rw-r--r--
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 37592
diff changeset
     1
(*  Title:      HOL/Tools/Quotient/quotient_def.ML
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
    Author:     Cezary Kaliszyk and Christian Urban
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
35788
f1deaca15ca3 observe standard header format;
wenzelm
parents: 35624
diff changeset
     4
Definitions for constants on quotient types.
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
*)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
signature QUOTIENT_DEF =
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
sig
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
     9
  val add_quotient_def:
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    10
    ((binding * mixfix) * Attrib.binding) * (term * term) -> thm ->
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    11
    local_theory -> Quotient_Info.quotconsts * local_theory
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    12
45598
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
    13
  val quotient_def:
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
    14
    (binding * typ option * mixfix) option * (Attrib.binding * (term * term)) ->
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    15
    local_theory -> Proof.state
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
45598
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
    17
  val quotient_def_cmd:
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
    18
    (binding * string option * mixfix) option * (Attrib.binding * (string * string)) ->
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    19
    local_theory -> Proof.state
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
37561
19fca77829ea mixfix can be given for automatically lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 37532
diff changeset
    21
  val lift_raw_const: typ list -> (string * term * mixfix) -> local_theory ->
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 43663
diff changeset
    22
    Quotient_Info.quotconsts * local_theory
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
end;
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
structure Quotient_Def: QUOTIENT_DEF =
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
struct
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
(** Interface and Syntax Setup **)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
(* The ML-interface for a quotient definition takes
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
   as argument:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
    - an optional binding and mixfix annotation
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
    - attributes
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
    - the new constant as term
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
    - the rhs of the definition as term
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    37
    - respectfulness theorem for the rhs
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 43663
diff changeset
    39
   It stores the qconst_info in the quotconsts data slot.
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    41
   Restriction: At the moment the left- and right-hand
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    42
   side of the definition must be a constant.
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
*)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    44
fun error_msg bind str =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    45
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    46
    val name = Binding.name_of bind
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    47
    val pos = Position.str_of (Binding.pos_of bind)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    48
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    49
    error ("Head of quotient_definition " ^
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    50
      quote str ^ " differs from declaration " ^ name ^ pos)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    51
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    53
fun add_quotient_def ((var, (name, atts)), (lhs, rhs)) rsp_thm lthy =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    54
  let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    55
    val absrep_trm = 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    56
      Quotient_Term.absrep_fun lthy Quotient_Term.AbsF (fastype_of rhs, fastype_of lhs) $ rhs
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    57
    val prop = Syntax.check_term lthy (Logic.mk_equals (lhs, absrep_trm))
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    58
    val (_, prop') = Local_Defs.cert_def lthy prop
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    59
    val (_, newrhs) = Local_Defs.abs_def prop'
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    60
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    61
    val ((trm, (_ , thm)), lthy') =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    62
      Local_Theory.define (var, ((Thm.def_binding_optional (#1 var) name, atts), newrhs)) lthy
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    63
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    64
    (* data storage *)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    65
    val qconst_data = {qconst = trm, rconst = rhs, def = thm}
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    66
    fun get_rsp_thm_name (lhs_name, _) = Binding.suffix_name "_rsp" lhs_name
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    67
    
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    68
    val lthy'' = lthy'
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    69
      |> Local_Theory.declaration {syntax = false, pervasive = true}
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    70
        (fn phi =>
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    71
          (case Quotient_Info.transform_quotconsts phi qconst_data of
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    72
            qcinfo as {qconst = Const (c, _), ...} =>
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    73
              Quotient_Info.update_quotconsts c qcinfo
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    74
          | _ => I))
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    75
      |> (snd oo Local_Theory.note) 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    76
        ((get_rsp_thm_name var, [Attrib.internal (K Quotient_Info.rsp_rules_add)]),
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    77
        [rsp_thm])
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    78
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    79
  in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    80
    (qconst_data, lthy'')
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    81
  end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    82
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    83
fun mk_readable_rsp_thm_eq tm lthy =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    84
  let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    85
    val ctm = cterm_of (Proof_Context.theory_of lthy) tm
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    86
    
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    87
    fun norm_fun_eq ctm = 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    88
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    89
        fun abs_conv2 cv = Conv.abs_conv (K (Conv.abs_conv (K cv) lthy)) lthy
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    90
        fun erase_quants ctm' =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    91
          case (Thm.term_of ctm') of
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    92
            Const ("HOL.eq", _) $ _ $ _ => Conv.all_conv ctm'
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    93
            | _ => (Conv.binder_conv (K erase_quants) lthy then_conv 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    94
              Conv.rewr_conv @{thm fun_eq_iff[symmetric, THEN eq_reflection]}) ctm'
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    95
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    96
        (abs_conv2 erase_quants then_conv Thm.eta_conversion) ctm
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    97
      end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    98
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    99
    fun simp_arrows_conv ctm =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   100
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   101
        val unfold_conv = Conv.rewrs_conv 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   102
          [@{thm fun_rel_eq_rel[THEN eq_reflection]}, @{thm fun_rel_def[THEN eq_reflection]}]
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   103
        val left_conv = simp_arrows_conv then_conv Conv.try_conv norm_fun_eq
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   104
        fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   105
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   106
        case (Thm.term_of ctm) of
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   107
          Const (@{const_name "fun_rel"}, _) $ _ $ _ => 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   108
            (binop_conv2  left_conv simp_arrows_conv then_conv unfold_conv) ctm
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   109
          | _ => Conv.all_conv ctm
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   110
      end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   111
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   112
    val simp_conv = Conv.arg_conv (Conv.fun2_conv simp_arrows_conv)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   113
    val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]}
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   114
    val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) lthy
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   115
    val eq_thm = 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   116
      (simp_conv then_conv univq_prenex_conv then_conv Thm.beta_conversion true) ctm
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   117
  in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   118
    Object_Logic.rulify(eq_thm RS Drule.equal_elim_rule2)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   119
  end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   120
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   121
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   122
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   123
fun gen_quotient_def prep_vars prep_term (raw_var, (attr, (lhs_raw, rhs_raw))) lthy =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   124
  let
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   125
    val (vars, ctxt) = prep_vars (the_list raw_var) lthy
43663
e8c80bbc0c5d re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents: 43608
diff changeset
   126
    val T_opt = (case vars of [(_, SOME T, _)] => SOME T | _ => NONE)
e8c80bbc0c5d re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents: 43608
diff changeset
   127
    val lhs = prep_term T_opt ctxt lhs_raw
e8c80bbc0c5d re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents: 43608
diff changeset
   128
    val rhs = prep_term NONE ctxt rhs_raw
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   129
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   130
    val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined."
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   131
    val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction"
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   132
    val _ = if is_Const rhs then () else warning "The definiens is not a constant"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   134
    val var =
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   135
      (case vars of 
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   136
        [] => (Binding.name lhs_str, NoSyn)
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   137
      | [(binding, _, mx)] =>
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   138
          if Variable.check_name binding = lhs_str then (binding, mx)
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   139
          else error_msg binding lhs_str
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   140
      | _ => raise Match)
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   141
    
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   142
    fun try_to_prove_refl thm = 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   143
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   144
        val lhs_eq =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   145
          thm
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   146
          |> prop_of
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   147
          |> Logic.dest_implies
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   148
          |> fst
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   149
          |> strip_all_body
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   150
          |> try HOLogic.dest_Trueprop
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   151
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   152
        case lhs_eq of
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   153
          SOME (Const ("HOL.eq", _) $ _ $ _) => SOME (@{thm refl} RS thm)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   154
          | SOME _ => (case body_type (fastype_of lhs) of
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   155
            Type (typ_name, _) => ( SOME
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   156
              (#equiv_thm (the (Quotient_Info.lookup_quotients lthy typ_name)) 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   157
                RS @{thm Equiv_Relations.equivp_reflp} RS thm)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   158
              handle _ => NONE)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   159
            | _ => NONE
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   160
            )
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   161
          | _ => NONE
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   162
      end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   164
    val rsp_rel = Quotient_Term.equiv_relation lthy (fastype_of rhs, lhs_ty)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   165
    val internal_rsp_tm = HOLogic.mk_Trueprop (Syntax.check_term lthy (rsp_rel $ rhs $ rhs))
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   166
    val readable_rsp_thm_eq = mk_readable_rsp_thm_eq internal_rsp_tm lthy
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   167
    val maybe_proven_rsp_thm = try_to_prove_refl readable_rsp_thm_eq
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   168
    val (readable_rsp_tm, _) = Logic.dest_implies (prop_of readable_rsp_thm_eq)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   169
  
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   170
    fun after_qed thm_list lthy = 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   171
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   172
        val internal_rsp_thm =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   173
          case thm_list of
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   174
            [] => the maybe_proven_rsp_thm
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   175
          | [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm 
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   176
            (fn _ => rtac readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac [thm] 1)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   177
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   178
        snd (add_quotient_def ((var, attr), (lhs, rhs)) internal_rsp_thm lthy)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   179
      end
37530
70d03844b2f9 export of proper information in the ML-interface of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 36960
diff changeset
   180
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   181
  in
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   182
    case maybe_proven_rsp_thm of
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   183
      SOME _ => Proof.theorem NONE after_qed [] lthy
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   184
      | NONE =>  Proof.theorem NONE after_qed [[(readable_rsp_tm,[])]] lthy
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   185
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
43663
e8c80bbc0c5d re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents: 43608
diff changeset
   187
fun check_term' cnstr ctxt =
e8c80bbc0c5d re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents: 43608
diff changeset
   188
  Syntax.check_term ctxt o (case cnstr of SOME T => Type.constraint T | _ => I)
45598
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
   189
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
   190
fun read_term' cnstr ctxt =
45929
d7d6c8cfb6f6 removing some debug output in quotient_definition
bulwahn
parents: 45826
diff changeset
   191
  check_term' cnstr ctxt o Syntax.parse_term ctxt
43663
e8c80bbc0c5d re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents: 43608
diff changeset
   192
e8c80bbc0c5d re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents: 43608
diff changeset
   193
val quotient_def = gen_quotient_def Proof_Context.cert_vars check_term'
45598
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
   194
val quotient_def_cmd = gen_quotient_def Proof_Context.read_vars read_term'
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   195
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
37532
8a9a34be09e0 slight cleaning and simplification of the automatic wrapper for quotient definitions
Christian Urban <urbanc@in.tum.de>
parents: 37530
diff changeset
   197
(* a wrapper for automatically lifting a raw constant *)
37561
19fca77829ea mixfix can be given for automatically lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 37532
diff changeset
   198
fun lift_raw_const qtys (qconst_name, rconst, mx) ctxt =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   199
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   200
    val rty = fastype_of rconst
41451
892e67be8304 do not open ML structures;
wenzelm
parents: 41444
diff changeset
   201
    val qty = Quotient_Term.derive_qtyp ctxt qtys rty
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   202
    val lhs = Free (qconst_name, qty)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   203
  in
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   204
    (*quotient_def (SOME (Binding.name qconst_name, NONE, mx), (Attrib.empty_binding, (lhs, rconst))) ctxt*)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   205
    ({qconst = lhs, rconst = lhs, def = @{thm refl}}, ctxt)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   206
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   207
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   208
(* parser and command *)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   209
val quotdef_parser =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   210
  Scan.option Parse_Spec.constdecl --
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   211
    Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term --| @{keyword "is"} -- Parse.term))
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   212
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   213
val _ =
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   214
  Outer_Syntax.local_theory_to_proof @{command_spec "quotient_definition"}
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   215
    "definition for constants over the quotient type"
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   216
      (quotdef_parser >> quotient_def_cmd)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   217
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   218
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   219
end; (* structure *)