src/HOL/Tools/Quotient/quotient_def.ML
author wenzelm
Wed, 04 Mar 2015 19:53:18 +0100
changeset 59582 0fbed69ff081
parent 59580 cbc38731d42f
child 59621 291934bac95e
permissions -rw-r--r--
tuned signature -- prefer qualified names;
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
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
end;
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
structure Quotient_Def: QUOTIENT_DEF =
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
struct
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
(** Interface and Syntax Setup **)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
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
    28
(* 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
    29
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47504
diff changeset
    30
open Lifting_Util
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
    31
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47504
diff changeset
    32
infix 0 MRSL
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
    33
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
(* The ML-interface for a quotient definition takes
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
   as argument:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
    - an optional binding and mixfix annotation
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
    - attributes
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
    - the new constant as term
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
    - 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
    41
    - respectfulness theorem for the rhs
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 43663
diff changeset
    43
   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
    44
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    45
   Restriction: At the moment the left- and right-hand
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    46
   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
    47
*)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    48
fun error_msg bind str =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    49
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    50
    val name = Binding.name_of bind
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 47938
diff changeset
    51
    val pos = Position.here (Binding.pos_of bind)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    52
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    53
    error ("Head of quotient_definition " ^
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    54
      quote str ^ " differs from declaration " ^ name ^ pos)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    55
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    57
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
    58
  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
    59
    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
    60
    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
    61
    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
    62
      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
    63
    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
    64
    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
    65
    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
    66
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
    67
    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
    68
      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
    69
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    70
    (* 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
    71
    val qconst_data = {qconst = trm, rconst = rhs, def = def_thm}
47198
cfd8ff62eab1 use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents: 47181
diff changeset
    72
     
cfd8ff62eab1 use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents: 47181
diff changeset
    73
    fun qualify defname suffix = Binding.name suffix
cfd8ff62eab1 use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents: 47181
diff changeset
    74
      |> Binding.qualify true defname
cfd8ff62eab1 use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents: 47181
diff changeset
    75
cfd8ff62eab1 use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents: 47181
diff changeset
    76
    val lhs_name = Binding.name_of (#1 var)
cfd8ff62eab1 use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents: 47181
diff changeset
    77
    val rsp_thm_name = qualify lhs_name "rsp"
47091
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
    val lthy'' = lthy'
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    80
      |> 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
    81
        (fn phi =>
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    82
          (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
    83
            qcinfo as {qconst = Const (c, _), ...} =>
59157
949829bae42a just one data slot per program unit;
wenzelm
parents: 57960
diff changeset
    84
              Quotient_Info.update_quotconsts (c, qcinfo)
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    85
          | _ => I))
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    86
      |> (snd oo Local_Theory.note) 
57960
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 56524
diff changeset
    87
        ((rsp_thm_name, @{attributes [quot_respect]}), [rsp_thm])
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    88
  in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    89
    (qconst_data, lthy'')
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    90
  end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    91
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    92
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
    93
  let
59580
cbc38731d42f eliminated some clones of Proof_Context.cterm_of
traytel
parents: 59157
diff changeset
    94
    val ctm = Proof_Context.cterm_of lthy tm
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    95
    
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    96
    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
    97
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    98
        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
    99
        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
   100
          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
   101
            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
   102
            | _ => (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
   103
              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
   104
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   105
        (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
   106
      end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   107
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   108
    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
   109
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   110
        val unfold_conv = Conv.rewrs_conv 
56519
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 55945
diff changeset
   111
          [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, @{thm rel_fun_eq_rel[THEN eq_reflection]}, 
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 54742
diff changeset
   112
            @{thm rel_fun_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
   113
        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
   114
        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
   115
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   116
        case (Thm.term_of ctm) of
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 54742
diff changeset
   117
          Const (@{const_name rel_fun}, _) $ _ $ _ => 
47181
b351ad77eb78 some tuning while reviewing the current state of the quotient_def package
bulwahn
parents: 47156
diff changeset
   118
            (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
   119
          | _ => 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
   120
      end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   121
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
   122
    val unfold_ret_val_invs = Conv.bottom_conv 
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56519
diff changeset
   123
      (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) lthy 
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   124
    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
   125
    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
   126
    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
   127
    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
   128
    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
   129
      (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
   130
  in
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 50902
diff changeset
   131
    Object_Logic.rulify lthy (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
   132
  end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   133
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   134
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   135
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   136
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
   137
  let
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   138
    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
   139
    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
   140
    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
   141
    val rhs = prep_term NONE ctxt rhs_raw
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   142
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   143
    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
   144
    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
   145
    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
   146
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   147
    val var =
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   148
      (case vars of 
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   149
        [] => (Binding.name lhs_str, NoSyn)
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   150
      | [(binding, _, mx)] =>
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   151
          if Variable.check_name binding = lhs_str then (binding, mx)
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   152
          else error_msg binding lhs_str
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   153
      | _ => raise Match)
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   154
    
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   155
    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
   156
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   157
        val lhs_eq =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   158
          thm
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   159
          |> Thm.prop_of
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   160
          |> Logic.dest_implies
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   161
          |> fst
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   162
          |> strip_all_body
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   163
          |> 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
   164
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   165
        case lhs_eq of
47181
b351ad77eb78 some tuning while reviewing the current state of the quotient_def package
bulwahn
parents: 47156
diff changeset
   166
          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
   167
          | SOME _ => (case body_type (fastype_of lhs) of
50902
cb2b940e2fdf avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
wenzelm
parents: 48992
diff changeset
   168
            Type (typ_name, _) =>
cb2b940e2fdf avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
wenzelm
parents: 48992
diff changeset
   169
              try (fn () =>
cb2b940e2fdf avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
wenzelm
parents: 48992
diff changeset
   170
                #equiv_thm (the (Quotient_Info.lookup_quotients lthy typ_name)) 
cb2b940e2fdf avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
wenzelm
parents: 48992
diff changeset
   171
                  RS @{thm Equiv_Relations.equivp_reflp} RS thm) ()
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   172
            | _ => NONE
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   173
            )
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   174
          | _ => NONE
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
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
    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
   178
    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
   179
    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
   180
    val maybe_proven_rsp_thm = try_to_prove_refl readable_rsp_thm_eq
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   181
    val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq)
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   182
  
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   183
    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
   184
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   185
        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
   186
          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
   187
            [] => 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
   188
          | [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm 
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 50902
diff changeset
   189
            (fn _ => rtac readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac ctxt [thm] 1)
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   190
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   191
        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
   192
      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
   193
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   194
  in
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   195
    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
   196
      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
   197
      | NONE =>  Proof.theorem NONE after_qed [[(readable_rsp_tm,[])]] lthy
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   198
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   199
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
   200
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
   201
  Syntax.check_term ctxt o (case cnstr of SOME T => Type.constraint T | _ => I)
45598
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
   202
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
   203
fun read_term' cnstr ctxt =
45929
d7d6c8cfb6f6 removing some debug output in quotient_definition
bulwahn
parents: 45826
diff changeset
   204
  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
   205
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
   206
val quotient_def = gen_quotient_def Proof_Context.cert_vars check_term'
45598
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
   207
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
   208
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   209
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   210
(* parser and command *)
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   211
val quotdef_parser =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   212
  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
   213
    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
   214
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   215
val _ =
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   216
  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
   217
    "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
   218
      (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
   219
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   220
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   221
end; (* structure *)