author | Christian Urban <urbanc@in.tum.de> |
Sat, 12 Dec 2009 15:07:59 +0100 | |
changeset 734 | ac2ed047988d |
parent 719 | a9e55e1ef64c |
child 760 | c1989de100b4 |
permissions | -rw-r--r-- |
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 |
709 | 6 |
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
|
7 |
Proof.context -> (term * thm) * local_theory |
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
8 |
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
|
9 |
local_theory -> local_theory |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
end; |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
|
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
structure Quotient_Def: QUOTIENT_DEF = |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
struct |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
|
279 | 15 |
(* wrapper for define *) |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
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
|
17 |
let |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
val ((rhs, (_ , thm)), lthy') = |
331
345c422b1cb5
updated to Isabelle 22nd November
Christian Urban <urbanc@in.tum.de>
parents:
329
diff
changeset
|
19 |
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
|
20 |
in |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
((rhs, thm), lthy') |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
end |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
|
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
datatype flag = absF | repF |
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 negF absF = repF |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
| negF repF = absF |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
|
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
|
29 |
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
|
30 |
|
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
|
31 |
fun get_fun_aux lthy s fs = |
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
|
32 |
case (maps_lookup (ProofContext.theory_of lthy) s) of |
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 |
SOME info => list_comb (Const (#mapfun info, dummyT), fs) |
719 | 34 |
| NONE => raise |
374
980fdf92a834
fixed the problem with generalising variables; at the moment it is quite a hack
Christian Urban <urbanc@in.tum.de>
parents:
365
diff
changeset
|
35 |
(LIFT_MATCH (space_implode " " ["get_fun_aux: no map for type", quote s, "."])) |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
|
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
|
37 |
fun get_const flag lthy _ 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
|
38 |
(* FIXME: check here that _ and qty are related *) |
719 | 39 |
let |
279 | 40 |
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
|
41 |
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
|
42 |
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
|
43 |
case flag of |
719 | 44 |
absF => Const (Sign.full_bname thy ("abs_" ^ qty_name), dummyT) |
45 |
| 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
|
46 |
end |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
47 |
|
329 | 48 |
|
49 |
(* calculates the aggregate abs and rep functions for a given type; |
|
50 |
repF is for constants' arguments; absF is for constants; |
|
51 |
function types need to be treated specially, since repF and absF |
|
52 |
change *) |
|
53 |
||
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
|
54 |
fun get_fun flag lthy (rty, qty) = |
389 | 55 |
if rty = qty then mk_identity qty else |
719 | 56 |
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
|
57 |
(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
|
58 |
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
|
59 |
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
|
60 |
val fs_ty2 = get_fun flag lthy (ty2, ty2') |
719 | 61 |
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
|
62 |
get_fun_aux lthy "fun" [fs_ty1, fs_ty2] |
719 | 63 |
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
|
64 |
| (Type (s, []), Type (s', [])) => |
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
|
65 |
if s = s' |
719 | 66 |
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
|
67 |
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
|
68 |
| (Type (s, tys), Type (s', tys')) => |
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 |
if s = s' |
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 |
then get_fun_aux lthy s' (map (get_fun flag lthy) (tys ~~ tys')) |
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 |
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
|
72 |
| (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
|
73 |
if x = x' |
719 | 74 |
then mk_identity qty |
374
980fdf92a834
fixed the problem with generalising variables; at the moment it is quite a hack
Christian Urban <urbanc@in.tum.de>
parents:
365
diff
changeset
|
75 |
else raise (LIFT_MATCH "get_fun") |
980fdf92a834
fixed the problem with generalising variables; at the moment it is quite a hack
Christian Urban <urbanc@in.tum.de>
parents:
365
diff
changeset
|
76 |
| (TVar _, TVar _) => raise (LIFT_MATCH "get_fun") |
980fdf92a834
fixed the problem with generalising variables; at the moment it is quite a hack
Christian Urban <urbanc@in.tum.de>
parents:
365
diff
changeset
|
77 |
| _ => raise (LIFT_MATCH "get_fun") |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
78 |
|
709 | 79 |
(* interface and syntax setup *) |
80 |
||
81 |
(* the ML-interface takes a 4-tuple consisting of *) |
|
82 |
(* *) |
|
83 |
(* - the mixfix annotation *) |
|
84 |
(* - name and attributes of the meta eq *) |
|
85 |
(* - the new constant including its type *) |
|
86 |
(* - the rhs of the definition *) |
|
87 |
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
|
88 |
let |
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
89 |
val Free (lhs_str, lhs_ty) = lhs; |
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
90 |
val qconst_bname = Binding.name lhs_str; |
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
91 |
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
|
92 |
|> 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
|
93 |
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
|
94 |
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
|
95 |
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
|
96 |
|
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
|
97 |
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
|
98 |
|
496
8f1bf5266ebc
Added the definition to quotient constant data.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
389
diff
changeset
|
99 |
fun qcinfo phi = qconsts_transfer phi {qconst = trm, rconst = rhs, def = thm} |
319 | 100 |
val lthy'' = Local_Theory.declaration true |
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
101 |
(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
|
102 |
in |
310
fec6301a1989
added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
parents:
307
diff
changeset
|
103 |
((trm, thm), lthy'') |
293
653460d3e849
tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents:
290
diff
changeset
|
104 |
end |
653460d3e849
tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents:
290
diff
changeset
|
105 |
|
709 | 106 |
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
|
107 |
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
|
108 |
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
|
109 |
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
|
110 |
in |
709 | 111 |
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
|
112 |
end |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
113 |
|
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
114 |
val _ = OuterKeyword.keyword "as"; |
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
115 |
|
279 | 116 |
val quotdef_parser = |
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
117 |
(SpecParse.opt_thm_name ":" -- |
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
118 |
OuterParse.term) -- |
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
119 |
(OuterParse.opt_mixfix' --| OuterParse.$$$ "as" -- |
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
120 |
OuterParse.term) |
279 | 121 |
|
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
122 |
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
|
123 |
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
|
124 |
|
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
125 |
end; (* structure *) |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
126 |
|
307
9aa3aba71ecc
Modifications while preparing the goal-directed version.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
297
diff
changeset
|
127 |
open Quotient_Def; |
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
|
128 |
|
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
|
129 |