src/HOL/Tools/Quotient/quotient_def.ML
author bulwahn
Wed, 28 Mar 2012 10:44:04 +0200
changeset 47181 b351ad77eb78
parent 47156 861f53bd95fe
child 47198 cfd8ff62eab1
permissions -rw-r--r--
some tuning while reviewing the current state of the quotient_def package
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
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    30
(* Generation of the code certificate from the rsp theorem *)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    31
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    32
infix 0 MRSL
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    33
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    34
fun ants MRSL thm = fold (fn rl => fn thm => rl RS thm) ants thm
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    35
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    36
fun get_body_types (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_types (U, V)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    37
  | get_body_types (U, V)  = (U, V)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    38
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    39
fun get_binder_types (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types (U, W)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    40
  | get_binder_types _ = []
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    41
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    42
fun unabs_def ctxt def = 
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    43
  let
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    44
    val (_, rhs) = Thm.dest_equals (cprop_of def)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    45
    fun dest_abs (Abs (var_name, T, _)) = (var_name, T)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    46
      | dest_abs tm = raise TERM("get_abs_var",[tm])
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    47
    val (var_name, T) = dest_abs (term_of rhs)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    48
    val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    49
    val thy = Proof_Context.theory_of ctxt'
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    50
    val refl_thm = Thm.reflexive (cterm_of thy (Free (hd new_var_names, T)))
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    51
  in
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    52
    Thm.combination def refl_thm |>
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    53
    singleton (Proof_Context.export ctxt' ctxt)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    54
  end
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    55
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    56
fun unabs_all_def ctxt def = 
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    57
  let
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    58
    val (_, rhs) = Thm.dest_equals (cprop_of def)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    59
    val xs = strip_abs_vars (term_of rhs)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    60
  in  
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    61
    fold (K (unabs_def ctxt)) xs def
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    62
  end
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    63
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    64
val map_fun_unfolded = 
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    65
  @{thm map_fun_def[abs_def]} |>
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    66
  unabs_def @{context} |>
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    67
  unabs_def @{context} |>
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    68
  Local_Defs.unfold @{context} [@{thm comp_def}]
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    69
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    70
fun unfold_fun_maps ctm =
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    71
  let
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    72
    fun unfold_conv ctm =
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    73
      case (Thm.term_of ctm) of
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    74
        Const (@{const_name "map_fun"}, _) $ _ $ _ => 
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    75
          (Conv.arg_conv unfold_conv then_conv Conv.rewr_conv map_fun_unfolded) ctm
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    76
        | _ => Conv.all_conv ctm
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    77
    val try_beta_conv = Conv.try_conv (Thm.beta_conversion false)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    78
  in
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    79
    (Conv.arg_conv (Conv.fun_conv unfold_conv then_conv try_beta_conv)) ctm
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    80
  end
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    81
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    82
fun prove_rel ctxt rsp_thm (rty, qty) =
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    83
  let
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    84
    val ty_args = get_binder_types (rty, qty)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    85
    fun disch_arg args_ty thm = 
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    86
      let
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    87
        val quot_thm = Quotient_Term.prove_quot_theorem ctxt args_ty
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    88
      in
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    89
        [quot_thm, thm] MRSL @{thm apply_rsp''}
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    90
      end
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    91
  in
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    92
    fold disch_arg ty_args rsp_thm
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    93
  end
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    94
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    95
exception CODE_CERT_GEN of string
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    96
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    97
fun simplify_code_eq ctxt def_thm = 
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    98
  Local_Defs.unfold ctxt [@{thm o_def}, @{thm map_fun_def}, @{thm id_def}] def_thm
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
    99
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   100
fun generate_code_cert ctxt def_thm rsp_thm (rty, qty) =
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   101
  let
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   102
    val quot_thm = Quotient_Term.prove_quot_theorem ctxt (get_body_types (rty, qty))
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   103
    val fun_rel = prove_rel ctxt rsp_thm (rty, qty)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   104
    val abs_rep_thm = [quot_thm, fun_rel] MRSL @{thm Quotient_rep_abs}
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   105
    val abs_rep_eq = 
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   106
      case (HOLogic.dest_Trueprop o prop_of) fun_rel of
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   107
        Const (@{const_name HOL.eq}, _) $ _ $ _ => abs_rep_thm
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   108
        | Const (@{const_name invariant}, _) $ _ $ _ $ _ => abs_rep_thm RS @{thm invariant_to_eq}
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   109
        | _ => raise CODE_CERT_GEN "relation is neither equality nor invariant"
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   110
    val unfolded_def = Conv.fconv_rule unfold_fun_maps def_thm
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   111
    val unabs_def = unabs_all_def ctxt unfolded_def
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   112
    val rep = (snd o Thm.dest_comb o snd o Thm.dest_comb o cprop_of) quot_thm
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   113
    val rep_refl = Thm.reflexive rep RS @{thm meta_eq_to_obj_eq}
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   114
    val repped_eq = [rep_refl, unabs_def RS @{thm meta_eq_to_obj_eq}] MRSL @{thm cong}
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   115
    val code_cert = [repped_eq, abs_rep_eq] MRSL @{thm trans}
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   116
  in
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   117
    simplify_code_eq ctxt code_cert
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   118
  end
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   119
47156
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   120
fun define_code_cert code_eqn_thm_name def_thm rsp_thm (rty, qty) lthy = 
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   121
  let
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   122
    val quot_thm = Quotient_Term.prove_quot_theorem lthy (get_body_types (rty, qty))
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   123
  in
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   124
    if Quotient_Type.can_generate_code_cert quot_thm then
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   125
      let
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   126
        val code_cert = generate_code_cert lthy def_thm rsp_thm (rty, qty)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   127
        val add_abs_eqn_attribute = 
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   128
          Thm.declaration_attribute (fn thm => Context.mapping (Code.add_abs_eqn thm) I)
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   129
        val add_abs_eqn_attrib = Attrib.internal (K add_abs_eqn_attribute);
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   130
      in
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   131
        lthy
47156
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   132
          |> (snd oo Local_Theory.note) ((code_eqn_thm_name, [add_abs_eqn_attrib]), [code_cert])
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   133
      end
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   134
    else
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   135
      lthy
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   136
  end
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   137
47156
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   138
fun define_code_eq code_eqn_thm_name def_thm lthy =
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   139
  let
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   140
    val unfolded_def = Conv.fconv_rule unfold_fun_maps def_thm
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   141
    val code_eq = unabs_all_def lthy unfolded_def
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   142
    val simp_code_eq = simplify_code_eq lthy code_eq
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   143
  in
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   144
    lthy
47156
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   145
      |> (snd oo Local_Theory.note) ((code_eqn_thm_name, [Code.add_default_eqn_attrib]), [simp_code_eq])
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   146
  end
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   147
47156
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   148
fun define_code code_eqn_thm_name def_thm rsp_thm (rty, qty) lthy =
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   149
  if body_type rty = body_type qty then 
47156
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   150
    define_code_eq code_eqn_thm_name def_thm lthy
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   151
  else 
47156
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   152
    define_code_cert code_eqn_thm_name def_thm rsp_thm (rty, qty) lthy
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   153
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
(* The ML-interface for a quotient definition takes
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
   as argument:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
    - an optional binding and mixfix annotation
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
    - attributes
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
    - the new constant as term
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
    - 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
   161
    - respectfulness theorem for the rhs
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   162
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 43663
diff changeset
   163
   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
   164
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   165
   Restriction: At the moment the left- and right-hand
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   166
   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
   167
*)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   168
fun error_msg bind str =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   169
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   170
    val name = Binding.name_of bind
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   171
    val pos = Position.str_of (Binding.pos_of bind)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   172
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   173
    error ("Head of quotient_definition " ^
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   174
      quote str ^ " differs from declaration " ^ name ^ pos)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   175
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   176
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   177
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
   178
  let
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   179
    val rty = fastype_of rhs
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   180
    val qty = fastype_of lhs
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   181
    val absrep_trm = 
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   182
      Quotient_Term.absrep_fun lthy Quotient_Term.AbsF (rty, qty) $ rhs
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   183
    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
   184
    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
   185
    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
   186
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   187
    val ((trm, (_ , def_thm)), lthy') =
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   188
      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
   189
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   190
    (* data storage *)
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   191
    val qconst_data = {qconst = trm, rconst = rhs, def = def_thm}
47156
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   192
    val lhs_name = #1 var
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   193
    val rsp_thm_name = Binding.suffix_name "_rsp" lhs_name
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   194
    val code_eqn_thm_name = Binding.suffix_name "_code_eqn" lhs_name
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   195
    
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   196
    val lthy'' = lthy'
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   197
      |> 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
   198
        (fn phi =>
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   199
          (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
   200
            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
   201
              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
   202
          | _ => I))
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   203
      |> (snd oo Local_Theory.note) 
47156
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   204
        ((rsp_thm_name, [Attrib.internal (K Quotient_Info.rsp_rules_add)]),
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   205
        [rsp_thm])
47156
861f53bd95fe note a code eqn in quotient_def
kuncar
parents: 47096
diff changeset
   206
      |> define_code code_eqn_thm_name def_thm rsp_thm (rty, qty)
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   207
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   208
  in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   209
    (qconst_data, lthy'')
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   210
  end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   211
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   212
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
   213
  let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   214
    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
   215
    
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   216
    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
   217
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   218
        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
   219
        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
   220
          case (Thm.term_of ctm') of
47181
b351ad77eb78 some tuning while reviewing the current state of the quotient_def package
bulwahn
parents: 47156
diff changeset
   221
            Const (@{const_name HOL.eq}, _) $ _ $ _ => Conv.all_conv ctm'
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   222
            | _ => (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
   223
              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
   224
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   225
        (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
   226
      end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   227
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   228
    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
   229
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   230
        val unfold_conv = Conv.rewrs_conv 
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   231
          [@{thm fun_rel_eq_invariant[THEN eq_reflection]}, @{thm fun_rel_eq_rel[THEN eq_reflection]}, 
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   232
            @{thm fun_rel_def[THEN eq_reflection]}]
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   233
        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
   234
        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
   235
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   236
        case (Thm.term_of ctm) of
47181
b351ad77eb78 some tuning while reviewing the current state of the quotient_def package
bulwahn
parents: 47156
diff changeset
   237
          Const (@{const_name fun_rel}, _) $ _ $ _ => 
b351ad77eb78 some tuning while reviewing the current state of the quotient_def package
bulwahn
parents: 47156
diff changeset
   238
            (binop_conv2 left_conv simp_arrows_conv then_conv unfold_conv) ctm
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   239
          | _ => 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
   240
      end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   241
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   242
    val unfold_ret_val_invs = Conv.bottom_conv 
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   243
      (K (Conv.try_conv (Conv.rewr_conv @{thm invariant_same_args}))) lthy 
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   244
    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
   245
    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
   246
    val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) lthy
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   247
    val beta_conv = Thm.beta_conversion true
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   248
    val eq_thm = 
47096
3ea48c19673e generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents: 47091
diff changeset
   249
      (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs) ctm
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   250
  in
47181
b351ad77eb78 some tuning while reviewing the current state of the quotient_def package
bulwahn
parents: 47156
diff changeset
   251
    Object_Logic.rulify (eq_thm RS Drule.equal_elim_rule2)
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   252
  end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   253
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   254
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   255
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   256
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
   257
  let
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   258
    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
   259
    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
   260
    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
   261
    val rhs = prep_term NONE ctxt rhs_raw
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   262
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   263
    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
   264
    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
   265
    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
   266
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   267
    val var =
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   268
      (case vars of 
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   269
        [] => (Binding.name lhs_str, NoSyn)
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   270
      | [(binding, _, mx)] =>
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   271
          if Variable.check_name binding = lhs_str then (binding, mx)
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   272
          else error_msg binding lhs_str
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   273
      | _ => raise Match)
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   274
    
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   275
    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
   276
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   277
        val lhs_eq =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   278
          thm
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   279
          |> prop_of
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   280
          |> Logic.dest_implies
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   281
          |> fst
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   282
          |> strip_all_body
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   283
          |> 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
   284
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   285
        case lhs_eq of
47181
b351ad77eb78 some tuning while reviewing the current state of the quotient_def package
bulwahn
parents: 47156
diff changeset
   286
          SOME (Const (@{const_name HOL.eq}, _) $ _ $ _) => SOME (@{thm refl} RS thm)
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   287
          | 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
   288
            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
   289
              (#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
   290
                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
   291
              handle _ => NONE)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   292
            | _ => NONE
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   293
            )
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   294
          | _ => NONE
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   295
      end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   296
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   297
    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
   298
    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
   299
    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
   300
    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
   301
    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
   302
  
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   303
    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
   304
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   305
        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
   306
          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
   307
            [] => 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
   308
          | [[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
   309
            (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
   310
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   311
        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
   312
      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
   313
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   314
  in
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   315
    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
   316
      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
   317
      | NONE =>  Proof.theorem NONE after_qed [[(readable_rsp_tm,[])]] lthy
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   318
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   319
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
   320
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
   321
  Syntax.check_term ctxt o (case cnstr of SOME T => Type.constraint T | _ => I)
45598
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
   322
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
   323
fun read_term' cnstr ctxt =
45929
d7d6c8cfb6f6 removing some debug output in quotient_definition
bulwahn
parents: 45826
diff changeset
   324
  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
   325
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
   326
val quotient_def = gen_quotient_def Proof_Context.cert_vars check_term'
45598
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
   327
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
   328
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   329
37532
8a9a34be09e0 slight cleaning and simplification of the automatic wrapper for quotient definitions
Christian Urban <urbanc@in.tum.de>
parents: 37530
diff changeset
   330
(* 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
   331
fun lift_raw_const qtys (qconst_name, rconst, mx) ctxt =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   332
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   333
    val rty = fastype_of rconst
41451
892e67be8304 do not open ML structures;
wenzelm
parents: 41444
diff changeset
   334
    val qty = Quotient_Term.derive_qtyp ctxt qtys rty
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   335
    val lhs = Free (qconst_name, qty)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   336
  in
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   337
    (*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
   338
    ({qconst = lhs, rconst = lhs, def = @{thm refl}}, ctxt)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   339
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   340
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   341
(* parser and command *)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   342
val quotdef_parser =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   343
  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
   344
    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
   345
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   346
val _ =
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   347
  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
   348
    "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
   349
      (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
   350
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   351
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   352
end; (* structure *)