author | wenzelm |
Thu, 09 Oct 2008 20:53:12 +0200 | |
changeset 28545 | 2fb2d48de366 |
parent 27684 | f45fd1159a4b |
child 28716 | ee6f9e50f9c8 |
permissions | -rw-r--r-- |
24914 | 1 |
(* Title: Pure/Isar/subclass.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
25196 | 5 |
User interface for proving subclass relationship between type classes. |
24914 | 6 |
*) |
7 |
||
8 |
signature SUBCLASS = |
|
9 |
sig |
|
10 |
val subclass: class -> local_theory -> Proof.state |
|
11 |
val subclass_cmd: xstring -> local_theory -> Proof.state |
|
25002 | 12 |
val prove_subclass: tactic -> class -> local_theory -> local_theory |
24914 | 13 |
end; |
14 |
||
15 |
structure Subclass : SUBCLASS = |
|
16 |
struct |
|
17 |
||
18 |
local |
|
19 |
||
27684
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
20 |
fun the_option [x] = SOME x |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
21 |
| the_option [] = NONE |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
22 |
| the_option _ = raise Empty; |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
23 |
|
25002 | 24 |
fun gen_subclass prep_class do_proof raw_sup lthy = |
24914 | 25 |
let |
25620 | 26 |
val thy = ProofContext.theory_of lthy; |
24914 | 27 |
val sup = prep_class thy raw_sup; |
25011 | 28 |
val sub = case TheoryTarget.peek lthy |
29 |
of {is_class = false, ...} => error "No class context" |
|
30 |
| {target, ...} => target; |
|
27684
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
31 |
val _ = if Sign.subsort thy ([sup], [sub]) |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
32 |
then error ("Class " ^ Syntax.string_of_sort lthy [sup] |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
33 |
^ " is subclass of class " ^ Syntax.string_of_sort lthy [sub]) |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
34 |
else (); |
25002 | 35 |
val sub_params = map fst (Class.these_params thy [sub]); |
36 |
val sup_params = map fst (Class.these_params thy [sup]); |
|
37 |
val err_params = subtract (op =) sub_params sup_params; |
|
38 |
val _ = if null err_params then [] else |
|
26950 | 39 |
error ("Class " ^ Syntax.string_of_sort lthy [sub] ^ " lacks parameter(s) " ^ |
40 |
commas_quote err_params ^ " of " ^ Syntax.string_of_sort lthy [sup]); |
|
25002 | 41 |
val sublocale_prop = |
25027
44b60657f54f
removed obsolete Class.class_of_locale/locale_of_class;
wenzelm
parents:
25011
diff
changeset
|
42 |
Locale.global_asms_of thy sup |
25002 | 43 |
|> maps snd |
27684
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
44 |
|> the_option |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
45 |
|> Option.map (ObjectLogic.ensure_propT thy); |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
46 |
fun after_qed some_thm = |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
47 |
LocalTheory.theory (Class.prove_subclass (sub, sup) some_thm) |
26276 | 48 |
#> (TheoryTarget.init (SOME sub) o ProofContext.theory_of); |
25002 | 49 |
in |
50 |
do_proof after_qed sublocale_prop lthy |
|
51 |
end; |
|
24934
9f5d60fe9086
renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents:
24914
diff
changeset
|
52 |
|
27684
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
53 |
fun user_proof after_qed NONE = |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
54 |
Proof.theorem_i NONE (K (after_qed NONE)) [[]] |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
55 |
| user_proof after_qed (SOME prop) = |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
56 |
Proof.theorem_i NONE (after_qed o the_option o the_single) [[(prop, [])]]; |
25002 | 57 |
|
27684
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
58 |
fun tactic_proof tac after_qed NONE lthy = |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
59 |
after_qed NONE lthy |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
60 |
| tactic_proof tac after_qed (SOME prop) lthy = |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
61 |
after_qed (SOME (Goal.prove (LocalTheory.target_of lthy) [] [] prop |
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
26950
diff
changeset
|
62 |
(K tac))) lthy; |
24914 | 63 |
|
64 |
in |
|
65 |
||
25002 | 66 |
val subclass = gen_subclass (K I) user_proof; |
67 |
val subclass_cmd = gen_subclass Sign.read_class user_proof; |
|
68 |
fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac); |
|
24914 | 69 |
|
70 |
end; (*local*) |
|
71 |
||
72 |
end; |