quotient_def.ML
author Christian Urban <urbanc@in.tum.de>
Fri, 06 Nov 2009 09:48:37 +0100
changeset 293 653460d3e849
parent 290 a0be84b0c707
child 297 28b264299590
permissions -rw-r--r--
tuned the code in quotient and quotient_def
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
signature QUOTIENT_DEF =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
sig
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
  datatype flag = absF | repF
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
     5
  val get_fun: flag -> (typ * typ) list -> Proof.context -> typ -> term
293
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
     6
  val make_def: binding -> typ -> mixfix -> Attrib.binding -> term ->
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
    Proof.context -> (term * thm) * local_theory
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
  val quotdef: (binding * typ * mixfix) * (Attrib.binding * term) ->
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
    local_theory -> (term * thm) * local_theory
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  val quotdef_cmd: (binding * string * mixfix) * (Attrib.binding * string) ->
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
    local_theory -> local_theory
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
end;
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
structure Quotient_Def: QUOTIENT_DEF =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
struct
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
279
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
    18
(* wrapper for define *)
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
fun define name mx attr rhs lthy =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  val ((rhs, (_ , thm)), lthy') =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
     LocalTheory.define Thm.internalK ((name, mx), (attr, rhs)) lthy
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  ((rhs, thm), lthy')
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
(* calculates the aggregate abs and rep functions for a given type; 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
   repF is for constants' arguments; absF is for constants;
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
   function types need to be treated specially, since repF and absF
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
   change *)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
datatype flag = absF | repF
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
fun negF absF = repF
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  | negF repF = absF
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
fun get_fun flag qenv lthy ty =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    41
  fun get_fun_aux s fs =
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
   (case (maps_lookup (ProofContext.theory_of lthy) s) of
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    43
      SOME info => list_comb (Const (#mapfun info, dummyT), fs)
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
    | NONE      => error ("no map association for type " ^ s))
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    46
  fun get_const flag qty =
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  let 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
    val thy = ProofContext.theory_of lthy
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
    val qty_name = Long_Name.base_name (fst (dest_Type qty))
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
    case flag of
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    52
      absF => Const (Sign.full_bname thy ("ABS_" ^ qty_name), dummyT)
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    53
    | repF => Const (Sign.full_bname thy ("REP_" ^ qty_name), dummyT)
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  fun mk_identity ty = Abs ("", ty, Bound 0)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  if (AList.defined (op=) qenv ty)
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    60
  then (get_const flag ty)
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  else (case ty of
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    62
          TFree _ => mk_identity ty
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    63
        | Type (_, []) => mk_identity ty 
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
        | Type ("fun" , [ty1, ty2]) => 
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    65
            let
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    66
              val fs_ty1 = get_fun (negF flag) qenv lthy ty1
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    67
              val fs_ty2 = get_fun flag qenv lthy ty2
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    68
            in  
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    69
              get_fun_aux "fun" [fs_ty1, fs_ty2]
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    70
            end 
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
        | Type (s, tys) => get_fun_aux s (map (get_fun flag qenv lthy) tys)
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
    72
        | _ => error ("no type variables allowed"))
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
(* returns all subterms where two types differ *)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
fun diff (T, S) Ds =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  case (T, S) of
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
    (TVar v, TVar u) => if v = u then Ds else (T, S)::Ds 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  | (TFree x, TFree y) => if x = y then Ds else (T, S)::Ds
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
  | (Type (a, Ts), Type (b, Us)) => 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
      if a = b then diffs (Ts, Us) Ds else (T, S)::Ds
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
  | _ => (T, S)::Ds
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
and diffs (T::Ts, U::Us) Ds = diffs (Ts, Us) (diff (T, U) Ds)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  | diffs ([], []) Ds = Ds
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
  | diffs _ _ = error "Unequal length of type arguments"
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
279
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
    88
(* sanity check that the calculated quotient environment
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
    89
   matches with the stored quotient environment. *)
293
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
    90
fun sanity_chk qenv lthy =
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
let
286
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
    92
  val global_qenv = Quotient_Info.mk_qenv lthy
279
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
    93
  val thy = ProofContext.theory_of lthy
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
293
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
    95
  fun error_msg lthy (qty, rty) =
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
    96
  let 
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
    97
    val qtystr = quote (Syntax.string_of_typ lthy qty)
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
    98
    val rtystr = quote (Syntax.string_of_typ lthy rty)
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
    99
  in
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   100
    error (implode ["Quotient type ", qtystr, " does not match with ", rtystr])
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   101
  end
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   102
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
   103
  fun is_inst (qty, rty) (qty', rty') =
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
   104
    if Sign.typ_instance thy (qty, qty')
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
   105
    then let
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
   106
           val inst = Sign.typ_match thy (qty', qty) Vartab.empty
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
   107
         in
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
   108
           rty = Envir.subst_type inst rty'       
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
   109
         end
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
   110
    else false
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
279
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   112
  fun chk_inst (qty, rty) = 
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
   113
    if exists (is_inst (qty, rty)) global_qenv 
286
Christian Urban <urbanc@in.tum.de>
parents: 280
diff changeset
   114
    then true
279
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   115
    else error_msg lthy (qty, rty)
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
in
290
a0be84b0c707 removed typing information from get_fun in quotient_def; *potentially* dangerous
Christian Urban <urbanc@in.tum.de>
parents: 287
diff changeset
   117
  map chk_inst qenv
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
293
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   120
fun make_def nconst_bname qty mx attr rhs lthy =
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   121
let
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   122
  val (arg_tys, res_ty) = strip_type qty
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   123
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   124
  val rty = fastype_of rhs
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   125
  val qenv = distinct (op=) (diff (qty, rty) [])   
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   126
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   127
  val _ = sanity_chk qenv lthy
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   128
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   129
  val rep_fns = map (get_fun repF qenv lthy) arg_tys
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   130
  val abs_fn  = (get_fun absF qenv lthy) res_ty
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   131
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   132
  fun mk_fun_map t s =  
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   133
        Const (@{const_name "fun_map"}, dummyT) $ t $ s
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   134
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   135
  val absrep_trm = (fold_rev mk_fun_map rep_fns abs_fn $ rhs)
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   136
                   |> Syntax.check_term lthy 
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   137
in
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   138
  define nconst_bname mx attr absrep_trm lthy
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   139
end
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   140
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   141
(* interface and syntax setup *)
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   142
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   143
(* the ML-interface takes a 5-tuple consisting of  *)
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   144
(*                                                 *)
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   145
(* - the name of the constant to be lifted         *)
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   146
(* - its type                                      *)
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   147
(* - its mixfix annotation                         *)
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   148
(* - a meta-equation defining the constant,        *)
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   149
(*   and the attributes of for this meta-equality  *)
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
fun quotdef ((bind, qty, mx), (attr, prop)) lthy =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
let   
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  val (_, prop') = PrimitiveDefs.dest_def lthy (K true) (K false) (K false) prop
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
  val (_, rhs) = PrimitiveDefs.abs_def prop'
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
in
293
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   156
  make_def bind qty mx attr rhs lthy 
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
fun quotdef_cmd ((bind, qtystr, mx), (attr, propstr)) lthy = 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
let
287
fc72f5b2f9d7 replaced check_term o parse_term by read_term
Christian Urban <urbanc@in.tum.de>
parents: 286
diff changeset
   161
  val qty  = Syntax.read_typ lthy qtystr
fc72f5b2f9d7 replaced check_term o parse_term by read_term
Christian Urban <urbanc@in.tum.de>
parents: 286
diff changeset
   162
  val prop = Syntax.read_prop lthy propstr
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
  quotdef ((bind, qty, mx), (attr, prop)) lthy |> snd
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
279
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   167
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   168
val quotdef_parser =
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   169
  (OuterParse.binding --
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   170
    (OuterParse.$$$ "::" |-- OuterParse.!!! (OuterParse.typ -- 
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   171
      OuterParse.opt_mixfix' --| OuterParse.where_)) >> OuterParse.triple2) -- 
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   172
       (SpecParse.opt_thm_name ":" -- OuterParse.prop)
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   173
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
val _ = OuterSyntax.local_theory "quotient_def" "lifted definition of constants"
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
  OuterKeyword.thy_decl (quotdef_parser >> quotdef_cmd)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
end; (* structure *)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
open Quotient_Def;