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