quotient_def.ML
author Christian Urban <urbanc@in.tum.de>
Wed, 04 Nov 2009 11:59:15 +0100
changeset 277 37636f2b1c19
child 279 b2fd070c8833
permissions -rw-r--r--
separated the quotient_def into a separate file
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
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
  val get_fun: flag -> (typ * typ) list -> Proof.context -> typ -> term * (typ * typ)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
  val make_def: binding -> term -> typ -> mixfix -> Attrib.binding -> (typ * typ) list ->
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
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
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
    19
let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  val ((rhs, (_ , thm)), lthy') =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
     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
    22
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  ((rhs, thm), lthy')
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
fun lookup_qenv qenv qty =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
  (case (AList.lookup (op=) qenv qty) of
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
    SOME rty => SOME (qty, rty)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
  | NONE => NONE)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
(* 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
    33
   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
    34
   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
    35
   change *)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
datatype flag = absF | repF
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
fun negF absF = repF
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  | negF repF = absF
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
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
    43
let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
  fun get_fun_aux s fs_tys =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
  let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
    val (fs, tys) = split_list fs_tys
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
    val (otys, ntys) = split_list tys
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
    val oty = Type (s, otys)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
    val nty = Type (s, ntys)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
    val ftys = map (op -->) tys
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
   (case (maps_lookup (ProofContext.theory_of lthy) s) of
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
      SOME info => (list_comb (Const (#mapfun info, ftys ---> (oty --> nty)), fs), (oty, nty))
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
    | 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
    56
  end
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
  fun get_fun_fun fs_tys =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
    val (fs, tys) = split_list fs_tys
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
    val ([oty1, oty2], [nty1, nty2]) = split_list tys
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
    val oty = nty1 --> oty2
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
    val nty = oty1 --> nty2
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
    val ftys = map (op -->) tys
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
    (list_comb (Const (@{const_name "fun_map"}, ftys ---> oty --> nty), fs), (oty, nty))
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
  end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  fun get_const flag (qty, rty) =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
  let 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
    val thy = ProofContext.theory_of lthy
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
    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
    73
  in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
    case flag of
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
      absF => (Const (Sign.full_bname thy ("ABS_" ^ qty_name), rty --> qty), (rty, qty))
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
    | repF => (Const (Sign.full_bname thy ("REP_" ^ qty_name), qty --> rty), (qty, rty))
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  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
    80
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
  if (AList.defined (op=) qenv ty)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  then (get_const flag (the (lookup_qenv qenv ty)))
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  else (case ty of
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
          TFree _ => (mk_identity ty, (ty, ty))
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
        | Type (_, []) => (mk_identity ty, (ty, ty)) 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
        | Type ("fun" , [ty1, ty2]) => 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
                 get_fun_fun [get_fun (negF flag) qenv lthy ty1, get_fun flag qenv lthy ty2]
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
        | Type (s, tys) => get_fun_aux s (map (get_fun flag qenv lthy) tys)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
        | _ => raise ERROR ("no type variables"))
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
fun make_def nconst_bname rhs qty mx attr qenv lthy =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  val (arg_tys, res_ty) = strip_type qty
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
  val rep_fns = map (fst o get_fun repF qenv lthy) arg_tys
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
  val abs_fn  = (fst o get_fun absF qenv lthy) res_ty
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
  fun mk_fun_map t s = 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
        Const (@{const_name "fun_map"}, dummyT) $ t $ s
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
  val absrep_fn = fold_rev mk_fun_map rep_fns abs_fn
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
                  |> Syntax.check_term lthy 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
  define nconst_bname mx attr (absrep_fn $ rhs) lthy
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
(* 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
   111
fun diff (T, S) Ds =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
  case (T, S) of
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
    (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
   114
  | (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
   115
  | (Type (a, Ts), Type (b, Us)) => 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
      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
   117
  | _ => (T, S)::Ds
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
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
   119
  | diffs ([], []) Ds = Ds
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
  | 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
   121
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
fun error_msg lthy (qty, rty) =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
let 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
  val qtystr = quote (Syntax.string_of_typ lthy qty)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  val rtystr = quote (Syntax.string_of_typ lthy rty)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
  error (implode ["Quotient type ", qtystr, " does not match with ", rtystr])
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
fun sanity_chk lthy qenv =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
   val qenv' = Quotient_Info.mk_qenv lthy
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
   val thy = ProofContext.theory_of lthy
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
   fun is_inst thy (qty, rty) (qty', rty') =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
   if Sign.typ_instance thy (qty, qty')
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
   then let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
          val inst = Sign.typ_match thy (qty', qty) Vartab.empty
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
        in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
          rty = Envir.subst_type inst rty'
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
        end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
   else false
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
   fun chk_inst (qty, rty) = 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
     if exists (is_inst thy (qty, rty)) qenv' then true
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
     else error_msg lthy (qty, rty)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
  forall chk_inst qenv
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
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
   154
let   
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  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
   156
  val (_, rhs) = PrimitiveDefs.abs_def prop'
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
  val rty = fastype_of rhs
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
  val qenv = distinct (op=) (diff (qty, rty) []) 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
  sanity_chk lthy qenv;
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
  make_def bind rhs qty mx attr qenv lthy 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
val quotdef_parser =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  (OuterParse.binding --
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
    (OuterParse.$$$ "::" |-- OuterParse.!!! (OuterParse.typ -- 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
      OuterParse.opt_mixfix' --| OuterParse.where_)) >> OuterParse.triple2) -- 
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
       (SpecParse.opt_thm_name ":" -- OuterParse.prop)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
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
   173
let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
  val qty  = (Syntax.check_typ lthy o Syntax.parse_typ lthy) qtystr
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
  val prop = (Syntax.check_prop lthy o Syntax.parse_prop lthy) propstr
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  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
   178
end
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
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
   181
  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
   182
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
end; (* structure *)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
open Quotient_Def;