Quot/quotient_def.ML
author Christian Urban <urbanc@in.tum.de>
Mon, 21 Dec 2009 22:36:31 +0100
changeset 768 e9e205b904e2
parent 767 37285ec4387d
child 774 b4ffb8826105
permissions -rw-r--r--
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
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
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
     5
  val get_fun: flag -> Proof.context -> typ * typ -> term
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
     6
  
709
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
     7
  val quotient_def: mixfix -> Attrib.binding -> term -> term ->
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
    Proof.context -> (term * thm) * local_theory
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 670
diff changeset
     9
  val quotdef_cmd: (Attrib.binding * string) * (mixfix * string) ->
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
    local_theory -> local_theory
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
end;
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
structure Quotient_Def: QUOTIENT_DEF =
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
struct
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
762
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
    16
open Quotient_Info;
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
    17
open Quotient_Type;
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
    18
279
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
    19
(* wrapper for define *)
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
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
    21
let
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  val ((rhs, (_ , thm)), lthy') =
331
345c422b1cb5 updated to Isabelle 22nd November
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
    23
     Local_Theory.define ((name, mx), (attr, rhs)) lthy
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
in
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
  ((rhs, thm), lthy')
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
end
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
datatype flag = absF | repF
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
fun negF absF = repF
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  | negF repF = absF
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    33
fun mk_identity ty = Const (@{const_name "id"}, ty --> ty)
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
768
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    35
fun mk_compose flag (trm1, trm2) = 
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    36
  case flag of
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    37
    absF => Const (@{const_name "comp"}, dummyT) $ trm1 $ trm2
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    38
  | repF => Const (@{const_name "comp"}, dummyT) $ trm2 $ trm1
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    39
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    40
fun get_fun_aux lthy s fs =
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
    41
let
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
    42
  val thy = ProofContext.theory_of lthy
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
    43
  val exc = LIFT_MATCH (space_implode " " ["get_fun_aux: no map for type", quote s, "."])
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
    44
  val info = maps_lookup thy s handle NotFound => raise exc
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
    45
in
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
    46
  list_comb (Const (#mapfun info, dummyT), fs)
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
    47
end
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    49
fun get_const flag lthy _ qty =
768
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    50
(* FIXME: check here that the type-constructors of _ and qty are related *)
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 709
diff changeset
    51
let
279
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
    52
  val thy = ProofContext.theory_of lthy
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    53
  val qty_name = Long_Name.base_name (fst (dest_Type qty))
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    54
in
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    55
  case flag of
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 709
diff changeset
    56
    absF => Const (Sign.full_bname thy ("abs_" ^ qty_name), dummyT)
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 709
diff changeset
    57
  | repF => Const (Sign.full_bname thy ("rep_" ^ qty_name), dummyT)
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    58
end
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
329
5d06e1dba69a slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
    60
5d06e1dba69a slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
    61
(* calculates the aggregate abs and rep functions for a given type; 
5d06e1dba69a slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
    62
   repF is for constants' arguments; absF is for constants;
5d06e1dba69a slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
    63
   function types need to be treated specially, since repF and absF
5d06e1dba69a slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
    64
   change *)
5d06e1dba69a slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
    65
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    66
fun get_fun flag lthy (rty, qty) =
389
d67240113f68 applic_prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 374
diff changeset
    67
  if rty = qty then mk_identity qty else
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 709
diff changeset
    68
  case (rty, qty) of
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    69
    (Type ("fun", [ty1, ty2]), Type ("fun", [ty1', ty2'])) =>
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    70
     let
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    71
       val fs_ty1 = get_fun (negF flag) lthy (ty1, ty1')
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    72
       val fs_ty2 = get_fun flag lthy (ty2, ty2')
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 709
diff changeset
    73
     in
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    74
       get_fun_aux lthy "fun" [fs_ty1, fs_ty2]
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 709
diff changeset
    75
     end
768
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    76
  | (Type (s, _), Type (s', [])) =>
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    77
     if s = s'
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 709
diff changeset
    78
     then mk_identity qty
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    79
     else get_const flag lthy rty qty
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    80
  | (Type (s, tys), Type (s', tys')) =>
768
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    81
     let
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    82
        val args = map (get_fun flag lthy) (tys ~~ tys')
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    83
     in
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    84
        if s = s'
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    85
        then get_fun_aux lthy s args
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    86
        else mk_compose flag (get_const flag lthy rty qty, get_fun_aux lthy s args)
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    87
     end
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    88
  | (TFree x, TFree x') =>
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
    89
     if x = x'
719
a9e55e1ef64c Renaming
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 709
diff changeset
    90
     then mk_identity qty
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
    91
     else raise (LIFT_MATCH "get_fun (frees)")
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
    92
  | (TVar _, TVar _) => raise (LIFT_MATCH "get_fun (vars)")
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
    93
  | _ => raise (LIFT_MATCH "get_fun (default)")
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
709
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
    95
(* interface and syntax setup *)
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
    96
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
    97
(* the ML-interface takes a 4-tuple consisting of *)
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
    98
(*                                                *)
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
    99
(* - the mixfix annotation                        *)
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
   100
(* - name and attributes of the meta eq           *)
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
   101
(* - the new constant including its type          *)
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
   102
(* - the rhs of the definition                    *)
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
   103
709
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
   104
fun quotient_def mx attr lhs rhs lthy =
293
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   105
let
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 670
diff changeset
   106
  val Free (lhs_str, lhs_ty) = lhs;
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 670
diff changeset
   107
  val qconst_bname = Binding.name lhs_str;
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 670
diff changeset
   108
  val absrep_trm = get_fun absF lthy (fastype_of rhs, lhs_ty) $ rhs
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 598
diff changeset
   109
                   |> Syntax.check_term lthy
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 598
diff changeset
   110
  val prop = Logic.mk_equals (lhs, absrep_trm)
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 598
diff changeset
   111
  val (_, prop') = LocalDefs.cert_def lthy prop
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 598
diff changeset
   112
  val (_, newrhs) = Primitive_Defs.abs_def prop'
310
fec6301a1989 added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   113
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 598
diff changeset
   114
  val ((trm, thm), lthy') = define qconst_bname mx attr newrhs lthy
310
fec6301a1989 added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   115
496
8f1bf5266ebc Added the definition to quotient constant data.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 389
diff changeset
   116
  fun qcinfo phi = qconsts_transfer phi {qconst = trm, rconst = rhs, def = thm}
319
0ae9d9e66cb7 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   117
  val lthy'' = Local_Theory.declaration true
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 670
diff changeset
   118
                 (fn phi => qconsts_update_gen lhs_str (qcinfo phi)) lthy'
293
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   119
in
310
fec6301a1989 added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   120
  ((trm, thm), lthy'')
293
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   121
end
653460d3e849 tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   122
709
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
   123
fun quotdef_cmd ((attr, lhsstr), (mx, rhsstr)) lthy =
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
let
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 598
diff changeset
   125
  val lhs = Syntax.read_term lthy lhsstr
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 598
diff changeset
   126
  val rhs = Syntax.read_term lthy rhsstr
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
in
709
596467882518 Updated comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 705
diff changeset
   128
  quotient_def mx attr lhs rhs lthy |> snd
277
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
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 670
diff changeset
   131
val _ = OuterKeyword.keyword "as";
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 670
diff changeset
   132
279
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   133
val quotdef_parser =
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 670
diff changeset
   134
  (SpecParse.opt_thm_name ":" --
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
   135
     OuterParse.term) --
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
   136
      (OuterParse.opt_mixfix' --| OuterParse.$$$ "as" --
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 719
diff changeset
   137
         OuterParse.term)
279
b2fd070c8833 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 277
diff changeset
   138
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   139
val _ = OuterSyntax.local_theory "quotient_definition" 
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   140
	  "definition for constants over the quotient type"
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   141
             OuterKeyword.thy_decl (quotdef_parser >> quotdef_cmd)
277
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
end; (* structure *)
37636f2b1c19 separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
321
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
   145
f46dc0ca08c3 simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
   146
762
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   147