src/HOL/Tools/Quotient/quotient_def.ML
author wenzelm
Wed, 30 Mar 2016 23:46:44 +0200
changeset 62775 b486f512a471
parent 60752 b48830b670a1
child 62953 48d935524988
permissions -rw-r--r--
proper object-logic constraint (amending dd2914250ca7);
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
45598
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
    12
  val quotient_def:
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
    13
    (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
    14
    local_theory -> Proof.state
45598
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
    15
  val quotient_def_cmd:
87a2624f57e4 more uniform signature;
wenzelm
parents: 45292
diff changeset
    16
    (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
    17
    local_theory -> Proof.state
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
end;
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
structure Quotient_Def: QUOTIENT_DEF =
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
struct
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
(** Interface and Syntax Setup **)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
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
    25
(* 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
    26
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47504
diff changeset
    27
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
    28
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47504
diff changeset
    29
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
    30
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
(* The ML-interface for a quotient definition takes
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
   as argument:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
    - an optional binding and mixfix annotation
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
    - attributes
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
    - the new constant as term
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
    - 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
    38
    - respectfulness theorem for the rhs
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 43663
diff changeset
    40
   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
    41
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    42
   Restriction: At the moment the left- and right-hand
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    43
   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
    44
*)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    45
fun error_msg bind str =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    46
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    47
    val name = Binding.name_of bind
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 47938
diff changeset
    48
    val pos = Position.here (Binding.pos_of bind)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    49
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    50
    error ("Head of quotient_definition " ^
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    51
      quote str ^ " differs from declaration " ^ name ^ pos)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
    52
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    54
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
    55
  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
    56
    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
    57
    val qty = fastype_of lhs
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
    58
    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
    59
      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
    60
    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
    61
    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
    62
    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
    63
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
    64
    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
    65
      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
    66
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    67
    (* 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
    68
    val qconst_data = {qconst = trm, rconst = rhs, def = def_thm}
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
    69
47198
cfd8ff62eab1 use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents: 47181
diff changeset
    70
    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
    71
      |> Binding.qualify true defname
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
    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
    74
    val rsp_thm_name = qualify lhs_name "rsp"
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
    75
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    76
    val lthy'' = lthy'
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    77
      |> 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
    78
        (fn phi =>
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    79
          (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
    80
            qcinfo as {qconst = Const (c, _), ...} =>
59157
949829bae42a just one data slot per program unit;
wenzelm
parents: 57960
diff changeset
    81
              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
    82
          | _ => I))
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
    83
      |> (snd oo Local_Theory.note)
57960
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 56524
diff changeset
    84
        ((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
    85
  in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    86
    (qconst_data, lthy'')
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    87
  end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    88
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    89
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
    90
  let
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
    91
    val ctm = Thm.cterm_of lthy tm
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
    92
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
    93
    fun norm_fun_eq ctm =
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    94
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
    95
        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
    96
        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
    97
          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
    98
            Const (@{const_name HOL.eq}, _) $ _ $ _ => Conv.all_conv ctm'
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
    99
            | _ => (Conv.binder_conv (K erase_quants) lthy then_conv
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   100
              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
   101
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   102
        (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
   103
      end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   104
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   105
    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
   106
      let
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   107
        val unfold_conv = Conv.rewrs_conv
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   108
          [@{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
   109
            @{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
   110
        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
   111
        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
   112
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   113
        case (Thm.term_of ctm) of
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   114
          Const (@{const_name rel_fun}, _) $ _ $ _ =>
47181
b351ad77eb78 some tuning while reviewing the current state of the quotient_def package
bulwahn
parents: 47156
diff changeset
   115
            (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
   116
          | _ => 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
   117
      end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   118
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   119
    val unfold_ret_val_invs = Conv.bottom_conv
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   120
      (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
   121
    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
   122
    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
   123
    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
   124
    val beta_conv = Thm.beta_conversion true
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   125
    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
   126
      (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
   127
  in
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 50902
diff changeset
   128
    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
   129
  end
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   130
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   131
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   132
fun gen_quotient_def prep_var parse_term (raw_var, (attr, (raw_lhs, raw_rhs))) lthy =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   133
  let
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   134
    val (opt_var, ctxt) =
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   135
      (case raw_var of
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   136
        NONE => (NONE, lthy)
62775
b486f512a471 proper object-logic constraint (amending dd2914250ca7);
wenzelm
parents: 60752
diff changeset
   137
      | SOME var =>
b486f512a471 proper object-logic constraint (amending dd2914250ca7);
wenzelm
parents: 60752
diff changeset
   138
          Proof_Context.set_object_logic_constraint lthy
b486f512a471 proper object-logic constraint (amending dd2914250ca7);
wenzelm
parents: 60752
diff changeset
   139
          |> prep_var var
b486f512a471 proper object-logic constraint (amending dd2914250ca7);
wenzelm
parents: 60752
diff changeset
   140
          ||> Proof_Context.restore_object_logic_constraint lthy
b486f512a471 proper object-logic constraint (amending dd2914250ca7);
wenzelm
parents: 60752
diff changeset
   141
          |>> SOME)
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   142
    val lhs_constraint = (case opt_var of SOME (_, SOME T, _) => T | _ => dummyT)
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   143
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   144
    fun prep_term T = parse_term ctxt #> Type.constraint T #> Syntax.check_term ctxt;
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   145
    val lhs = prep_term lhs_constraint raw_lhs
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   146
    val rhs = prep_term dummyT raw_rhs
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   147
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   148
    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
   149
    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
   150
    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
   151
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   152
    val var =
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   153
      (case opt_var of
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   154
        NONE => (Binding.name lhs_str, NoSyn)
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   155
      | SOME (binding, _, mx) =>
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   156
          if Variable.check_name binding = lhs_str then (binding, mx)
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   157
          else error_msg binding lhs_str);
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   158
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   159
    fun try_to_prove_refl thm =
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   160
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   161
        val lhs_eq =
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   162
          thm
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   163
          |> 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
   164
          |> Logic.dest_implies
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   165
          |> fst
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   166
          |> strip_all_body
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   167
          |> 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
   168
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   169
        case lhs_eq of
47181
b351ad77eb78 some tuning while reviewing the current state of the quotient_def package
bulwahn
parents: 47156
diff changeset
   170
          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
   171
          | 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
   172
            Type (typ_name, _) =>
cb2b940e2fdf avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
wenzelm
parents: 48992
diff changeset
   173
              try (fn () =>
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   174
                #equiv_thm (the (Quotient_Info.lookup_quotients lthy typ_name))
50902
cb2b940e2fdf avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
wenzelm
parents: 48992
diff changeset
   175
                  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
   176
            | _ => NONE
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   177
            )
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   178
          | _ => NONE
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   179
      end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   181
    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
   182
    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
   183
    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
   184
    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
   185
    val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq)
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   186
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   187
    fun after_qed thm_list lthy =
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   188
      let
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   189
        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
   190
          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
   191
            [] => the maybe_proven_rsp_thm
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   192
          | [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
   193
            (fn _ =>
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
   194
              resolve_tac ctxt [readable_rsp_thm_eq] 1 THEN
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
   195
              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
   196
      in
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   197
        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
   198
      end
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   199
  in
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   200
    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
   201
      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
   202
      | NONE =>  Proof.theorem NONE after_qed [[(readable_rsp_tm,[])]] lthy
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 38624
diff changeset
   203
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   204
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   205
val quotient_def = gen_quotient_def Proof_Context.cert_var (K I)
51d9dcd71ad7 tuned signature;
wenzelm
parents: 59936
diff changeset
   206
val quotient_def_cmd = gen_quotient_def Proof_Context.read_var Syntax.parse_term
43608
294570668f25 parse term in auxiliary context augmented with variable;
krauss
parents: 42494
diff changeset
   207
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   208
60669
wenzelm
parents: 60379
diff changeset
   209
(* command syntax *)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   210
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   211
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59621
diff changeset
   212
  Outer_Syntax.local_theory_to_proof @{command_keyword quotient_definition}
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   213
    "definition for constants over the quotient type"
60669
wenzelm
parents: 60379
diff changeset
   214
      (Scan.option Parse_Spec.constdecl --
wenzelm
parents: 60379
diff changeset
   215
        Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term -- (@{keyword "is"} |-- Parse.term)))
wenzelm
parents: 60379
diff changeset
   216
        >> quotient_def_cmd);
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   217
60669
wenzelm
parents: 60379
diff changeset
   218
end;