| author | wenzelm | 
| Sat, 08 Apr 2006 22:51:25 +0200 | |
| changeset 19369 | a4374b41c9bf | 
| parent 18948 | b6b19cc8454f | 
| child 20885 | e0223c1bd7e8 | 
| permissions | -rw-r--r-- | 
| 
14657
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/Isar/constdefs.ML  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
Author: Makarius, Hagia Maria Sion Abbey (Jerusalem)  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 18779 | 5  | 
Old-style constant definitions, with type-inference and optional  | 
| 
14657
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
structure context; specifications need to observe strictly sequential  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
dependencies; no support for overloading.  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
*)  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
signature CONSTDEFS =  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
sig  | 
| 18668 | 12  | 
val add_constdefs: (string * string option) list *  | 
13  | 
((bstring * string option * mixfix) option * ((bstring * Attrib.src list) * string)) list ->  | 
|
14  | 
theory -> theory  | 
|
15  | 
val add_constdefs_i: (string * typ option) list *  | 
|
| 18728 | 16  | 
((bstring * typ option * mixfix) option * ((bstring * attribute list) * term)) list ->  | 
| 18668 | 17  | 
theory -> theory  | 
| 
14657
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
end;  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
structure Constdefs: CONSTDEFS =  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
struct  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
(** add_constdefs(_i) **)  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
|
| 18668 | 25  | 
fun gen_constdef prep_vars prep_term prep_att  | 
26  | 
structs (raw_decl, ((raw_name, raw_atts), raw_prop)) thy =  | 
|
| 
14657
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
let  | 
| 18948 | 28  | 
fun err msg ts = error (cat_lines (msg :: map (Sign.string_of_term thy) ts));  | 
| 14664 | 29  | 
|
| 18948 | 30  | 
val thy_ctxt = ProofContext.init thy;  | 
31  | 
val struct_ctxt = #2 (ProofContext.add_fixes_i structs thy_ctxt);  | 
|
32  | 
val ((d, mx), var_ctxt) =  | 
|
| 18668 | 33  | 
(case raw_decl of  | 
| 18948 | 34  | 
NONE => ((NONE, NoSyn), struct_ctxt)  | 
| 18668 | 35  | 
| SOME raw_var =>  | 
| 18948 | 36  | 
struct_ctxt |> prep_vars [raw_var] |-> (fn [(x, T, mx)] =>  | 
| 18668 | 37  | 
ProofContext.add_fixes_legacy [(x, T, mx)] #> snd #> pair (SOME x, mx)));  | 
| 
14657
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
|
| 18948 | 39  | 
val prop = prep_term var_ctxt raw_prop;  | 
40  | 
val (c, T) = #1 (LocalDefs.cert_def thy_ctxt (Logic.strip_imp_concl prop));  | 
|
| 15531 | 41  | 
val _ = (case d of NONE => () | SOME c' =>  | 
| 14719 | 42  | 
if c <> c' then  | 
43  | 
        err ("Head of definition " ^ quote c ^ " differs from declaration " ^ quote c') []
 | 
|
| 14664 | 44  | 
else ());  | 
| 
14657
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
|
| 17065 | 46  | 
val def = Term.subst_atomic [(Free (c, T), Const (Sign.full_name thy c, T))] prop;  | 
| 
14657
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
val name = if raw_name = "" then Thm.def_name c else raw_name;  | 
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
val atts = map (prep_att thy) raw_atts;  | 
| 14664 | 49  | 
|
50  | 
val thy' =  | 
|
51  | 
thy  | 
|
52  | 
|> Theory.add_consts_i [(c, T, mx)]  | 
|
| 18668 | 53  | 
|> PureThy.add_defs_i false [((name, def), atts)] |> snd;  | 
| 17853 | 54  | 
in ((c, T), thy') end;  | 
| 
14657
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
|
| 18668 | 56  | 
fun gen_constdefs prep_vars prep_term prep_att (raw_structs, specs) thy =  | 
| 14664 | 57  | 
let  | 
| 18638 | 58  | 
val ctxt = ProofContext.init thy;  | 
| 18668 | 59  | 
val (structs, _) = prep_vars (map (fn (x, T) => (x, T, Structure)) raw_structs) ctxt;  | 
60  | 
val (decls, thy') = fold_map (gen_constdef prep_vars prep_term prep_att structs) specs thy;  | 
|
| 18873 | 61  | 
in Pretty.writeln (LocalTheory.pretty_consts ctxt (K true) decls); thy' end;  | 
| 
14657
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
|
| 18668 | 63  | 
val add_constdefs = gen_constdefs ProofContext.read_vars_legacy  | 
| 18728 | 64  | 
ProofContext.read_term_legacy Attrib.attribute;  | 
| 18668 | 65  | 
val add_constdefs_i = gen_constdefs ProofContext.cert_vars_legacy ProofContext.cert_term (K I);  | 
| 
14657
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
66  | 
|
| 
 
c7cc01735801
'constdefs' with automatic type-inference and structure context;
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
end;  |