renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
1 (* Title: Pure/facts.ML
4 Environment of named facts, optionally indexed by proposition.
9 val the_single: string -> thm list -> thm
10 datatype interval = FromTo of int * int | From of int | Single of int
12 Named of (string * Position.T) * interval list option |
14 val named: string -> ref
15 val string_of_ref: ref -> string
16 val name_of_ref: ref -> string
17 val pos_of_ref: ref -> Position.T
18 val map_name_of_ref: (string -> string) -> ref -> ref
19 val select: ref -> thm list -> thm list
20 val selections: string * thm list -> (ref * thm) list
23 val space_of: T -> Name_Space.T
24 val intern: T -> xstring -> string
25 val extern: T -> string -> xstring
26 val lookup: Context.generic -> T -> string -> (bool * thm list) option
27 val defined: T -> string -> bool
28 val fold_static: (string * thm list -> 'a -> 'a) -> T -> 'a -> 'a
29 val dest_static: T list -> T -> (string * thm list) list
30 val extern_static: T list -> T -> (xstring * thm list) list
31 val props: T -> thm list
32 val could_unify: T -> term -> thm list
34 val add_global: Name_Space.naming -> binding * thm list -> T -> string * T
35 val add_local: bool -> Name_Space.naming -> binding * thm list -> T -> string * T
36 val add_dynamic: Name_Space.naming -> binding * (Context.generic -> thm list) -> T -> string * T
37 val del: string -> T -> T
38 val hide: bool -> string -> T -> T
41 structure Facts: FACTS =
44 (** fact references **)
46 fun the_single _ [th] : thm = th
47 | the_single name _ = error ("Single theorem expected " ^ quote name);
50 (* datatype interval *)
57 fun string_of_interval (FromTo (i, j)) = string_of_int i ^ "-" ^ string_of_int j
58 | string_of_interval (From i) = string_of_int i ^ "-"
59 | string_of_interval (Single i) = string_of_int i;
62 let fun err () = raise Fail ("Bad interval specification " ^ string_of_interval iv) in
64 FromTo (i, j) => if i <= j then i upto j else err ()
65 | From i => if i <= n then i upto n else err ()
73 Named of (string * Position.T) * interval list option |
76 fun named name = Named ((name, Position.none), NONE);
78 fun name_pos_of_ref (Named (name_pos, _)) = name_pos
79 | name_pos_of_ref (Fact _) = error "Illegal literal fact";
81 val name_of_ref = #1 o name_pos_of_ref;
82 val pos_of_ref = #2 o name_pos_of_ref;
84 fun map_name_of_ref f (Named ((name, pos), is)) = Named ((f name, pos), is)
85 | map_name_of_ref _ r = r;
87 fun string_of_ref (Named ((name, _), NONE)) = name
88 | string_of_ref (Named ((name, _), SOME is)) =
89 name ^ enclose "(" ")" (commas (map string_of_interval is))
90 | string_of_ref (Fact _) = error "Illegal literal fact";
95 fun select (Fact _) ths = ths
96 | select (Named (_, NONE)) ths = ths
97 | select (Named ((name, pos), SOME ivs)) ths =
101 error (msg ^ " for " ^ quote name ^ " (length " ^ string_of_int n ^ ")" ^
102 Position.str_of pos);
104 if i < 1 orelse i > n then err ("Bad subscript " ^ string_of_int i)
105 else nth ths (i - 1);
106 val is = maps (interval n) ivs handle Fail msg => err msg;
112 fun selections (name, [th]) = [(Named ((name, Position.none), NONE), th)]
113 | selections (name, ths) = map2 (fn i => fn th =>
114 (Named ((name, Position.none), SOME [Single i]), th)) (1 upto length ths) ths;
118 (** fact environment **)
122 datatype fact = Static of thm list | Dynamic of Context.generic -> thm list;
124 datatype T = Facts of
125 {facts: fact Name_Space.table,
128 fun make_facts facts props = Facts {facts = facts, props = props};
129 val empty = make_facts Name_Space.empty_table Net.empty;
134 fun facts_of (Facts {facts, ...}) = facts;
136 val space_of = #1 o facts_of;
137 val table_of = #2 o facts_of;
139 val intern = Name_Space.intern o space_of;
140 val extern = Name_Space.extern o space_of;
142 val defined = Symtab.defined o table_of;
144 fun lookup context facts name =
145 (case Symtab.lookup (table_of facts) name of
147 | SOME (Static ths) => SOME (true, ths)
148 | SOME (Dynamic f) => SOME (false, f context));
151 Symtab.fold (fn (name, Static ths) => f (name, ths) | _ => I) o table_of;
154 (* content difference *)
156 fun diff_table prev_facts facts =
157 fold_static (fn (name, ths) =>
158 if exists (fn prev => defined prev name) prev_facts then I
159 else cons (name, ths)) facts [];
161 fun dest_static prev_facts facts =
162 sort_wrt #1 (diff_table prev_facts facts);
164 fun extern_static prev_facts facts =
165 sort_wrt #1 (diff_table prev_facts facts |> map (apfst (extern facts)));
170 val prop_ord = TermOrd.term_ord o pairself Thm.full_prop_of;
172 fun props (Facts {props, ...}) = sort_distinct prop_ord (Net.content props);
173 fun could_unify (Facts {props, ...}) = Net.unify_term props;
178 fun merge (Facts {facts = facts1, props = props1}, Facts {facts = facts2, props = props2}) =
180 val facts' = Name_Space.merge_tables (facts1, facts2);
181 val props' = Net.merge (is_equal o prop_ord) (props1, props2);
182 in make_facts facts' props' end;
185 (* add static entries *)
189 fun add strict do_props naming (b, ths) (Facts {facts, props}) =
192 if Binding.is_empty b then ("", facts)
193 else Name_Space.define strict naming (b, Static ths) facts;
196 then fold (fn th => Net.insert_term (K false) (Thm.full_prop_of th, th)) ths props
198 in (name, make_facts facts' props') end;
202 val add_global = add true false;
203 val add_local = add false;
208 (* add dynamic entries *)
210 fun add_dynamic naming (b, f) (Facts {facts, props}) =
211 let val (name, facts') = Name_Space.define true naming (b, Dynamic f) facts;
212 in (name, make_facts facts' props) end;
217 fun del name (Facts {facts = (space, tab), props}) =
219 val space' = Name_Space.hide true name space handle ERROR _ => space; (* FIXME handle?? *)
220 val tab' = Symtab.delete_safe name tab;
221 in make_facts (space', tab') props end;
223 fun hide fully name (Facts {facts = (space, tab), props}) =
224 make_facts (Name_Space.hide fully name space, tab) props;