src/HOL/Tools/Quotient/quotient_type.ML
author wenzelm
Thu, 24 Sep 2015 13:33:42 +0200
changeset 61260 e6f03fae14d5
parent 61110 6b7c2ecc6aea
child 62969 9f394a16c557
permissions -rw-r--r--
explicit indication of overloaded typedefs;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45680
a61510361b89 more conventional file name;
wenzelm
parents: 45676
diff changeset
     1
(*  Title:      HOL/Tools/Quotient/quotient_type.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
35806
a814cccce0b8 rollback of local typedef until problem with type-variables can be sorted out; fixed header
Christian Urban <urbanc@in.tum.de>
parents: 35790
diff changeset
     4
Definition of a quotient type.
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_TYPE =
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
sig
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
     9
  val add_quotient_type: {overloaded: bool} ->
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    10
    ((string list * binding * mixfix) * (typ * term * bool) *
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    11
      ((binding * binding) option * thm option)) * thm -> local_theory ->
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    12
      Quotient_Info.quotients * local_theory
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    13
  val quotient_type:  {overloaded: bool} ->
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    14
    (string list * binding * mixfix) * (typ * term * bool) *
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    15
      ((binding * binding) option * thm option) -> Proof.context -> Proof.state
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    16
  val quotient_type_cmd:  {overloaded: bool} ->
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    17
    (((((string list * binding) * mixfix) * string) * (bool * string)) *
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    18
      (binding * binding) option) * (Facts.ref * Token.src list) option ->
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    19
      Proof.context -> Proof.state
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
end;
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
structure Quotient_Type: QUOTIENT_TYPE =
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
struct
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
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
(*** definition of quotient types ***)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
60669
wenzelm
parents: 60225
diff changeset
    28
val mem_def1 = @{lemma "y \<in> Collect S \<Longrightarrow> S y" by simp}
wenzelm
parents: 60225
diff changeset
    29
val mem_def2 = @{lemma "S y \<Longrightarrow> y \<in> Collect S" by simp}
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
47100
f8f788c8b7f3 updated comment
kuncar
parents: 47096
diff changeset
    31
(* constructs the term {c. EX (x::rty). rel x x \<and> c = Collect (rel x)} *)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
fun typedef_term rel rty lthy =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    33
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    34
    val [x, c] =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    35
      [("x", rty), ("c", HOLogic.mk_setT rty)]
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    36
      |> Variable.variant_frees lthy [rel]
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    37
      |> map Free
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    38
  in
45312
bulwahn
parents: 45291
diff changeset
    39
    HOLogic.Collect_const (HOLogic.mk_setT rty) $ (lambda c (HOLogic.exists_const rty $
44204
3cdc4176638c Quotient Package: make quotient_type work with separate set type
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 43547
diff changeset
    40
        lambda x (HOLogic.mk_conj (rel $ x $ x,
45312
bulwahn
parents: 45291
diff changeset
    41
        HOLogic.mk_eq (c, HOLogic.Collect_const rty $ (rel $ x))))))
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    42
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
(* makes the new type definitions and proves non-emptyness *)
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    46
fun typedef_make overloaded (vs, qty_name, mx, rel, rty) equiv_thm lthy =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    47
  let
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
    48
    fun typedef_tac ctxt =
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
    49
      EVERY1 (map (resolve_tac ctxt o single) [@{thm part_equivp_typedef}, equiv_thm])
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    50
  in
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
    51
    Typedef.add_typedef overloaded (qty_name, map (rpair dummyS) vs, mx)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    52
      (typedef_term rel rty lthy) NONE typedef_tac lthy
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    53
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
(* tactic to prove the quot_type theorem for the new type *)
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58959
diff changeset
    57
fun typedef_quot_type_tac ctxt equiv_thm ((_, typedef_info): Typedef.info) =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    58
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    59
    val rep_thm = #Rep typedef_info RS mem_def1
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    60
    val rep_inv = #Rep_inverse typedef_info
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    61
    val abs_inv = #Abs_inverse typedef_info
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    62
    val rep_inj = #Rep_inject typedef_info
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    63
  in
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
    64
    (resolve_tac ctxt @{thms quot_type.intro} THEN' RANGE [
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
    65
      resolve_tac ctxt [equiv_thm],
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
    66
      resolve_tac ctxt [rep_thm],
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
    67
      resolve_tac ctxt [rep_inv],
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
    68
      resolve_tac ctxt [abs_inv] THEN' resolve_tac ctxt [mem_def2] THEN' assume_tac ctxt,
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60669
diff changeset
    69
      resolve_tac ctxt [rep_inj]]) 1
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    70
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
(* proves the quot_type theorem for the new type *)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
fun typedef_quot_type_thm (rel, abs, rep, equiv_thm, typedef_info) lthy =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    74
  let
45317
bf8b9ac6000c more robust, declarative and unsurprising computation of types in the quotient type definition
bulwahn
parents: 45314
diff changeset
    75
    val quot_type_const = Const (@{const_name "quot_type"},
bf8b9ac6000c more robust, declarative and unsurprising computation of types in the quotient type definition
bulwahn
parents: 45314
diff changeset
    76
      fastype_of rel --> fastype_of abs --> fastype_of rep --> @{typ bool})
bf8b9ac6000c more robust, declarative and unsurprising computation of types in the quotient type definition
bulwahn
parents: 45314
diff changeset
    77
    val goal = HOLogic.mk_Trueprop (quot_type_const $ rel $ abs $ rep)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    78
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    79
    Goal.prove lthy [] [] goal
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58959
diff changeset
    80
      (fn {context = ctxt, ...} => typedef_quot_type_tac ctxt equiv_thm typedef_info)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
    81
  end
60669
wenzelm
parents: 60225
diff changeset
    82
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47566
diff changeset
    83
open Lifting_Util
47362
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    84
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47566
diff changeset
    85
infix 0 MRSL
47362
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    86
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    87
fun define_cr_rel equiv_thm abs_fun lthy =
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    88
  let
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    89
    fun force_type_of_rel rel forced_ty =
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    90
      let
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    91
        val thy = Proof_Context.theory_of lthy
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    92
        val rel_ty = (domain_type o fastype_of) rel
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    93
        val ty_inst = Sign.typ_match thy (rel_ty, forced_ty) Vartab.empty
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    94
      in
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    95
        Envir.subst_term_types ty_inst rel
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    96
      end
60669
wenzelm
parents: 60225
diff changeset
    97
47362
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    98
    val (rty, qty) = (dest_funT o fastype_of) abs_fun
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
    99
    val abs_fun_graph = HOLogic.mk_eq(abs_fun $ Bound 1, Bound 0)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59487
diff changeset
   100
    val Abs_body = (case (HOLogic.dest_Trueprop o Thm.prop_of) equiv_thm of
47362
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   101
      Const (@{const_name equivp}, _) $ _ => abs_fun_graph
60669
wenzelm
parents: 60225
diff changeset
   102
      | Const (@{const_name part_equivp}, _) $ rel =>
47362
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   103
        HOLogic.mk_conj (force_type_of_rel rel rty $ Bound 1 $ Bound 1, abs_fun_graph)
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   104
      | _ => error "unsupported equivalence theorem"
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   105
      )
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   106
    val def_term = Abs ("x", rty, Abs ("y", qty, Abs_body));
50177
2006c50172c9 generate correct names
kuncar
parents: 49835
diff changeset
   107
    val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty
2006c50172c9 generate correct names
kuncar
parents: 49835
diff changeset
   108
    val cr_rel_name = Binding.prefix_name "cr_" qty_name
47362
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   109
    val (fixed_def_term, lthy') = yield_singleton (Variable.importT_terms) def_term lthy
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   110
    val ((_, (_ , def_thm)), lthy'') =
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   111
      Local_Theory.define ((cr_rel_name, NoSyn), ((Thm.def_binding cr_rel_name, []), fixed_def_term)) lthy'
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   112
  in
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   113
    (def_thm, lthy'')
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   114
  end;
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   115
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59157
diff changeset
   116
fun setup_lifting_package quot3_thm equiv_thm opt_par_thm lthy =
47362
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   117
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59487
diff changeset
   118
    val (_ $ _ $ abs_fun $ _) = (HOLogic.dest_Trueprop o Thm.prop_of) quot3_thm
47362
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   119
    val (T_def, lthy') = define_cr_rel equiv_thm abs_fun lthy
47502
4c949049cd78 note the Quotient theorem in quotient_type
kuncar
parents: 47362
diff changeset
   120
    val (rty, qty) = (dest_funT o fastype_of) abs_fun
50177
2006c50172c9 generate correct names
kuncar
parents: 49835
diff changeset
   121
    val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty
2006c50172c9 generate correct names
kuncar
parents: 49835
diff changeset
   122
    val quotient_thm_name = Binding.prefix_name "Quotient_" qty_name
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59487
diff changeset
   123
    val (reflp_thm, quot_thm) =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59487
diff changeset
   124
      (case (HOLogic.dest_Trueprop o Thm.prop_of) equiv_thm of
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59487
diff changeset
   125
        Const (@{const_name equivp}, _) $ _ =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59487
diff changeset
   126
          (SOME (equiv_thm RS @{thm equivp_reflp2}),
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59487
diff changeset
   127
            [quot3_thm, T_def, equiv_thm] MRSL @{thm Quotient3_to_Quotient_equivp})
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47502
diff changeset
   128
      | Const (@{const_name part_equivp}, _) $ _ =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59487
diff changeset
   129
          (NONE, [quot3_thm, T_def] MRSL @{thm Quotient3_to_Quotient})
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59487
diff changeset
   130
      | _ => error "unsupported equivalence theorem")
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   131
    val config = { notes = true }
47362
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   132
  in
47502
4c949049cd78 note the Quotient theorem in quotient_type
kuncar
parents: 47362
diff changeset
   133
    lthy'
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   134
      |> Lifting_Setup.setup_by_quotient config quot_thm reflp_thm opt_par_thm
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 59936
diff changeset
   135
      |> snd
47502
4c949049cd78 note the Quotient theorem in quotient_type
kuncar
parents: 47362
diff changeset
   136
      |> (snd oo Local_Theory.note) ((quotient_thm_name, []), [quot_thm])
47362
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   137
  end
b1f099bdfbba connect the Quotient package to the Lifting package
kuncar
parents: 47308
diff changeset
   138
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59157
diff changeset
   139
fun init_quotient_infr quot_thm equiv_thm opt_par_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: 47093
diff changeset
   140
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59487
diff changeset
   141
    val (_ $ rel $ abs $ rep) = (HOLogic.dest_Trueprop o Thm.prop_of) quot_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: 47093
diff changeset
   142
    val (qtyp, rtyp) = (dest_funT o fastype_of) rep
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: 47093
diff changeset
   143
    val qty_full_name = (fst o dest_Type) qtyp
60669
wenzelm
parents: 60225
diff changeset
   144
    val quotients = {qtyp = qtyp, rtyp = rtyp, equiv_rel = rel, equiv_thm = equiv_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: 47093
diff changeset
   145
      quot_thm = 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: 47093
diff changeset
   146
    fun quot_info phi = Quotient_Info.transform_quotients phi quotients
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: 47093
diff changeset
   147
    val abs_rep = {abs = abs, rep = rep}
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: 47093
diff changeset
   148
    fun abs_rep_info phi = Quotient_Info.transform_abs_rep phi abs_rep
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: 47093
diff changeset
   149
  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: 47093
diff changeset
   150
    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: 47093
diff changeset
   151
      |> Local_Theory.declaration {syntax = false, pervasive = true}
59157
949829bae42a just one data slot per program unit;
wenzelm
parents: 58963
diff changeset
   152
        (fn phi => Quotient_Info.update_quotients (qty_full_name, quot_info phi)
949829bae42a just one data slot per program unit;
wenzelm
parents: 58963
diff changeset
   153
          #> Quotient_Info.update_abs_rep (qty_full_name, abs_rep_info phi))
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59157
diff changeset
   154
      |> setup_lifting_package quot_thm equiv_thm opt_par_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: 47093
diff changeset
   155
  end
35222
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
(* main function for constructing a quotient type *)
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
   158
fun add_quotient_type overloaded
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
   159
    (((vs, qty_name, mx), (rty, rel, partial), (opt_morphs, opt_par_thm)), equiv_thm) lthy =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   160
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   161
    val part_equiv =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   162
      if partial
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   163
      then equiv_thm
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   164
      else equiv_thm RS @{thm equivp_implies_part_equivp}
37493
2377d246a631 Quotient package now uses Partial Equivalence instead place of equivalence
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36960
diff changeset
   165
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   166
    (* generates the typedef *)
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: 47093
diff changeset
   167
    val ((_, typedef_info), lthy1) =
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
   168
      typedef_make overloaded (vs, qty_name, mx, rel, rty) part_equiv lthy
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   169
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   170
    (* abs and rep functions from the typedef *)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   171
    val Abs_ty = #abs_type (#1 typedef_info)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   172
    val Rep_ty = #rep_type (#1 typedef_info)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   173
    val Abs_name = #Abs_name (#1 typedef_info)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   174
    val Rep_name = #Rep_name (#1 typedef_info)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   175
    val Abs_const = Const (Abs_name, Rep_ty --> Abs_ty)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   176
    val Rep_const = Const (Rep_name, Abs_ty --> Rep_ty)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   178
    (* more useful abs and rep definitions *)
45317
bf8b9ac6000c more robust, declarative and unsurprising computation of types in the quotient type definition
bulwahn
parents: 45314
diff changeset
   179
    val abs_const = Const (@{const_name quot_type.abs},
bf8b9ac6000c more robust, declarative and unsurprising computation of types in the quotient type definition
bulwahn
parents: 45314
diff changeset
   180
      (rty --> rty --> @{typ bool}) --> (Rep_ty --> Abs_ty) --> rty --> Abs_ty)
bf8b9ac6000c more robust, declarative and unsurprising computation of types in the quotient type definition
bulwahn
parents: 45314
diff changeset
   181
    val rep_const = Const (@{const_name quot_type.rep}, (Abs_ty --> Rep_ty) --> Abs_ty --> rty)
bf8b9ac6000c more robust, declarative and unsurprising computation of types in the quotient type definition
bulwahn
parents: 45314
diff changeset
   182
    val abs_trm = abs_const $ rel $ Abs_const
bf8b9ac6000c more robust, declarative and unsurprising computation of types in the quotient type definition
bulwahn
parents: 45314
diff changeset
   183
    val rep_trm = rep_const $ Rep_const
45676
fa46fef06590 alternative names of morphisms in the definition of a quotient type can be specified
kuncar
parents: 45534
diff changeset
   184
    val (rep_name, abs_name) =
fa46fef06590 alternative names of morphisms in the definition of a quotient type can be specified
kuncar
parents: 45534
diff changeset
   185
      (case opt_morphs of
fa46fef06590 alternative names of morphisms in the definition of a quotient type can be specified
kuncar
parents: 45534
diff changeset
   186
        NONE => (Binding.prefix_name "rep_" qty_name, Binding.prefix_name "abs_" qty_name)
fa46fef06590 alternative names of morphisms in the definition of a quotient type can be specified
kuncar
parents: 45534
diff changeset
   187
      | SOME morphs => morphs)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   188
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: 47093
diff changeset
   189
    val ((_, (_, abs_def)), lthy2) = lthy1
46909
3c73a121a387 more explicit indication of def names;
wenzelm
parents: 46727
diff changeset
   190
      |> Local_Theory.define ((abs_name, NoSyn), ((Thm.def_binding abs_name, []), abs_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: 47093
diff changeset
   191
    val ((_, (_, rep_def)), lthy3) = lthy2
46909
3c73a121a387 more explicit indication of def names;
wenzelm
parents: 46727
diff changeset
   192
      |> Local_Theory.define ((rep_name, NoSyn), ((Thm.def_binding rep_name, []), rep_trm))
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   193
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   194
    (* quot_type theorem *)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   195
    val quot_thm = typedef_quot_type_thm (rel, Abs_const, Rep_const, part_equiv, typedef_info) lthy3
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   197
    (* quotient theorem *)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47100
diff changeset
   198
    val quotient_thm_name = Binding.prefix_name "Quotient3_" qty_name
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   199
    val quotient_thm =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   200
      (quot_thm RS @{thm quot_type.Quotient})
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51374
diff changeset
   201
      |> fold_rule lthy3 [abs_def, rep_def]
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   202
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   203
    (* name equivalence theorem *)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   204
    val equiv_thm_name = Binding.suffix_name "_equivp" qty_name
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   205
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   206
    (* storing the quotients *)
60669
wenzelm
parents: 60225
diff changeset
   207
    val quotients = {qtyp = Abs_ty, rtyp = rty, equiv_rel = rel, equiv_thm = equiv_thm,
47093
0516a6c1ea59 store the quotient theorem for every quotient
kuncar
parents: 47091
diff changeset
   208
      quot_thm = quotient_thm}
37530
70d03844b2f9 export of proper information in the ML-interface of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 37493
diff changeset
   209
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   210
    val lthy4 = lthy3
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59157
diff changeset
   211
      |> init_quotient_infr quotient_thm equiv_thm opt_par_thm
45282
eaec1651709a eliminated aliases of standard functions;
wenzelm
parents: 45280
diff changeset
   212
      |> (snd oo Local_Theory.note)
eaec1651709a eliminated aliases of standard functions;
wenzelm
parents: 45280
diff changeset
   213
        ((equiv_thm_name,
57960
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 55467
diff changeset
   214
          if partial then [] else @{attributes [quot_equiv]}),
45282
eaec1651709a eliminated aliases of standard functions;
wenzelm
parents: 45280
diff changeset
   215
          [equiv_thm])
eaec1651709a eliminated aliases of standard functions;
wenzelm
parents: 45280
diff changeset
   216
      |> (snd oo Local_Theory.note)
57960
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 55467
diff changeset
   217
        ((quotient_thm_name, @{attributes [quot_thm]}), [quotient_thm])
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   218
  in
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   219
    (quotients, lthy4)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   220
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   221
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   223
(* sanity checks for the quotient type specifications *)
45676
fa46fef06590 alternative names of morphisms in the definition of a quotient type can be specified
kuncar
parents: 45534
diff changeset
   224
fun sanity_check ((vs, qty_name, _), (rty, rel, _), _) =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   225
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   226
    val rty_tfreesT = map fst (Term.add_tfreesT rty [])
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   227
    val rel_tfrees = map fst (Term.add_tfrees rel [])
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   228
    val rel_frees = map fst (Term.add_frees rel [])
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   229
    val rel_vars = Term.add_vars rel []
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   230
    val rel_tvars = Term.add_tvars rel []
43547
f3a8476285c6 clarified Binding.pretty/print: no quotes, only markup -- Binding.str_of is rendered obsolete;
wenzelm
parents: 42361
diff changeset
   231
    val qty_str = Binding.print qty_name ^ ": "
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   232
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   233
    val illegal_rel_vars =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   234
      if null rel_vars andalso null rel_tvars then []
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   235
      else [qty_str ^ "illegal schematic variable(s) in the relation."]
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   236
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   237
    val dup_vs =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   238
      (case duplicates (op =) vs of
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   239
        [] => []
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   240
      | dups => [qty_str ^ "duplicate type variable(s) on the lhs: " ^ commas_quote dups])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   241
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   242
    val extra_rty_tfrees =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   243
      (case subtract (op =) vs rty_tfreesT of
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   244
        [] => []
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   245
      | extras => [qty_str ^ "extra type variable(s) on the lhs: " ^ commas_quote extras])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   246
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   247
    val extra_rel_tfrees =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   248
      (case subtract (op =) vs rel_tfrees of
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   249
        [] => []
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   250
      | extras => [qty_str ^ "extra type variable(s) in the relation: " ^ commas_quote extras])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   251
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   252
    val illegal_rel_frees =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   253
      (case rel_frees of
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   254
        [] => []
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   255
      | xs => [qty_str ^ "illegal variable(s) in the relation: " ^ commas_quote xs])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   256
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   257
    val errs = illegal_rel_vars @ dup_vs @ extra_rty_tfrees @ extra_rel_tfrees @ illegal_rel_frees
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   258
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   259
    if null errs then () else error (cat_lines errs)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   260
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   261
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   262
(* check for existence of map functions *)
45795
2d8949268303 maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
kuncar
parents: 45698
diff changeset
   263
fun map_check ctxt (_, (rty, _, _), _) =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   264
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   265
    fun map_check_aux rty warns =
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   266
      (case rty of
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   267
        Type (_, []) => warns
45340
98ec8b51af9c prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
wenzelm
parents: 45317
diff changeset
   268
      | Type (s, _) =>
55467
a5c9002bc54d renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
blanchet
parents: 54742
diff changeset
   269
          if Symtab.defined (Functor.entries ctxt) s then warns else s :: warns
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   270
      | _ => warns)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   271
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   272
    val warns = map_check_aux rty []
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   273
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   274
    if null warns then ()
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   275
    else warning ("No map function defined for " ^ commas warns ^
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   276
      ". This will cause problems later on.")
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   277
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   278
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   279
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   280
(*** interface and syntax setup ***)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   281
45698
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   282
(* the ML-interface takes a list of tuples consisting of:
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   283
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   284
 - the name of the quotient type
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   285
 - its free type variables (first argument)
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   286
 - its mixfix annotation
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   287
 - the type to be quotient
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   288
 - the partial flag (a boolean)
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   289
 - the relation according to which the type is quotient
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   290
 - optional names of morphisms (rep/abs)
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47698
diff changeset
   291
 - flag if code should be generated by Lifting package
45698
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   292
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   293
 it opens a proof-state in which one has to show that the
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   294
 relations are equivalence relations
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   295
*)
fd8e140ae879 removed outdated comment moved back and updated (at the direct request of Christian Urban)
kuncar
parents: 45690
diff changeset
   296
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
   297
fun quotient_type overloaded quot lthy =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   298
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   299
    (* sanity check *)
47983
a5e699834f2d drop the feature that more than one quotient type can be defined by quotient_type -> it causes problems
kuncar
parents: 47938
diff changeset
   300
    val _ = sanity_check quot
a5e699834f2d drop the feature that more than one quotient type can be defined by quotient_type -> it causes problems
kuncar
parents: 47938
diff changeset
   301
    val _ = map_check lthy quot
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   302
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   303
    fun mk_goal (rty, rel, partial) =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   304
      let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   305
        val equivp_ty = ([rty, rty] ---> @{typ bool}) --> @{typ bool}
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   306
        val const =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   307
          if partial then @{const_name part_equivp} else @{const_name equivp}
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   308
      in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   309
        HOLogic.mk_Trueprop (Const (const, equivp_ty) $ rel)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   310
      end
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   311
47983
a5e699834f2d drop the feature that more than one quotient type can be defined by quotient_type -> it causes problems
kuncar
parents: 47938
diff changeset
   312
    val goal = (mk_goal o #2) quot
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   313
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
   314
    fun after_qed [[thm]] = (snd oo add_quotient_type overloaded) (quot, thm)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   315
  in
47983
a5e699834f2d drop the feature that more than one quotient type can be defined by quotient_type -> it causes problems
kuncar
parents: 47938
diff changeset
   316
    Proof.theorem NONE after_qed [[(goal, [])]] lthy
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   317
  end
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   318
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
   319
fun quotient_type_cmd overloaded spec lthy =
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   320
  let
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59157
diff changeset
   321
    fun parse_spec ((((((vs, qty_name), mx), rty_str), (partial, rel_str)), opt_morphs), opt_par_xthm) lthy =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   322
      let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   323
        val rty = Syntax.read_typ lthy rty_str
46727
0162a0d284ac Finish localizing the quotient package.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 45835
diff changeset
   324
        val tmp_lthy1 = Variable.declare_typ rty lthy
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   325
        val rel =
46727
0162a0d284ac Finish localizing the quotient package.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 45835
diff changeset
   326
          Syntax.parse_term tmp_lthy1 rel_str
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   327
          |> Type.constraint (rty --> rty --> @{typ bool})
46727
0162a0d284ac Finish localizing the quotient package.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 45835
diff changeset
   328
          |> Syntax.check_term tmp_lthy1
0162a0d284ac Finish localizing the quotient package.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 45835
diff changeset
   329
        val tmp_lthy2 = Variable.declare_term rel tmp_lthy1
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 50177
diff changeset
   330
        val opt_par_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_par_xthm
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   331
      in
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 59157
diff changeset
   332
        (((vs, qty_name, mx), (rty, rel, partial), (opt_morphs, opt_par_thm)), tmp_lthy2)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   333
      end
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   334
47983
a5e699834f2d drop the feature that more than one quotient type can be defined by quotient_type -> it causes problems
kuncar
parents: 47938
diff changeset
   335
    val (spec', _) = parse_spec spec lthy
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   336
  in
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
   337
    quotient_type overloaded spec' lthy
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   338
  end
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   339
37493
2377d246a631 Quotient package now uses Partial Equivalence instead place of equivalence
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36960
diff changeset
   340
60669
wenzelm
parents: 60225
diff changeset
   341
(* command syntax *)
47091
d5cd13aca90b respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents: 46961
diff changeset
   342
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   343
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59582
diff changeset
   344
  Outer_Syntax.local_theory_to_proof @{command_keyword quotient_type}
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 39288
diff changeset
   345
    "quotient type definitions (require equivalence proofs)"
60669
wenzelm
parents: 60225
diff changeset
   346
      (* FIXME Parse.type_args_constrained and standard treatment of sort constraints *)
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
   347
      (Parse_Spec.overloaded -- (Parse.type_args -- Parse.binding --
60669
wenzelm
parents: 60225
diff changeset
   348
        Parse.opt_mixfix -- (@{keyword "="} |-- Parse.typ) -- (@{keyword "/"} |--
wenzelm
parents: 60225
diff changeset
   349
          Scan.optional (Parse.reserved "partial" -- @{keyword ":"} >> K true) false -- Parse.term) --
wenzelm
parents: 60225
diff changeset
   350
        Scan.option (@{keyword "morphisms"} |-- Parse.!!! (Parse.binding -- Parse.binding)) --
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
   351
        Scan.option (@{keyword "parametric"} |-- Parse.!!! Parse.xthm))
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61110
diff changeset
   352
      >> (fn (overloaded, spec) => quotient_type_cmd {overloaded = overloaded} spec))
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: 47093
diff changeset
   353
60669
wenzelm
parents: 60225
diff changeset
   354
end