author | wenzelm |
Wed, 05 Nov 1997 19:39:34 +0100 | |
changeset 4176 | 84a0bfbd74e5 |
parent 4028 | 01745d56307d |
child 4409 | 2af86fcb97d7 |
permissions | -rw-r--r-- |
1475 | 1 |
(* Title: HOL/typedef.ML |
923 | 2 |
ID: $Id$ |
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
2851 | 5 |
Internal interface for typedefs. |
923 | 6 |
*) |
7 |
||
1475 | 8 |
signature TYPEDEF = |
923 | 9 |
sig |
10 |
val prove_nonempty: cterm -> thm list -> tactic option -> thm |
|
3955 | 11 |
val add_typedef: string -> bstring * string list * mixfix -> |
923 | 12 |
string -> string list -> thm list -> tactic option -> theory -> theory |
3955 | 13 |
val add_typedef_i: string -> bstring * string list * mixfix -> |
923 | 14 |
term -> string list -> thm list -> tactic option -> theory -> theory |
15 |
end; |
|
16 |
||
1475 | 17 |
structure Typedef: TYPEDEF = |
923 | 18 |
struct |
19 |
||
20 |
||
21 |
(* prove non-emptyness of a set *) (*exception ERROR*) |
|
22 |
||
3946 | 23 |
val is_def = Logic.is_equals o #prop o rep_thm; |
923 | 24 |
|
25 |
fun prove_nonempty cset thms usr_tac = |
|
26 |
let |
|
27 |
val {T = setT, t = set, maxidx, sign} = rep_cterm cset; |
|
3946 | 28 |
val T = HOLogic.dest_setT setT; |
29 |
val goal = cterm_of sign |
|
30 |
(HOLogic.mk_Trueprop (HOLogic.mk_mem (Var (("x", maxidx + 1), T), set))); |
|
923 | 31 |
val tac = |
32 |
TRY (rewrite_goals_tac (filter is_def thms)) THEN |
|
33 |
TRY (REPEAT_FIRST (resolve_tac (filter_out is_def thms))) THEN |
|
2928
c0e3f1ceabf2
Added trace output and replaced fast_tac set_cs by Fast_tac.
nipkow
parents:
2851
diff
changeset
|
34 |
if_none usr_tac (TRY (ALLGOALS Fast_tac)); |
923 | 35 |
in |
36 |
prove_goalw_cterm [] goal (K [tac]) |
|
37 |
end |
|
38 |
handle ERROR => |
|
2934 | 39 |
error ("Failed to prove nonemptiness of " ^ quote (string_of_cterm cset)); |
923 | 40 |
|
41 |
||
1475 | 42 |
(* ext_typedef *) |
923 | 43 |
|
1475 | 44 |
fun ext_typedef prep_term name (t, vs, mx) raw_set axms thms usr_tac thy = |
923 | 45 |
let |
3946 | 46 |
val _ = require_thy thy "Set" "typedefs"; |
923 | 47 |
val sign = sign_of thy; |
3946 | 48 |
val full_name = Sign.full_name sign; |
923 | 49 |
|
50 |
(*rhs*) |
|
51 |
val cset = prep_term sign raw_set; |
|
52 |
val {T = setT, t = set, ...} = rep_cterm cset; |
|
53 |
val rhs_tfrees = term_tfrees set; |
|
3946 | 54 |
val oldT = HOLogic.dest_setT setT handle TYPE _ => |
923 | 55 |
error ("Not a set type: " ^ quote (Sign.string_of_typ sign setT)); |
56 |
||
57 |
(*lhs*) |
|
58 |
val lhs_tfrees = |
|
3946 | 59 |
map (fn v => (v, if_none (assoc (rhs_tfrees, v)) HOLogic.termS)) vs; |
923 | 60 |
|
3946 | 61 |
val tname = Syntax.type_name t mx; |
923 | 62 |
val tlen = length vs; |
3946 | 63 |
val newT = Type (full_name tname, map TFree lhs_tfrees); |
923 | 64 |
|
65 |
val Rep_name = "Rep_" ^ name; |
|
66 |
val Abs_name = "Abs_" ^ name; |
|
3946 | 67 |
val setC = Const (full_name name, setT); |
68 |
val RepC = Const (full_name Rep_name, newT --> oldT); |
|
69 |
val AbsC = Const (full_name Abs_name, oldT --> newT); |
|
923 | 70 |
val x_new = Free ("x", newT); |
71 |
val y_old = Free ("y", oldT); |
|
72 |
||
73 |
(*axioms*) |
|
3946 | 74 |
val rep_type = HOLogic.mk_Trueprop (HOLogic.mk_mem (RepC $ x_new, setC)); |
75 |
val rep_type_inv = HOLogic.mk_Trueprop (HOLogic.mk_eq (AbsC $ (RepC $ x_new), x_new)); |
|
76 |
val abs_type_inv = Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (y_old, setC)), |
|
77 |
HOLogic.mk_Trueprop (HOLogic.mk_eq (RepC $ (AbsC $ y_old), y_old))); |
|
923 | 78 |
|
79 |
||
80 |
(* errors *) |
|
81 |
||
2238
c72a23bbe762
Eta-expanded some declarations that are illegal under value polymorphism
paulson
parents:
1475
diff
changeset
|
82 |
fun show_names pairs = commas_quote (map fst pairs); |
923 | 83 |
|
84 |
val illegal_vars = |
|
85 |
if null (term_vars set) andalso null (term_tvars set) then [] |
|
86 |
else ["Illegal schematic variable(s) on rhs"]; |
|
87 |
||
88 |
val dup_lhs_tfrees = |
|
89 |
(case duplicates lhs_tfrees of [] => [] |
|
90 |
| dups => ["Duplicate type variables on lhs: " ^ show_names dups]); |
|
91 |
||
92 |
val extra_rhs_tfrees = |
|
93 |
(case gen_rems (op =) (rhs_tfrees, lhs_tfrees) of [] => [] |
|
94 |
| extras => ["Extra type variables on rhs: " ^ show_names extras]); |
|
95 |
||
96 |
val illegal_frees = |
|
97 |
(case term_frees set of [] => [] |
|
98 |
| xs => ["Illegal variables on rhs: " ^ show_names (map dest_Free xs)]); |
|
99 |
||
100 |
val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees; |
|
101 |
in |
|
102 |
if null errs then () |
|
103 |
else error (cat_lines errs); |
|
104 |
||
2928
c0e3f1ceabf2
Added trace output and replaced fast_tac set_cs by Fast_tac.
nipkow
parents:
2851
diff
changeset
|
105 |
writeln("Proving nonemptiness of " ^ quote name ^ " ..."); |
923 | 106 |
prove_nonempty cset (map (get_axiom thy) axms @ thms) usr_tac; |
107 |
||
108 |
thy |
|
3768 | 109 |
|> Theory.add_types [(t, tlen, mx)] |
3793 | 110 |
|> Theory.add_arities_i |
3946 | 111 |
[(full_name tname, replicate tlen logicS, logicS), |
112 |
(full_name tname, replicate tlen HOLogic.termS, HOLogic.termS)] |
|
3768 | 113 |
|> Theory.add_consts_i |
923 | 114 |
[(name, setT, NoSyn), |
115 |
(Rep_name, newT --> oldT, NoSyn), |
|
116 |
(Abs_name, oldT --> newT, NoSyn)] |
|
4028
01745d56307d
PureThy.add_store_defs_i, PureThy.add_store_axioms_i;
wenzelm
parents:
3955
diff
changeset
|
117 |
|> PureThy.add_store_defs_i |
3946 | 118 |
[(name ^ "_def", Logic.mk_equals (setC, set))] |
4028
01745d56307d
PureThy.add_store_defs_i, PureThy.add_store_axioms_i;
wenzelm
parents:
3955
diff
changeset
|
119 |
|> PureThy.add_store_axioms_i |
923 | 120 |
[(Rep_name, rep_type), |
121 |
(Rep_name ^ "_inverse", rep_type_inv), |
|
122 |
(Abs_name ^ "_inverse", abs_type_inv)] |
|
123 |
end |
|
124 |
handle ERROR => |
|
2851 | 125 |
error ("The error(s) above occurred in typedef " ^ quote name); |
923 | 126 |
|
127 |
||
128 |
(* external interfaces *) |
|
129 |
||
3955 | 130 |
fun read_term sg str = |
131 |
read_cterm sg (str, HOLogic.termTVar); |
|
132 |
||
923 | 133 |
fun cert_term sg tm = |
134 |
cterm_of sg tm handle TERM (msg, _) => error msg; |
|
135 |
||
1475 | 136 |
val add_typedef = ext_typedef read_term; |
137 |
val add_typedef_i = ext_typedef cert_term; |
|
923 | 138 |
|
139 |
||
140 |
end; |