--- a/doc-src/Codegen/Thy/ML.thy Fri May 15 15:12:23 2009 -0700
+++ b/doc-src/Codegen/Thy/ML.thy Fri May 15 15:13:09 2009 -0700
@@ -25,11 +25,11 @@
@{index_ML Code.add_eqn: "thm -> theory -> theory"} \\
@{index_ML Code.del_eqn: "thm -> theory -> theory"} \\
@{index_ML Code.add_eqnl: "string * (thm * bool) list lazy -> theory -> theory"} \\
- @{index_ML Code.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
- @{index_ML Code.map_post: "(simpset -> simpset) -> theory -> theory"} \\
- @{index_ML Code.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option)
+ @{index_ML Code_Preproc.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
+ @{index_ML Code_Preproc.map_post: "(simpset -> simpset) -> theory -> theory"} \\
+ @{index_ML Code_Preproc.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option)
-> theory -> theory"} \\
- @{index_ML Code.del_functrans: "string -> theory -> theory"} \\
+ @{index_ML Code_Preproc.del_functrans: "string -> theory -> theory"} \\
@{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\
@{index_ML Code.get_datatype: "theory -> string
-> (string * sort) list * (string * typ list) list"} \\
@@ -48,10 +48,10 @@
suspended code equations @{text lthms} for constant
@{text const} to executable content.
- \item @{ML Code.map_pre}~@{text "f"}~@{text "thy"} changes
+ \item @{ML Code_Preproc.map_pre}~@{text "f"}~@{text "thy"} changes
the preprocessor simpset.
- \item @{ML Code.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
+ \item @{ML Code_Preproc.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
function transformer @{text f} (named @{text name}) to executable content;
@{text f} is a transformer of the code equations belonging
to a certain function definition, depending on the
@@ -59,7 +59,7 @@
transformation took place; otherwise, the whole process will be iterated
with the new code equations.
- \item @{ML Code.del_functrans}~@{text "name"}~@{text "thy"} removes
+ \item @{ML Code_Preproc.del_functrans}~@{text "name"}~@{text "thy"} removes
function transformer named @{text name} from executable content.
\item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds
@@ -78,20 +78,16 @@
text %mlref {*
\begin{mldecls}
- @{index_ML Code_Unit.read_const: "theory -> string -> string"} \\
- @{index_ML Code_Unit.head_eqn: "theory -> thm -> string * ((string * sort) list * typ)"} \\
- @{index_ML Code_Unit.rewrite_eqn: "simpset -> thm -> thm"} \\
+ @{index_ML Code.read_const: "theory -> string -> string"} \\
+ @{index_ML Code.rewrite_eqn: "simpset -> thm -> thm"} \\
\end{mldecls}
\begin{description}
- \item @{ML Code_Unit.read_const}~@{text thy}~@{text s}
+ \item @{ML Code.read_const}~@{text thy}~@{text s}
reads a constant as a concrete term expression @{text s}.
- \item @{ML Code_Unit.head_eqn}~@{text thy}~@{text thm}
- extracts the constant and its type from a code equation @{text thm}.
-
- \item @{ML Code_Unit.rewrite_eqn}~@{text ss}~@{text thm}
+ \item @{ML Code.rewrite_eqn}~@{text ss}~@{text thm}
rewrites a code equation @{text thm} with a simpset @{text ss};
only arguments and right hand side are rewritten,
not the head of the code equation.
--- a/doc-src/Codegen/Thy/document/Adaptation.tex Fri May 15 15:12:23 2009 -0700
+++ b/doc-src/Codegen/Thy/document/Adaptation.tex Fri May 15 15:13:09 2009 -0700
@@ -166,7 +166,7 @@
\isa{index} which is mapped to target-language built-in integers.
Useful for code setups which involve e.g. indexing of
target-language arrays.
- \item[\hyperlink{theory.Code-Message}{\mbox{\isa{Code{\isacharunderscore}Message}}}] provides an additional datatype
+ \item[\hyperlink{theory.String}{\mbox{\isa{String}}}] provides an additional datatype
\isa{message{\isacharunderscore}string} which is isomorphic to strings;
\isa{message{\isacharunderscore}string}s are mapped to target-language strings.
Useful for code setups which involve e.g. printing (error) messages.
--- a/doc-src/Codegen/Thy/document/Codegen.tex Fri May 15 15:12:23 2009 -0700
+++ b/doc-src/Codegen/Thy/document/Codegen.tex Fri May 15 15:13:09 2009 -0700
@@ -1550,20 +1550,20 @@
%
\begin{isamarkuptext}%
\begin{mldecls}
- \indexml{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\
- \indexml{Code\_Unit.head\_func}\verb|Code_Unit.head_func: thm -> string * ((string * sort) list * typ)| \\
- \indexml{Code\_Unit.rewrite\_func}\verb|Code_Unit.rewrite_func: MetaSimplifier.simpset -> thm -> thm| \\
+ \indexml{Code\_Unit.read\_const}\verb|Code.read_const: theory -> string -> string| \\
+ \indexml{Code\_Unit.head\_func}\verb|Code.head_func: thm -> string * ((string * sort) list * typ)| \\
+ \indexml{Code\_Unit.rewrite\_func}\verb|Code.rewrite_func: MetaSimplifier.simpset -> thm -> thm| \\
\end{mldecls}
\begin{description}
- \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s}
+ \item \verb|Code.read_const|~\isa{thy}~\isa{s}
reads a constant as a concrete term expression \isa{s}.
- \item \verb|Code_Unit.head_func|~\isa{thm}
+ \item \verb|Code.head_func|~\isa{thm}
extracts the constant and its type from a defining equation \isa{thm}.
- \item \verb|Code_Unit.rewrite_func|~\isa{ss}~\isa{thm}
+ \item \verb|Code.rewrite_func|~\isa{ss}~\isa{thm}
rewrites a defining equation \isa{thm} with a simpset \isa{ss};
only arguments and right hand side are rewritten,
not the head of the defining equation.
--- a/doc-src/Codegen/Thy/document/ML.tex Fri May 15 15:12:23 2009 -0700
+++ b/doc-src/Codegen/Thy/document/ML.tex Fri May 15 15:13:09 2009 -0700
@@ -55,11 +55,11 @@
\indexdef{}{ML}{Code.add\_eqn}\verb|Code.add_eqn: thm -> theory -> theory| \\
\indexdef{}{ML}{Code.del\_eqn}\verb|Code.del_eqn: thm -> theory -> theory| \\
\indexdef{}{ML}{Code.add\_eqnl}\verb|Code.add_eqnl: string * (thm * bool) list lazy -> theory -> theory| \\
- \indexdef{}{ML}{Code.map\_pre}\verb|Code.map_pre: (simpset -> simpset) -> theory -> theory| \\
- \indexdef{}{ML}{Code.map\_post}\verb|Code.map_post: (simpset -> simpset) -> theory -> theory| \\
- \indexdef{}{ML}{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline%
+ \indexdef{}{ML}{Code\_Preproc.map\_pre}\verb|Code_Preproc.map_pre: (simpset -> simpset) -> theory -> theory| \\
+ \indexdef{}{ML}{Code\_Preproc.map\_post}\verb|Code_Preproc.map_post: (simpset -> simpset) -> theory -> theory| \\
+ \indexdef{}{ML}{Code\_Preproc.add\_functrans}\verb|Code_Preproc.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline%
\verb| -> theory -> theory| \\
- \indexdef{}{ML}{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\
+ \indexdef{}{ML}{Code\_Preproc.del\_functrans}\verb|Code_Preproc.del_functrans: string -> theory -> theory| \\
\indexdef{}{ML}{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\
\indexdef{}{ML}{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline%
\verb| -> (string * sort) list * (string * typ list) list| \\
@@ -78,10 +78,10 @@
suspended code equations \isa{lthms} for constant
\isa{const} to executable content.
- \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes
+ \item \verb|Code_Preproc.map_pre|~\isa{f}~\isa{thy} changes
the preprocessor simpset.
- \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds
+ \item \verb|Code_Preproc.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds
function transformer \isa{f} (named \isa{name}) to executable content;
\isa{f} is a transformer of the code equations belonging
to a certain function definition, depending on the
@@ -89,7 +89,7 @@
transformation took place; otherwise, the whole process will be iterated
with the new code equations.
- \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes
+ \item \verb|Code_Preproc.del_functrans|~\isa{name}~\isa{thy} removes
function transformer named \isa{name} from executable content.
\item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds
@@ -124,20 +124,16 @@
%
\begin{isamarkuptext}%
\begin{mldecls}
- \indexdef{}{ML}{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\
- \indexdef{}{ML}{Code\_Unit.head\_eqn}\verb|Code_Unit.head_eqn: theory -> thm -> string * ((string * sort) list * typ)| \\
- \indexdef{}{ML}{Code\_Unit.rewrite\_eqn}\verb|Code_Unit.rewrite_eqn: simpset -> thm -> thm| \\
+ \indexdef{}{ML}{Code\_Unit.read\_const}\verb|Code.read_const: theory -> string -> string| \\
+ \indexdef{}{ML}{Code\_Unit.rewrite\_eqn}\verb|Code.rewrite_eqn: simpset -> thm -> thm| \\
\end{mldecls}
\begin{description}
- \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s}
+ \item \verb|Code.read_const|~\isa{thy}~\isa{s}
reads a constant as a concrete term expression \isa{s}.
- \item \verb|Code_Unit.head_eqn|~\isa{thy}~\isa{thm}
- extracts the constant and its type from a code equation \isa{thm}.
-
- \item \verb|Code_Unit.rewrite_eqn|~\isa{ss}~\isa{thm}
+ \item \verb|Code.rewrite_eqn|~\isa{ss}~\isa{thm}
rewrites a code equation \isa{thm} with a simpset \isa{ss};
only arguments and right hand side are rewritten,
not the head of the code equation.
--- a/doc-src/Codegen/Thy/document/Program.tex Fri May 15 15:12:23 2009 -0700
+++ b/doc-src/Codegen/Thy/document/Program.tex Fri May 15 15:13:09 2009 -0700
@@ -891,8 +891,8 @@
\hspace*{0pt}fun null [] = true\\
\hspace*{0pt} ~| null (x ::~xs) = false;\\
\hspace*{0pt}\\
-\hspace*{0pt}fun eq{\char95}nat (Suc a) Zero{\char95}nat = false\\
-\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat (Suc a) = false\\
+\hspace*{0pt}fun eq{\char95}nat (Suc nat') Zero{\char95}nat = false\\
+\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat (Suc nat') = false\\
\hspace*{0pt} ~| eq{\char95}nat (Suc nat) (Suc nat') = eq{\char95}nat nat nat'\\
\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat Zero{\char95}nat = true;\\
\hspace*{0pt}\\
--- a/doc-src/more_antiquote.ML Fri May 15 15:12:23 2009 -0700
+++ b/doc-src/more_antiquote.ML Fri May 15 15:13:09 2009 -0700
@@ -87,10 +87,10 @@
fun pretty_code_thm src ctxt raw_const =
let
val thy = ProofContext.theory_of ctxt;
- val const = Code_Unit.check_const thy raw_const;
- val (_, funcgr) = Code_Wellsorted.obtain thy [const] [];
+ val const = Code.check_const thy raw_const;
+ val (_, funcgr) = Code_Preproc.obtain thy [const] [];
fun holize thm = @{thm meta_eq_to_obj_eq} OF [thm];
- val thms = Code_Wellsorted.eqns funcgr const
+ val thms = Code_Preproc.eqns funcgr const
|> map_filter (fn (thm, linear) => if linear then SOME thm else NONE)
|> map (holize o no_vars ctxt o AxClass.overload thy);
in ThyOutput.output (ThyOutput.maybe_pretty_source (pretty_thm ctxt) src thms) end;
@@ -108,7 +108,7 @@
local
val parse_const_terms = Scan.repeat1 Args.term
- >> (fn ts => fn thy => map (Code_Unit.check_const thy) ts);
+ >> (fn ts => fn thy => map (Code.check_const thy) ts);
val parse_consts = Scan.lift (Args.parens (Args.$$$ "consts")) |-- parse_const_terms
>> (fn mk_cs => fn thy => fn naming => map_filter (Code_Thingol.lookup_const naming) (mk_cs thy));
val parse_types = Scan.lift (Args.parens (Args.$$$ "types") |-- Scan.repeat1 Args.name)
--- a/etc/isar-keywords-ZF.el Fri May 15 15:12:23 2009 -0700
+++ b/etc/isar-keywords-ZF.el Fri May 15 15:13:09 2009 -0700
@@ -155,8 +155,6 @@
"prop"
"pwd"
"qed"
- "quickcheck"
- "quickcheck_params"
"quit"
"realizability"
"realizers"
@@ -319,7 +317,6 @@
"print_trans_rules"
"prop"
"pwd"
- "quickcheck"
"remove_thy"
"term"
"thm"
@@ -397,7 +394,6 @@
"primrec"
"print_ast_translation"
"print_translation"
- "quickcheck_params"
"realizability"
"realizers"
"rep_datatype"
--- a/etc/isar-keywords.el Fri May 15 15:12:23 2009 -0700
+++ b/etc/isar-keywords.el Fri May 15 15:13:09 2009 -0700
@@ -168,6 +168,7 @@
"print_cases"
"print_claset"
"print_classes"
+ "print_codeproc"
"print_codesetup"
"print_commands"
"print_configs"
@@ -250,6 +251,7 @@
"use_thy"
"using"
"value"
+ "values"
"welcome"
"with"
"{"
@@ -369,6 +371,7 @@
"print_cases"
"print_claset"
"print_classes"
+ "print_codeproc"
"print_codesetup"
"print_commands"
"print_configs"
@@ -402,6 +405,7 @@
"unused_thms"
"use_thy"
"value"
+ "values"
"welcome"))
(defconst isar-keywords-theory-begin
--- a/lib/jedit/isabelle.xml Fri May 15 15:12:23 2009 -0700
+++ b/lib/jedit/isabelle.xml Fri May 15 15:13:09 2009 -0700
@@ -238,6 +238,7 @@
<LABEL>print_cases</LABEL>
<LABEL>print_claset</LABEL>
<LABEL>print_classes</LABEL>
+ <LABEL>print_codeproc</LABEL>
<LABEL>print_codesetup</LABEL>
<LABEL>print_commands</LABEL>
<LABEL>print_configs</LABEL>
@@ -335,6 +336,7 @@
<KEYWORD4>uses</KEYWORD4>
<OPERATOR>using</OPERATOR>
<LABEL>value</LABEL>
+ <LABEL>values</LABEL>
<LABEL>welcome</LABEL>
<KEYWORD4>where</KEYWORD4>
<OPERATOR>with</OPERATOR>
--- a/src/HOL/Bali/State.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Bali/State.thy Fri May 15 15:13:09 2009 -0700
@@ -29,7 +29,7 @@
types vn = "fspec + int" --{* variable name *}
record obj =
tag :: "obj_tag" --{* generalized object *}
- values :: "(vn, val) table"
+ "values" :: "(vn, val) table"
translations
"fspec" <= (type) "vname \<times> qtname"
--- a/src/HOL/Code_Eval.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Code_Eval.thy Fri May 15 15:13:09 2009 -0700
@@ -28,30 +28,15 @@
lemma term_of_anything: "term_of x \<equiv> t"
by (rule eq_reflection) (cases "term_of x", cases t, simp)
-ML {*
-structure Eval =
-struct
-
-fun mk_term f g (Const (c, ty)) =
- @{term Const} $ HOLogic.mk_message_string c $ g ty
- | mk_term f g (t1 $ t2) =
- @{term App} $ mk_term f g t1 $ mk_term f g t2
- | mk_term f g (Free v) = f v
- | mk_term f g (Bound i) = Bound i
- | mk_term f g (Abs (v, _, t)) = Abs (v, @{typ term}, mk_term f g t);
-
-fun mk_term_of ty t = Const (@{const_name term_of}, ty --> @{typ term}) $ t;
-
-end;
-*}
-
subsubsection {* @{text term_of} instances *}
setup {*
let
- fun add_term_of_def ty vs tyco thy =
+ fun add_term_of tyco raw_vs thy =
let
+ val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
+ val ty = Type (tyco, map TFree vs);
val lhs = Const (@{const_name term_of}, ty --> @{typ term})
$ Free ("x", ty);
val rhs = @{term "undefined \<Colon> term"};
@@ -64,66 +49,57 @@
|> `(fn lthy => Syntax.check_term lthy eq)
|-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
|> snd
- |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
- |> LocalTheory.exit_global
+ |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
end;
- fun interpretator ("prop", (raw_vs, _)) thy = thy
- | interpretator (tyco, (raw_vs, _)) thy =
- let
- val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
- val constrain_sort =
- curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort term_of};
- val vs = (map o apsnd) constrain_sort raw_vs;
- val ty = Type (tyco, map TFree vs);
- in
- thy
- |> Typerep.perhaps_add_def tyco
- |> not has_inst ? add_term_of_def ty vs tyco
- end;
+ fun ensure_term_of (tyco, (raw_vs, _)) thy =
+ let
+ val need_inst = not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of})
+ andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
+ in
+ thy
+ |> need_inst ? add_term_of tyco raw_vs
+ end;
in
- Code.type_interpretation interpretator
+ Code.type_interpretation ensure_term_of
end
*}
setup {*
let
- fun mk_term_of_eq ty vs tyco (c, tys) =
+ fun mk_term_of_eq thy ty vs tyco (c, tys) =
let
val t = list_comb (Const (c, tys ---> ty),
map Free (Name.names Name.context "a" tys));
- in (map_aterms (fn Free (v, ty) => Var ((v, 0), ty) | t => t) t, Eval.mk_term
- (fn (v, ty) => Eval.mk_term_of ty (Var ((v, 0), ty)))
- (Typerep.mk (fn (v, sort) => Typerep.typerep (TFree (v, sort)))) t)
+ val (arg, rhs) = pairself (Thm.cterm_of thy o map_types Logic.unvarifyT o Logic.varify)
+ (t, (map_aterms (fn t as Free (v, ty) => HOLogic.mk_term_of ty t | t => t) o HOLogic.reflect_term) t)
+ val cty = Thm.ctyp_of thy ty;
+ in
+ @{thm term_of_anything}
+ |> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs]
+ |> Thm.varifyT
end;
- fun prove_term_of_eq ty eq thy =
+ fun add_term_of_code tyco raw_vs raw_cs thy =
let
- val cty = Thm.ctyp_of thy ty;
- val (arg, rhs) = pairself (Thm.cterm_of thy) eq;
- val thm = @{thm term_of_anything}
- |> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs]
- |> Thm.varifyT;
+ val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
+ val ty = Type (tyco, map TFree vs);
+ val cs = (map o apsnd o map o map_atyps)
+ (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
+ val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
+ val eqs = map (mk_term_of_eq thy ty vs tyco) cs;
+ in
+ thy
+ |> Code.del_eqns const
+ |> fold Code.add_eqn eqs
+ end;
+ fun ensure_term_of_code (tyco, (raw_vs, cs)) thy =
+ let
+ val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
in
thy
- |> Code.add_eqn thm
+ |> has_inst ? add_term_of_code tyco raw_vs cs
end;
- fun interpretator ("prop", (raw_vs, _)) thy = thy
- | interpretator (tyco, (raw_vs, raw_cs)) thy =
- let
- val constrain_sort =
- curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort term_of};
- val vs = (map o apsnd) constrain_sort raw_vs;
- val cs = (map o apsnd o map o map_atyps)
- (fn TFree (v, sort) => TFree (v, constrain_sort sort)) raw_cs;
- val ty = Type (tyco, map TFree vs);
- val eqs = map (mk_term_of_eq ty vs tyco) cs;
- val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
- in
- thy
- |> Code.del_eqns const
- |> fold (prove_term_of_eq ty) eqs
- end;
in
- Code.type_interpretation interpretator
+ Code.type_interpretation ensure_term_of_code
end
*}
@@ -164,8 +140,6 @@
ML {*
signature EVAL =
sig
- val mk_term: ((string * typ) -> term) -> (typ -> term) -> term -> term
- val mk_term_of: typ -> term -> term
val eval_ref: (unit -> term) option ref
val eval_term: theory -> term -> term
end;
@@ -173,14 +147,10 @@
structure Eval : EVAL =
struct
-open Eval;
-
val eval_ref = ref (NONE : (unit -> term) option);
fun eval_term thy t =
- t
- |> Eval.mk_term_of (fastype_of t)
- |> (fn t => Code_ML.eval NONE ("Eval.eval_ref", eval_ref) I thy t []);
+ Code_ML.eval NONE ("Eval.eval_ref", eval_ref) I thy (HOLogic.mk_term_of (fastype_of t) t) [];
end;
*}
--- a/src/HOL/Decision_Procs/Approximation.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Decision_Procs/Approximation.thy Fri May 15 15:13:09 2009 -0700
@@ -460,7 +460,7 @@
proof (cases "even n")
case True
obtain n' where "Suc n' = get_odd n" and "odd (Suc n')" using get_odd_ex by auto
- hence "even n'" unfolding even_nat_Suc by auto
+ hence "even n'" unfolding even_Suc by auto
have "arctan (real x) \<le> real (x * ub_arctan_horner prec (get_odd n) 1 (x * x))"
unfolding `Suc n' = get_odd n`[symmetric] using arctan_0_1_bounds'[OF `0 \<le> real x` `real x \<le> 1` `even n'`] by auto
moreover
@@ -470,7 +470,7 @@
next
case False hence "0 < n" by (rule odd_pos)
from gr0_implies_Suc[OF this] obtain n' where "n = Suc n'" ..
- from False[unfolded this even_nat_Suc]
+ from False[unfolded this even_Suc]
have "even n'" and "even (Suc (Suc n'))" by auto
have "get_odd n = Suc n'" unfolding get_odd_def if_P[OF False] using `n = Suc n'` .
--- a/src/HOL/HOL.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/HOL.thy Fri May 15 15:13:09 2009 -0700
@@ -1331,7 +1331,7 @@
of Abs (_, _, t') => count_loose t' 0 <= 1
| _ => true;
in fn _ => fn ss => fn ct => if is_trivial_let (Thm.term_of ct)
- then SOME @{thm Let_def} (*no or one ocurrenc of bound variable*)
+ then SOME @{thm Let_def} (*no or one ocurrence of bound variable*)
else let (*Norbert Schirmer's case*)
val ctxt = Simplifier.the_context ss;
val thy = ProofContext.theory_of ctxt;
@@ -1775,6 +1775,13 @@
end
*}
+subsubsection {* Generic code generator preprocessor setup *}
+
+setup {*
+ Code_Preproc.map_pre (K HOL_basic_ss)
+ #> Code_Preproc.map_post (K HOL_basic_ss)
+*}
+
subsubsection {* Equality *}
class eq =
@@ -1788,7 +1795,7 @@
lemma eq_refl: "eq x x \<longleftrightarrow> True"
unfolding eq by rule+
-lemma equals_eq [code inline]: "(op =) \<equiv> eq"
+lemma equals_eq: "(op =) \<equiv> eq"
by (rule eq_reflection) (rule ext, rule ext, rule sym, rule eq_equals)
declare equals_eq [symmetric, code post]
@@ -1797,6 +1804,14 @@
declare equals_eq [code]
+setup {*
+ Code_Preproc.map_pre (fn simpset =>
+ simpset addsimprocs [Simplifier.simproc_i @{theory} "eq" [@{term "op ="}]
+ (fn thy => fn _ => fn t as Const (_, T) => case strip_type T
+ of ((T as Type _) :: _, _) => SOME @{thm equals_eq}
+ | _ => NONE)])
+*}
+
subsubsection {* Generic code generator foundation *}
@@ -1836,17 +1851,32 @@
lemmas [code, code unfold, symmetric, code post] = imp_conv_disj
+instantiation itself :: (type) eq
+begin
+
+definition eq_itself :: "'a itself \<Rightarrow> 'a itself \<Rightarrow> bool" where
+ "eq_itself x y \<longleftrightarrow> x = y"
+
+instance proof
+qed (fact eq_itself_def)
+
+end
+
+lemma eq_itself_code [code]:
+ "eq_class.eq TYPE('a) TYPE('a) \<longleftrightarrow> True"
+ by (simp add: eq)
+
text {* Equality *}
declare simp_thms(6) [code nbe]
+setup {*
+ Code.add_const_alias @{thm equals_eq}
+*}
+
hide (open) const eq
hide const eq
-setup {*
- Code_Unit.add_const_alias @{thm equals_eq}
-*}
-
text {* Cases *}
lemma Let_case_cert:
@@ -1867,13 +1897,6 @@
code_abort undefined
-subsubsection {* Generic code generator preprocessor *}
-
-setup {*
- Code.map_pre (K HOL_basic_ss)
- #> Code.map_post (K HOL_basic_ss)
-*}
-
subsubsection {* Generic code generator target languages *}
text {* type bool *}
--- a/src/HOL/IsaMakefile Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/IsaMakefile Fri May 15 15:13:09 2009 -0700
@@ -92,10 +92,10 @@
$(SRC)/Tools/auto_solve.ML \
$(SRC)/Tools/code/code_haskell.ML \
$(SRC)/Tools/code/code_ml.ML \
+ $(SRC)/Tools/code/code_preproc.ML \
$(SRC)/Tools/code/code_printer.ML \
$(SRC)/Tools/code/code_target.ML \
$(SRC)/Tools/code/code_thingol.ML \
- $(SRC)/Tools/code/code_wellsorted.ML \
$(SRC)/Tools/coherent.ML \
$(SRC)/Tools/eqsubst.ML \
$(SRC)/Tools/induct.ML \
@@ -854,7 +854,7 @@
ex/Sudoku.thy ex/Tarski.thy ex/Term_Of_Syntax.thy \
ex/Termination.thy ex/Unification.thy ex/document/root.bib \
ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy \
- ex/Predicate_Compile.thy ex/predicate_compile.ML
+ ex/Predicate_Compile.thy ex/predicate_compile.ML ex/Predicate_Compile_ex.thy
@$(ISABELLE_TOOL) usedir $(OUT)/HOL ex
--- a/src/HOL/Library/Efficient_Nat.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Library/Efficient_Nat.thy Fri May 15 15:13:09 2009 -0700
@@ -179,7 +179,7 @@
else NONE
end;
-val eqn_suc_preproc = Code.simple_functrans (gen_eqn_suc_preproc
+val eqn_suc_preproc = Code_Preproc.simple_functrans (gen_eqn_suc_preproc
@{thm Suc_if_eq} I (fst o Logic.dest_equals));
fun eqn_suc_preproc' thy thms = gen_eqn_suc_preproc
@@ -229,7 +229,7 @@
Codegen.add_preprocessor eqn_suc_preproc'
#> Codegen.add_preprocessor clause_suc_preproc
- #> Code.add_functrans ("eqn_Suc", eqn_suc_preproc)
+ #> Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc)
end;
*}
--- a/src/HOL/Library/Formal_Power_Series.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Library/Formal_Power_Series.thy Fri May 15 15:13:09 2009 -0700
@@ -917,8 +917,7 @@
proof-
have eq: "(1 + X) * ?r = 1"
unfolding minus_one_power_iff
- apply (auto simp add: ring_simps fps_eq_iff)
- by presburger+
+ by (auto simp add: ring_simps fps_eq_iff)
show ?thesis by (auto simp add: eq intro: fps_inverse_unique)
qed
@@ -2286,9 +2285,7 @@
(is "inverse ?l = ?r")
proof-
have th: "?l * ?r = 1"
- apply (auto simp add: ring_simps fps_eq_iff X_mult_nth minus_one_power_iff)
- apply presburger+
- done
+ by (auto simp add: ring_simps fps_eq_iff minus_one_power_iff)
have th': "?l $ 0 \<noteq> 0" by (simp add: )
from fps_inverse_unique[OF th' th] show ?thesis .
qed
--- a/src/HOL/Library/Quickcheck.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Library/Quickcheck.thy Fri May 15 15:13:09 2009 -0700
@@ -75,7 +75,7 @@
val tys = (map snd o fst o strip_abs) t;
val t' = mk_generator_expr thy t tys;
val f = Code_ML.eval (SOME target) ("Quickcheck.eval_ref", eval_ref)
- (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
+ (fn proc => fn g => fn s => g (s + 1) #>> (Option.map o map) proc) thy t' [];
in f #> Random_Engine.run end;
end
--- a/src/HOL/Library/Sum_Of_Squares.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Library/Sum_Of_Squares.thy Fri May 15 15:13:09 2009 -0700
@@ -9,30 +9,20 @@
uses "positivstellensatz.ML" "sum_of_squares.ML"
begin
-method_setup sos = {*
-let
- fun strip_all ct =
- case term_of ct of
- Const("all",_) $ Abs (xn,xT,p) =>
- let val (a,(v,t')) = (apsnd (Thm.dest_abs (SOME xn)) o Thm.dest_comb) ct
- in apfst (cons v) (strip_all t')
- end
- | _ => ([],ct)
+(* Note:
+
+In order to use the method sos, install CSDP (https://projects.coin-or.org/Csdp/) and put the executable csdp on your path.
+
+*)
- fun core_sos_conv ctxt t = Drule.arg_cong_rule @{cterm Trueprop} (Sos.real_sos ctxt (Thm.dest_arg t) RS @{thm Eq_TrueI})
- fun core_sos_tac ctxt = CSUBGOAL (fn (ct, i) =>
- let val (avs, p) = strip_all ct
- val th = standard (fold_rev forall_intr avs (Sos.real_sos ctxt (Thm.dest_arg p)))
- in rtac th i end) (* CONVERSION o core_sos_conv *)
-in Scan.succeed (SIMPLE_METHOD' o core_sos_tac)
-end
-*} "Prove universal problems over the reals using sums of squares"
+method_setup sos = {* Scan.succeed (SIMPLE_METHOD' o Sos.sos_tac) *}
+ "Prove universal problems over the reals using sums of squares"
-text{* Tests -- commented since they work only when csdp is installed *}
+text{* Tests -- commented since they work only when csdp is installed -- see above *}
(*
-lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x --> a < 0" by sos
+lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x \<Longrightarrow> a < 0" by sos
lemma "a1 >= 0 & a2 >= 0 \<and> (a1 * a1 + a2 * a2 = b1 * b1 + b2 * b2 + 2) \<and> (a1 * b1 + a2 * b2 = 0) --> a1 * a2 - b1 * b2 >= (0::real)" by sos
@@ -69,8 +59,8 @@
(* ------------------------------------------------------------------------- *)
(*
lemma "2 <= (x::real) & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 12 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos
+
*)
-
(* ------------------------------------------------------------------------- *)
(* Inequality from sci.math (see "Leon-Sotelo, por favor"). *)
(* ------------------------------------------------------------------------- *)
@@ -110,5 +100,20 @@
*)
(*
lemma "((x::real) - y - 2 * x^4 = 0) & 0 <= x & x <= 2 & 0 <= y & y <= 3 --> y^2 - 7 * y - 12 * x + 17 >= 0" by sos *) (* Too hard?*)
+(*
+lemma "(0::real) <= x --> (1 + x + x^2)/(1 + x^2) <= 1 + x"
+apply sos
+done
+
+lemma "(0::real) <= x --> 1 - x <= 1 / (1 + x + x^2)"
+apply sos
+done
+
+lemma "(x::real) <= 1 / 2 --> - x - 2 * x^2 <= - x / (1 - x)"
+apply sos
+done
+
+lemma "4*r^2 = p^2 - 4*q & r >= (0::real) & x^2 + p*x + q = 0 --> 2*(x::real) = - p + 2*r | 2*x = -p - 2*r" by sos
+*)
end
--- a/src/HOL/Library/sum_of_squares.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Library/sum_of_squares.ML Fri May 15 15:13:09 2009 -0700
@@ -1609,4 +1609,57 @@
fun real_sos ctxt t = gen_prover_real_arith ctxt (real_nonlinear_subst_prover ctxt) t;
end;
+(* A tactic *)
+fun strip_all ct =
+ case term_of ct of
+ Const("all",_) $ Abs (xn,xT,p) =>
+ let val (a,(v,t')) = (apsnd (Thm.dest_abs (SOME xn)) o Thm.dest_comb) ct
+ in apfst (cons v) (strip_all t')
+ end
+| _ => ([],ct)
+
+fun core_sos_conv ctxt t = Drule.arg_cong_rule @{cterm Trueprop} (real_sos ctxt (Thm.dest_arg t) RS @{thm Eq_TrueI})
+fun core_sos_tac ctxt = CSUBGOAL (fn (ct, i) =>
+ let val (avs, p) = strip_all ct
+ val th = standard (fold_rev forall_intr avs (real_sos ctxt (Thm.dest_arg p)))
+ in rtac th i end);
+
+fun default_SOME f NONE v = SOME v
+ | default_SOME f (SOME v) _ = SOME v;
+
+fun lift_SOME f NONE a = f a
+ | lift_SOME f (SOME a) _ = SOME a;
+
+
+local
+ val is_numeral = can (HOLogic.dest_number o term_of)
+in
+fun get_denom b ct = case term_of ct of
+ @{term "op / :: real => _"} $ _ $ _ =>
+ if is_numeral (Thm.dest_arg ct) then get_denom b (Thm.dest_arg1 ct)
+ else default_SOME (get_denom b) (get_denom b (Thm.dest_arg ct)) (Thm.dest_arg ct, b)
+ | @{term "op < :: real => _"} $ _ $ _ => lift_SOME (get_denom true) (get_denom true (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
+ | @{term "op <= :: real => _"} $ _ $ _ => lift_SOME (get_denom true) (get_denom true (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
+ | _ $ _ => lift_SOME (get_denom b) (get_denom b (Thm.dest_fun ct)) (Thm.dest_arg ct)
+ | _ => NONE
+end;
+
+fun elim_one_denom_tac ctxt =
+CSUBGOAL (fn (P,i) =>
+ case get_denom false P of
+ NONE => no_tac
+ | SOME (d,ord) =>
+ let
+ val ss = simpset_of (ProofContext.theory_of ctxt) addsimps @{thms field_simps}
+ addsimps [@{thm nonzero_power_divide}, @{thm power_divide}]
+ val th = instantiate' [] [SOME d, SOME (Thm.dest_arg P)]
+ (if ord then @{lemma "(d=0 --> P) & (d>0 --> P) & (d<(0::real) --> P) ==> P" by auto}
+ else @{lemma "(d=0 --> P) & (d ~= (0::real) --> P) ==> P" by blast})
+ in (rtac th i THEN Simplifier.asm_full_simp_tac ss i) end);
+
+fun elim_denom_tac ctxt i = REPEAT (elim_one_denom_tac ctxt i);
+
+fun sos_tac ctxt = ObjectLogic.full_atomize_tac THEN' elim_denom_tac ctxt THEN' core_sos_tac ctxt
+
+
end;
\ No newline at end of file
--- a/src/HOL/List.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/List.thy Fri May 15 15:13:09 2009 -0700
@@ -1299,6 +1299,25 @@
show ?case by (clarsimp simp add: Cons nth_append)
qed
+lemma Skolem_list_nth:
+ "(ALL i<k. EX x. P i x) = (EX xs. size xs = k & (ALL i<k. P i (xs!i)))"
+ (is "_ = (EX xs. ?P k xs)")
+proof(induct k)
+ case 0 show ?case by simp
+next
+ case (Suc k)
+ show ?case (is "?L = ?R" is "_ = (EX xs. ?P' xs)")
+ proof
+ assume "?R" thus "?L" using Suc by auto
+ next
+ assume "?L"
+ with Suc obtain x xs where "?P k xs & P k x" by (metis less_Suc_eq)
+ hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq)
+ thus "?R" ..
+ qed
+qed
+
+
subsubsection {* @{text list_update} *}
lemma length_list_update [simp]: "length(xs[i:=x]) = length xs"
@@ -3646,10 +3665,14 @@
lemmas in_set_code [code unfold] = mem_iff [symmetric]
-lemma empty_null [code inline]:
+lemma empty_null:
"xs = [] \<longleftrightarrow> null xs"
by (cases xs) simp_all
+lemma [code inline]:
+ "eq_class.eq xs [] \<longleftrightarrow> null xs"
+by (simp add: eq empty_null)
+
lemmas null_empty [code post] =
empty_null [symmetric]
--- a/src/HOL/MacLaurin.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/MacLaurin.thy Fri May 15 15:13:09 2009 -0700
@@ -552,10 +552,6 @@
"[|x = y; abs u \<le> (v::real) |] ==> \<bar>(x + u) - y\<bar> \<le> v"
by auto
-text {* TODO: move to Parity.thy *}
-lemma nat_odd_1 [simp]: "odd (1::nat)"
- unfolding even_nat_def by simp
-
lemma Maclaurin_sin_bound:
"abs(sin x - (\<Sum>m=0..<n. (if even m then 0 else (-1 ^ ((m - Suc 0) div 2)) / real (fact m)) *
x ^ m)) \<le> inverse(real (fact n)) * \<bar>x\<bar> ^ n"
--- a/src/HOL/Nat.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Nat.thy Fri May 15 15:13:09 2009 -0700
@@ -1199,7 +1199,7 @@
definition funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
funpow_code_def [code post]: "funpow = compow"
-lemmas [code inline] = funpow_code_def [symmetric]
+lemmas [code unfold] = funpow_code_def [symmetric]
lemma [code]:
"funpow 0 f = id"
--- a/src/HOL/Option.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Option.thy Fri May 15 15:13:09 2009 -0700
@@ -63,10 +63,8 @@
lemma set_empty_eq [simp]: "(set xo = {}) = (xo = None)"
by (cases xo) auto
-definition
- map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a option \<Rightarrow> 'b option"
-where
- [code del]: "map = (%f y. case y of None => None | Some x => Some (f x))"
+definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a option \<Rightarrow> 'b option" where
+ "map = (%f y. case y of None => None | Some x => Some (f x))"
lemma option_map_None [simp, code]: "map f None = None"
by (simp add: map_def)
@@ -95,14 +93,21 @@
subsubsection {* Code generator setup *}
-definition
- is_none :: "'a option \<Rightarrow> bool" where
- is_none_none [code post, symmetric, code inline]: "is_none x \<longleftrightarrow> x = None"
+definition is_none :: "'a option \<Rightarrow> bool" where
+ [code post]: "is_none x \<longleftrightarrow> x = None"
lemma is_none_code [code]:
shows "is_none None \<longleftrightarrow> True"
and "is_none (Some x) \<longleftrightarrow> False"
- unfolding is_none_none [symmetric] by simp_all
+ unfolding is_none_def by simp_all
+
+lemma is_none_none:
+ "is_none x \<longleftrightarrow> x = None"
+ by (simp add: is_none_def)
+
+lemma [code inline]:
+ "eq_class.eq x None \<longleftrightarrow> is_none x"
+ by (simp add: eq is_none_none)
hide (open) const is_none
--- a/src/HOL/Parity.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Parity.thy Fri May 15 15:13:09 2009 -0700
@@ -29,6 +29,18 @@
end
+lemma even_zero_int[simp]: "even (0::int)" by presburger
+
+lemma odd_one_int[simp]: "odd (1::int)" by presburger
+
+lemma even_zero_nat[simp]: "even (0::nat)" by presburger
+
+lemma odd_zero_nat [simp]: "odd (1::nat)" by presburger
+
+declare even_def[of "number_of v", standard, simp]
+
+declare even_nat_def[of "number_of v", standard, simp]
+
subsection {* Even and odd are mutually exclusive *}
lemma int_pos_lt_two_imp_zero_or_one:
@@ -54,66 +66,47 @@
lemma odd_times_odd: "odd (x::int) ==> odd y ==> odd (x * y)"
by (simp add: even_def zmod_zmult1_eq)
-lemma even_product[presburger]: "even((x::int) * y) = (even x | even y)"
+lemma even_product[simp,presburger]: "even((x::int) * y) = (even x | even y)"
apply (auto simp add: even_times_anything anything_times_even)
apply (rule ccontr)
apply (auto simp add: odd_times_odd)
done
lemma even_plus_even: "even (x::int) ==> even y ==> even (x + y)"
- by presburger
+by presburger
lemma even_plus_odd: "even (x::int) ==> odd y ==> odd (x + y)"
- by presburger
+by presburger
lemma odd_plus_even: "odd (x::int) ==> even y ==> odd (x + y)"
- by presburger
+by presburger
lemma odd_plus_odd: "odd (x::int) ==> odd y ==> even (x + y)" by presburger
-lemma even_sum[presburger]: "even ((x::int) + y) = ((even x & even y) | (odd x & odd y))"
- by presburger
+lemma even_sum[simp,presburger]:
+ "even ((x::int) + y) = ((even x & even y) | (odd x & odd y))"
+by presburger
-lemma even_neg[presburger, algebra]: "even (-(x::int)) = even x" by presburger
+lemma even_neg[simp,presburger,algebra]: "even (-(x::int)) = even x"
+by presburger
-lemma even_difference:
+lemma even_difference[simp]:
"even ((x::int) - y) = ((even x & even y) | (odd x & odd y))" by presburger
-lemma even_pow_gt_zero:
- "even (x::int) ==> 0 < n ==> even (x^n)"
- by (induct n) (auto simp add: even_product)
-
-lemma odd_pow_iff[presburger, algebra]:
- "odd ((x::int) ^ n) \<longleftrightarrow> (n = 0 \<or> odd x)"
- apply (induct n, simp_all)
- apply presburger
- apply (case_tac n, auto)
- apply (simp_all add: even_product)
- done
+lemma even_power[simp,presburger]: "even ((x::int)^n) = (even x & n \<noteq> 0)"
+by (induct n) auto
-lemma odd_pow: "odd x ==> odd((x::int)^n)" by (simp add: odd_pow_iff)
-
-lemma even_power[presburger]: "even ((x::int)^n) = (even x & 0 < n)"
- apply (auto simp add: even_pow_gt_zero)
- apply (erule contrapos_pp, erule odd_pow)
- apply (erule contrapos_pp, simp add: even_def)
- done
-
-lemma even_zero[presburger]: "even (0::int)" by presburger
-
-lemma odd_one[presburger]: "odd (1::int)" by presburger
-
-lemmas even_odd_simps [simp] = even_def[of "number_of v",standard] even_zero
- odd_one even_product even_sum even_neg even_difference even_power
+lemma odd_pow: "odd x ==> odd((x::int)^n)" by simp
subsection {* Equivalent definitions *}
lemma two_times_even_div_two: "even (x::int) ==> 2 * (x div 2) = x"
- by presburger
+by presburger
-lemma two_times_odd_div_two_plus_one: "odd (x::int) ==>
- 2 * (x div 2) + 1 = x" by presburger
+lemma two_times_odd_div_two_plus_one:
+ "odd (x::int) ==> 2 * (x div 2) + 1 = x"
+by presburger
lemma even_equiv_def: "even (x::int) = (EX y. x = 2 * y)" by presburger
@@ -122,45 +115,45 @@
subsection {* even and odd for nats *}
lemma pos_int_even_equiv_nat_even: "0 \<le> x ==> even x = even (nat x)"
- by (simp add: even_nat_def)
-
-lemma even_nat_product[presburger, algebra]: "even((x::nat) * y) = (even x | even y)"
- by (simp add: even_nat_def int_mult)
+by (simp add: even_nat_def)
-lemma even_nat_sum[presburger, algebra]: "even ((x::nat) + y) =
- ((even x & even y) | (odd x & odd y))" by presburger
+lemma even_product_nat[simp,presburger,algebra]:
+ "even((x::nat) * y) = (even x | even y)"
+by (simp add: even_nat_def int_mult)
-lemma even_nat_difference[presburger, algebra]:
- "even ((x::nat) - y) = (x < y | (even x & even y) | (odd x & odd y))"
+lemma even_sum_nat[simp,presburger,algebra]:
+ "even ((x::nat) + y) = ((even x & even y) | (odd x & odd y))"
by presburger
-lemma even_nat_Suc[presburger, algebra]: "even (Suc x) = odd x" by presburger
-
-lemma even_nat_power[presburger, algebra]: "even ((x::nat)^y) = (even x & 0 < y)"
- by (simp add: even_nat_def int_power)
+lemma even_difference_nat[simp,presburger,algebra]:
+ "even ((x::nat) - y) = (x < y | (even x & even y) | (odd x & odd y))"
+by presburger
-lemma even_nat_zero[presburger]: "even (0::nat)" by presburger
+lemma even_Suc[simp,presburger,algebra]: "even (Suc x) = odd x"
+by presburger
-lemmas even_odd_nat_simps [simp] = even_nat_def[of "number_of v",standard]
- even_nat_zero even_nat_Suc even_nat_product even_nat_sum even_nat_power
+lemma even_power_nat[simp,presburger,algebra]:
+ "even ((x::nat)^y) = (even x & 0 < y)"
+by (simp add: even_nat_def int_power)
subsection {* Equivalent definitions *}
-lemma nat_lt_two_imp_zero_or_one: "(x::nat) < Suc (Suc 0) ==>
- x = 0 | x = Suc 0" by presburger
+lemma nat_lt_two_imp_zero_or_one:
+ "(x::nat) < Suc (Suc 0) ==> x = 0 | x = Suc 0"
+by presburger
lemma even_nat_mod_two_eq_zero: "even (x::nat) ==> x mod (Suc (Suc 0)) = 0"
- by presburger
+by presburger
lemma odd_nat_mod_two_eq_one: "odd (x::nat) ==> x mod (Suc (Suc 0)) = Suc 0"
by presburger
lemma even_nat_equiv_def: "even (x::nat) = (x mod Suc (Suc 0) = 0)"
- by presburger
+by presburger
lemma odd_nat_equiv_def: "odd (x::nat) = (x mod Suc (Suc 0) = Suc 0)"
- by presburger
+by presburger
lemma even_nat_div_two_times_two: "even (x::nat) ==>
Suc (Suc 0) * (x div Suc (Suc 0)) = x" by presburger
@@ -169,10 +162,10 @@
Suc( Suc (Suc 0) * (x div Suc (Suc 0))) = x" by presburger
lemma even_nat_equiv_def2: "even (x::nat) = (EX y. x = Suc (Suc 0) * y)"
- by presburger
+by presburger
lemma odd_nat_equiv_def2: "odd (x::nat) = (EX y. x = Suc(Suc (Suc 0) * y))"
- by presburger
+by presburger
subsection {* Parity and powers *}
@@ -183,7 +176,7 @@
apply (induct x)
apply (rule conjI)
apply simp
- apply (insert even_nat_zero, blast)
+ apply (insert even_zero_nat, blast)
apply (simp add: power_Suc)
done
--- a/src/HOL/Power.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Power.thy Fri May 15 15:13:09 2009 -0700
@@ -452,4 +452,13 @@
from power_strict_increasing_iff [OF this] less show ?thesis ..
qed
+
+subsection {* Code generator tweak *}
+
+lemma power_power_power [code, code unfold, code inline del]:
+ "power = power.power (1::'a::{power}) (op *)"
+ unfolding power_def power.power_def ..
+
+declare power.power.simps [code]
+
end
--- a/src/HOL/Predicate.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Predicate.thy Fri May 15 15:13:09 2009 -0700
@@ -610,7 +610,7 @@
(simp_all add: Seq_def single_less_eq_eval contained_less_eq)
lemma eq_pred_code [code]:
- fixes P Q :: "'a::eq pred"
+ fixes P Q :: "'a pred"
shows "eq_class.eq P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P"
unfolding eq by auto
@@ -669,11 +669,26 @@
code_const Seq and Empty and Insert and Join
(Eval "Predicate.Seq" and "Predicate.Empty" and "Predicate.Insert/ (_,/ _)" and "Predicate.Join/ (_,/ _)")
-text {* dummy setup for @{text code_pred} keyword *}
+text {* dummy setup for @{text code_pred} and @{text values} keywords *}
ML {*
-OuterSyntax.local_theory_to_proof "code_pred" "sets up goal for cases rule from given introduction rules and compiles predicate"
- OuterKeyword.thy_goal (OuterParse.term_group >> (K (Proof.theorem_i NONE (K I) [[]])))
+local
+
+structure P = OuterParse;
+
+val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
+
+in
+
+val _ = OuterSyntax.local_theory_to_proof "code_pred" "sets up goal for cases rule from given introduction rules and compiles predicate"
+ OuterKeyword.thy_goal (P.term_group >> (K (Proof.theorem_i NONE (K I) [[]])));
+
+val _ = OuterSyntax.improper_command "values" "evaluate and print enumerations"
+ OuterKeyword.diag ((opt_modes -- P.term)
+ >> (fn (modes, t) => Toplevel.no_timing o Toplevel.keep
+ (K ())));
+
+end
*}
no_notation
--- a/src/HOL/Tools/datatype_codegen.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Tools/datatype_codegen.ML Fri May 15 15:13:09 2009 -0700
@@ -6,7 +6,7 @@
signature DATATYPE_CODEGEN =
sig
- val mk_eq: theory -> string -> thm list
+ val mk_eq_eqns: theory -> string -> (thm * bool) list
val mk_case_cert: theory -> string -> thm
val setup: theory -> theory
end;
@@ -309,18 +309,6 @@
(** generic code generator **)
-(* specification *)
-
-fun add_datatype_spec vs dtco cos thy =
- let
- val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
- in
- thy
- |> try (Code.add_datatype cs)
- |> the_default thy
- end;
-
-
(* case certificates *)
fun mk_case_cert thy tyco =
@@ -354,88 +342,40 @@
|> Thm.varifyT
end;
-fun add_datatype_cases dtco thy =
- let
- val {case_rewrites, ...} = DatatypePackage.the_datatype thy dtco;
- val cert = mk_case_cert thy dtco;
- fun add_case_liberal thy = thy
- |> try (Code.add_case cert)
- |> the_default thy;
- in
- thy
- |> add_case_liberal
- |> fold_rev Code.add_default_eqn case_rewrites
- end;
-
(* equality *)
-local
-
-val not_sym = @{thm HOL.not_sym};
-val not_false_true = iffD2 OF [nth @{thms HOL.simp_thms} 7, TrueI];
-val refl = @{thm refl};
-val eqTrueI = @{thm eqTrueI};
-
-fun mk_distinct cos =
- let
- fun sym_product [] = []
- | sym_product (x::xs) = map (pair x) xs @ sym_product xs;
- fun mk_co_args (co, tys) ctxt =
- let
- val names = Name.invents ctxt "a" (length tys);
- val ctxt' = fold Name.declare names ctxt;
- val vs = map2 (curry Free) names tys;
- in (vs, ctxt') end;
- fun mk_dist ((co1, tys1), (co2, tys2)) =
- let
- val ((xs1, xs2), _) = Name.context
- |> mk_co_args (co1, tys1)
- ||>> mk_co_args (co2, tys2);
- val prem = HOLogic.mk_eq
- (list_comb (co1, xs1), list_comb (co2, xs2));
- val t = HOLogic.mk_not prem;
- in HOLogic.mk_Trueprop t end;
- in map mk_dist (sym_product cos) end;
-
-in
-
-fun mk_eq thy dtco =
+fun mk_eq_eqns thy dtco =
let
- val (vs, cs) = DatatypePackage.the_datatype_spec thy dtco;
- fun mk_triv_inject co =
- let
- val ct' = Thm.cterm_of thy
- (Const (co, Type (dtco, map (fn (v, sort) => TVar ((v, 0), sort)) vs)))
- val cty' = Thm.ctyp_of_term ct';
- val SOME (ct, cty) = fold_aterms (fn Var (v, ty) =>
- (K o SOME) (Thm.cterm_of thy (Var (v, Thm.typ_of cty')), Thm.ctyp_of thy ty) | _ => I)
- (Thm.prop_of refl) NONE;
- in eqTrueI OF [Thm.instantiate ([(cty, cty')], [(ct, ct')]) refl] end;
- val inject1 = map_filter (fn (co, []) => SOME (mk_triv_inject co) | _ => NONE) cs
- val inject2 = (#inject o DatatypePackage.the_datatype thy) dtco;
- val ctxt = ProofContext.init thy;
- val simpset = Simplifier.context ctxt
- (Simplifier.empty_ss addsimprocs [DatatypePackage.distinct_simproc]);
- val cos = map (fn (co, tys) =>
- (Const (co, tys ---> Type (dtco, map TFree vs)), tys)) cs;
- val tac = ALLGOALS (simp_tac simpset)
- THEN ALLGOALS (ProofContext.fact_tac [not_false_true, TrueI]);
- val distinct =
- mk_distinct cos
- |> map (fn t => Goal.prove_global thy [] [] t (K tac))
- |> (fn thms => thms @ map (fn thm => not_sym OF [thm]) thms)
- in inject1 @ inject2 @ distinct end;
+ val (vs, cos) = DatatypePackage.the_datatype_spec thy dtco;
+ val { descr, index, inject = inject_thms, ... } = DatatypePackage.the_datatype thy dtco;
+ val ty = Type (dtco, map TFree vs);
+ fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
+ $ t1 $ t2;
+ fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
+ fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
+ val triv_injects = map_filter
+ (fn (c, []) => SOME (HOLogic.mk_Trueprop (true_eq (Const (c, ty), Const (c, ty))))
+ | _ => NONE) cos;
+ fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) =
+ trueprop $ (equiv $ mk_eq (t1, t2) $ rhs);
+ val injects = map prep_inject (nth (DatatypeProp.make_injs [descr] vs) index);
+ fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) =
+ [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)];
+ val distincts = maps prep_distinct (snd (nth (DatatypeProp.make_distincts [descr] vs) index));
+ val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
+ val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss
+ addsimps (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms))
+ addsimprocs [DatatypePackage.distinct_simproc]);
+ fun prove prop = Goal.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
+ |> Simpdata.mk_eq;
+ in map (rpair true o prove) (triv_injects @ injects @ distincts) @ [(prove refl, false)] end;
-end;
-
-fun add_datatypes_equality vs dtcos thy =
+fun add_equality vs dtcos thy =
let
- val vs' = (map o apsnd)
- (curry (Sorts.inter_sort (Sign.classes_of thy)) [HOLogic.class_eq]) vs;
fun add_def dtco lthy =
let
- val ty = Type (dtco, map TFree vs');
+ val ty = Type (dtco, map TFree vs);
fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
$ Free ("x", ty) $ Free ("y", ty);
val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
@@ -448,52 +388,60 @@
in (thm', lthy') end;
fun tac thms = Class.intro_classes_tac []
THEN ALLGOALS (ProofContext.fact_tac thms);
- fun mk_eq' thy dtco = mk_eq thy dtco
- |> map (Code_Unit.constrain_thm thy [HOLogic.class_eq])
- |> map Simpdata.mk_eq
- |> map (MetaSimplifier.rewrite_rule [Thm.transfer thy @{thm equals_eq}])
- |> map (AxClass.unoverload thy);
fun add_eq_thms dtco thy =
let
- val ty = Type (dtco, map TFree vs');
+ val const = AxClass.param_of_inst thy (@{const_name eq_class.eq}, dtco);
val thy_ref = Theory.check_thy thy;
- val const = AxClass.param_of_inst thy (@{const_name eq_class.eq}, dtco);
- val eq_refl = @{thm HOL.eq_refl}
- |> Thm.instantiate
- ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT ty)], [])
- |> Simpdata.mk_eq
- |> AxClass.unoverload thy;
- fun mk_thms () = (eq_refl, false)
- :: rev (map (rpair true) (mk_eq' (Theory.deref thy_ref) dtco));
+ fun mk_thms () = rev ((mk_eq_eqns (Theory.deref thy_ref) dtco));
in
Code.add_eqnl (const, Lazy.lazy mk_thms) thy
end;
in
thy
- |> TheoryTarget.instantiation (dtcos, vs', [HOLogic.class_eq])
+ |> TheoryTarget.instantiation (dtcos, vs, [HOLogic.class_eq])
|> fold_map add_def dtcos
- |-> (fn thms => Class.prove_instantiation_instance (K (tac thms))
- #> LocalTheory.exit_global
- #> fold Code.del_eqn thms
- #> fold add_eq_thms dtcos)
+ |-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
+ (fn _ => fn def_thms => tac def_thms) def_thms)
+ |-> (fn def_thms => fold Code.del_eqn def_thms)
+ |> fold add_eq_thms dtcos
+ end;
+
+
+(* liberal addition of code data for datatypes *)
+
+fun mk_constr_consts thy vs dtco cos =
+ let
+ val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
+ val cs' = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) cs;
+ in if is_some (try (Code.constrset_of_consts thy) cs')
+ then SOME cs
+ else NONE
end;
+fun add_all_code dtcos thy =
+ let
+ val (vs :: _, coss) = (split_list o map (DatatypePackage.the_datatype_spec thy)) dtcos;
+ val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
+ val css = if exists is_none any_css then []
+ else map_filter I any_css;
+ val case_rewrites = maps (#case_rewrites o DatatypePackage.the_datatype thy) dtcos;
+ val certs = map (mk_case_cert thy) dtcos;
+ in
+ if null css then thy
+ else thy
+ |> fold Code.add_datatype css
+ |> fold_rev Code.add_default_eqn case_rewrites
+ |> fold Code.add_case certs
+ |> add_equality vs dtcos
+ end;
+
+
(** theory setup **)
-fun add_datatype_code dtcos thy =
- let
- val (vs :: _, coss) = (split_list o map (DatatypePackage.the_datatype_spec thy)) dtcos;
- in
- thy
- |> fold2 (add_datatype_spec vs) dtcos coss
- |> fold add_datatype_cases dtcos
- |> add_datatypes_equality vs dtcos
- end;
-
val setup =
add_codegen "datatype" datatype_codegen
#> add_tycodegen "datatype" datatype_tycodegen
- #> DatatypePackage.interpretation add_datatype_code
+ #> DatatypePackage.interpretation add_all_code
end;
--- a/src/HOL/Tools/hologic.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Tools/hologic.ML Fri May 15 15:13:09 2009 -0700
@@ -119,6 +119,9 @@
val message_stringT: typ
val mk_message_string: string -> term
val dest_message_string: term -> string
+ val mk_typerep: typ -> term
+ val mk_term_of: typ -> term -> term
+ val reflect_term: term -> term
end;
structure HOLogic: HOLOGIC =
@@ -591,4 +594,26 @@
dest_string t
| dest_message_string t = raise TERM ("dest_message_string", [t]);
+
+(* typerep and term *)
+
+val typerepT = Type ("Typerep.typerep", []);
+
+fun mk_typerep T = Const ("Typerep.typerep_class.typerep",
+ Term.itselfT T --> typerepT) $ Logic.mk_type T;
+
+val termT = Type ("Code_Eval.term", []);
+
+fun mk_term_of T t = Const ("Code_Eval.term_of_class.term_of", T --> termT) $ t;
+
+fun reflect_term (Const (c, T)) =
+ Const ("Code_Eval.Const", message_stringT --> typerepT --> termT)
+ $ mk_message_string c $ mk_typerep T
+ | reflect_term (t1 $ t2) =
+ Const ("Code_Eval.App", termT --> termT --> termT)
+ $ reflect_term t1 $ reflect_term t2
+ | reflect_term (t as Free _) = t
+ | reflect_term (t as Bound _) = t
+ | reflect_term (Abs (v, _, t)) = Abs (v, termT, reflect_term t);
+
end;
--- a/src/HOL/Tools/recfun_codegen.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Tools/recfun_codegen.ML Fri May 15 15:13:09 2009 -0700
@@ -26,10 +26,10 @@
fun add_thm NONE thm thy = Code.add_eqn thm thy
| add_thm (SOME module_name) thm thy =
let
- val (thm', _) = Code_Unit.mk_eqn thy (K false) (thm, true)
+ val (thm', _) = Code.mk_eqn thy (K false) (thm, true)
in
thy
- |> ModuleData.map (Symtab.update (Code_Unit.const_eqn thy thm', module_name))
+ |> ModuleData.map (Symtab.update (Code.const_eqn thy thm', module_name))
|> Code.add_eqn thm'
end;
@@ -44,21 +44,21 @@
fun expand_eta thy [] = []
| expand_eta thy (thms as thm :: _) =
let
- val (_, ty) = Code_Unit.const_typ_eqn thm;
+ val (_, ty) = Code.const_typ_eqn thm;
in if null (Term.add_tvarsT ty []) orelse (null o fst o strip_type) ty
then thms
- else map (Code_Unit.expand_eta thy 1) thms
+ else map (Code.expand_eta thy 1) thms
end;
fun retrieve_equations thy (c, T) = if c = @{const_name "op ="} then NONE else
let
val c' = AxClass.unoverload_const thy (c, T);
val opt_name = Symtab.lookup (ModuleData.get thy) c';
- val thms = Code.these_raw_eqns thy c'
+ val thms = Code.these_eqns thy c'
|> map_filter (fn (thm, linear) => if linear then SOME thm else NONE)
|> expand_eta thy
|> map_filter (meta_eq_to_obj_eq thy)
- |> Code_Unit.norm_varnames thy
+ |> Code.norm_varnames thy
|> map (rpair opt_name)
in if null thms then NONE else SOME thms end;
--- a/src/HOL/Tools/record_package.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Tools/record_package.ML Fri May 15 15:13:09 2009 -0700
@@ -1464,7 +1464,7 @@
val tname = Binding.name (Long_Name.base_name name);
in
thy
- |> TypecopyPackage.add_typecopy (Binding.suffix_name ext_typeN tname, alphas) repT NONE
+ |> TypecopyPackage.typecopy (Binding.suffix_name ext_typeN tname, alphas) repT NONE
|-> (fn (name, _) => `(fn thy => get_thms thy name))
end;
--- a/src/HOL/Tools/typecopy_package.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Tools/typecopy_package.ML Fri May 15 15:13:09 2009 -0700
@@ -14,12 +14,12 @@
proj: string * typ,
proj_def: thm
}
- val add_typecopy: binding * string list -> typ -> (binding * binding) option
+ val typecopy: binding * string list -> typ -> (binding * binding) option
-> theory -> (string * info) * theory
val get_typecopies: theory -> string list
val get_info: theory -> string -> info option
val interpretation: (string -> theory -> theory) -> theory -> theory
- val add_typecopy_default_code: string -> theory -> theory
+ val add_default_code: string -> theory -> theory
val print_typecopies: theory -> unit
val setup: theory -> theory
end;
@@ -71,7 +71,10 @@
structure TypecopyInterpretation = InterpretationFun(type T = string val eq = op =);
val interpretation = TypecopyInterpretation.interpretation;
-fun add_typecopy (raw_tyco, raw_vs) raw_ty constr_proj thy =
+
+(* declaring typecopies *)
+
+fun typecopy (raw_tyco, raw_vs) raw_ty constr_proj thy =
let
val ty = Sign.certify_typ thy raw_ty;
val vs =
@@ -108,28 +111,26 @@
end;
-(* code generator setup *)
+(* default code setup *)
-fun add_typecopy_default_code tyco thy =
+fun add_default_code tyco thy =
let
val SOME { constr = constr_name, proj = (proj, _), proj_def = proj_eq, vs = raw_vs,
- typ = raw_ty_rep, ... } = get_info thy tyco;
- val inst_meet = Sorts.meet_sort_typ (Sign.classes_of thy)
- (Logic.varifyT raw_ty_rep, [HOLogic.class_eq]) handle Sorts.CLASS_ERROR _ => I;
- val ty_inst = Logic.unvarifyT o inst_meet o Logic.varifyT;
- val vs = (map dest_TFree o snd o dest_Type o ty_inst)
- (Type (tyco, map TFree raw_vs));
- val ty_rep = ty_inst raw_ty_rep;
+ typ = ty_rep, ... } = get_info thy tyco;
val SOME { Rep_inject = proj_inject, ... } = TypedefPackage.get_info thy tyco;
- val ty_constr = Logic.unvarifyT (Sign.the_const_type thy constr_name);
- val constr = (constr_name, ty_constr)
+ val constr = (constr_name, Logic.unvarifyT (Sign.the_const_type thy constr_name));
+ val vs = (map dest_TFree o snd o dest_Type) (Type (tyco, map TFree raw_vs));
val ty = Type (tyco, map TFree vs);
- fun mk_eq ty t_x t_y = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
+ val proj = Const (proj, ty --> ty_rep);
+ val (t_x, t_y) = (Free ("x", ty), Free ("y", ty));
+ val eq_lhs = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
$ t_x $ t_y;
- fun mk_proj t = Const (proj, ty --> ty_rep) $ t;
- val (t_x, t_y) = (Free ("x", ty), Free ("y", ty));
- val def_eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq)
- (mk_eq ty t_x t_y, HOLogic.mk_eq (mk_proj t_x, mk_proj t_y));
+ val eq_rhs = HOLogic.mk_eq (proj $ t_x, proj $ t_y);
+ val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (eq_lhs, eq_rhs);
+ fun tac eq_thm = Class.intro_classes_tac []
+ THEN (Simplifier.rewrite_goals_tac
+ (map Simpdata.mk_eq [eq_thm, @{thm eq}, proj_inject]))
+ THEN ALLGOALS (rtac @{thm refl});
fun mk_eq_refl thy = @{thm HOL.eq_refl}
|> Thm.instantiate
([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT ty)], [])
@@ -139,22 +140,18 @@
|> Code.add_datatype [constr]
|> Code.add_eqn proj_eq
|> TheoryTarget.instantiation ([tyco], vs, [HOLogic.class_eq])
- |> `(fn lthy => Syntax.check_term lthy def_eq)
- |-> (fn def_eq => Specification.definition
- (NONE, (Attrib.empty_binding, def_eq)))
- |-> (fn (_, (_, def_thm)) =>
+ |> `(fn lthy => Syntax.check_term lthy eq)
+ |-> (fn eq => Specification.definition
+ (NONE, (Attrib.empty_binding, eq)))
+ |-> (fn (_, (_, eq_thm)) =>
Class.prove_instantiation_exit_result Morphism.thm
- (fn _ => fn def_thm => Class.intro_classes_tac []
- THEN (Simplifier.rewrite_goals_tac
- (map Simpdata.mk_eq [def_thm, @{thm eq}, proj_inject]))
- THEN ALLGOALS (rtac @{thm refl})) def_thm)
- |-> (fn def_thm => Code.add_eqn def_thm)
- |> `(fn thy => mk_eq_refl thy)
- |-> (fn refl_thm => Code.add_nbe_eqn refl_thm)
+ (fn _ => fn eq_thm => tac eq_thm) eq_thm)
+ |-> (fn eq_thm => Code.add_eqn eq_thm)
+ |> (fn thy => Code.add_nbe_eqn (mk_eq_refl thy) thy)
end;
val setup =
TypecopyInterpretation.init
- #> interpretation add_typecopy_default_code
+ #> interpretation add_default_code
end;
--- a/src/HOL/Transcendental.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Transcendental.thy Fri May 15 15:13:09 2009 -0700
@@ -173,7 +173,7 @@
have "(\<lambda> n. if even n then f (n div 2) else 0) sums y"
unfolding sums_def setsum_shift_lb_Suc0_0_upt[where f="?s", OF `?s 0 = 0`, symmetric]
image_Suc_atLeastLessThan[symmetric] setsum_reindex[OF inj_Suc, unfolded comp_def]
- even_nat_Suc Suc_m1 if_eq .
+ even_Suc Suc_m1 if_eq .
} from sums_add[OF g_sums this]
show ?thesis unfolding if_sum .
qed
--- a/src/HOL/Typerep.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/Typerep.thy Fri May 15 15:13:09 2009 -0700
@@ -35,28 +35,18 @@
end
*}
-ML {*
-structure Typerep =
-struct
+setup {*
+let
-fun mk f (Type (tyco, tys)) =
- @{term Typerep} $ HOLogic.mk_message_string tyco
- $ HOLogic.mk_list @{typ typerep} (map (mk f) tys)
- | mk f (TFree v) =
- f v;
-
-fun typerep ty =
- Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
- $ Logic.mk_type ty;
-
-fun add_def tyco thy =
+fun add_typerep tyco thy =
let
val sorts = replicate (Sign.arity_number thy tyco) @{sort typerep};
val vs = Name.names Name.context "'a" sorts;
val ty = Type (tyco, map TFree vs);
val lhs = Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
$ Free ("T", Term.itselfT ty);
- val rhs = mk (typerep o TFree) ty;
+ val rhs = @{term Typerep} $ HOLogic.mk_message_string tyco
+ $ HOLogic.mk_list @{typ typerep} (map (HOLogic.mk_typerep o TFree) vs);
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
in
thy
@@ -64,23 +54,20 @@
|> `(fn lthy => Syntax.check_term lthy eq)
|-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
|> snd
- |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
- |> LocalTheory.exit_global
+ |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
end;
-fun perhaps_add_def tyco thy =
- let
- val inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep}
- in if inst then thy else add_def tyco thy end;
+fun ensure_typerep tyco thy = if not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep})
+ andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort type}
+ then add_typerep tyco thy else thy;
+
+in
-end;
-*}
+add_typerep @{type_name fun}
+#> TypedefPackage.interpretation ensure_typerep
+#> Code.type_interpretation (ensure_typerep o fst)
-setup {*
- Typerep.add_def @{type_name fun}
- #> Typerep.add_def @{type_name itself}
- #> Typerep.add_def @{type_name bool}
- #> TypedefPackage.interpretation Typerep.perhaps_add_def
+end
*}
lemma [code]:
--- a/src/HOL/ex/Predicate_Compile.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/ex/Predicate_Compile.thy Fri May 15 15:13:09 2009 -0700
@@ -7,11 +7,6 @@
setup {* Predicate_Compile.setup *}
-ML {*
- OuterSyntax.local_theory_to_proof "code_pred" "sets up goal for cases rule from given introduction rules and compiles predicate"
- OuterKeyword.thy_goal (OuterParse.term_group >> Predicate_Compile.code_pred_cmd)
-*}
-
text {* Experimental code *}
@@ -27,9 +22,7 @@
val pred_ref = ref (NONE : (unit -> term Predicate.pred) option);
fun eval_pred thy t =
- t
- |> Eval.mk_term_of (fastype_of t)
- |> (fn t => Code_ML.eval NONE ("Predicate.pred_ref", pred_ref) @{code pred_map} thy t []);
+ Code_ML.eval NONE ("Predicate.pred_ref", pred_ref) @{code pred_map} thy (HOLogic.mk_term_of (fastype_of t) t) [];
fun eval_pred_elems thy t T length =
t |> eval_pred thy |> yieldn length |> fst |> HOLogic.mk_list T;
@@ -51,53 +44,4 @@
end;
*}
-
-text {* Example(s) *}
-
-inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
- "even 0"
- | "even n \<Longrightarrow> odd (Suc n)"
- | "odd n \<Longrightarrow> even (Suc n)"
-
-setup {* pred_compile "even" *}
-thm even_codegen
-
-
-inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
- append_Nil: "append [] xs xs"
- | append_Cons: "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
-
-setup {* pred_compile "append" *}
-thm append_codegen
-
-
-inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
- for f where
- "partition f [] [] []"
- | "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"
- | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
-
-setup {* pred_compile "partition" *}
-thm partition_codegen
-
-setup {* pred_compile "tranclp" *}
-thm tranclp_codegen
-
-ML_val {* Predicate_Compile.modes_of @{theory} @{const_name partition} *}
-ML_val {* Predicate_Compile.modes_of @{theory} @{const_name tranclp} *}
-
-ML_val {* Predicate.analyze_compr @{theory} @{term "{n. odd n}"} *}
-
-section {* Example for user interface *}
-
-inductive append2 :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
- append2_Nil: "append2 [] xs xs"
- | append2_Cons: "append2 xs ys zs \<Longrightarrow> append2 (x # xs) ys (x # zs)"
-
-(*code_pred append2
- using assms by (rule append2.cases)
-
-thm append2_codegen
-thm append2_cases*)
-
end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Predicate_Compile_ex.thy Fri May 15 15:13:09 2009 -0700
@@ -0,0 +1,43 @@
+theory Predicate_Compile_ex
+imports Complex_Main Predicate_Compile
+begin
+
+inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
+ "even 0"
+ | "even n \<Longrightarrow> odd (Suc n)"
+ | "odd n \<Longrightarrow> even (Suc n)"
+
+code_pred even
+ using assms by (rule even.cases)
+
+thm even.equation
+
+
+inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
+ append_Nil: "append [] xs xs"
+ | append_Cons: "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
+
+code_pred append
+ using assms by (rule append.cases)
+
+thm append.equation
+
+
+inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
+ for f where
+ "partition f [] [] []"
+ | "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"
+ | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
+
+code_pred partition
+ using assms by (rule partition.cases)
+
+thm partition.equation
+
+
+code_pred tranclp
+ using assms by (rule tranclp.cases)
+
+thm tranclp.equation
+
+end
\ No newline at end of file
--- a/src/HOL/ex/Quickcheck_Generators.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/ex/Quickcheck_Generators.thy Fri May 15 15:13:09 2009 -0700
@@ -54,7 +54,7 @@
val t_indices = map (curry ( op * ) 2) (length tys - 1 downto 0);
val c_indices = map (curry ( op + ) 1) t_indices;
val c_t = list_comb (c, map Bound c_indices);
- val t_t = Abs ("", @{typ unit}, Eval.mk_term Free Typerep.typerep
+ val t_t = Abs ("", @{typ unit}, HOLogic.reflect_term
(list_comb (c, map (fn k => Bound (k + 1)) t_indices))
|> map_aterms (fn t as Bound _ => t $ @{term "()"} | t => t));
val return = StateMonad.return (term_ty this_ty) @{typ seed}
--- a/src/HOL/ex/ROOT.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/ex/ROOT.ML Fri May 15 15:13:09 2009 -0700
@@ -16,7 +16,8 @@
"Codegenerator_Pretty",
"NormalForm",
"../NumberTheory/Factorization",
- "Predicate_Compile"
+ "Predicate_Compile",
+ "Predicate_Compile_ex"
];
use_thys [
--- a/src/HOL/ex/Term_Of_Syntax.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/ex/Term_Of_Syntax.thy Fri May 15 15:13:09 2009 -0700
@@ -31,9 +31,9 @@
setup {*
let
- val subst_rterm_of = Eval.mk_term
- (fn (v, _) => error ("illegal free variable in term quotation: " ^ quote v))
- (Typerep.mk (fn (v, sort) => Typerep.typerep (TFree (v, sort))));
+ val subst_rterm_of = map_aterms
+ (fn Free (v, _) => error ("illegal free variable in term quotation: " ^ quote v) | t => t)
+ o HOLogic.reflect_term;
fun subst_rterm_of' (Const (@{const_name rterm_of}, _), [t]) = subst_rterm_of t
| subst_rterm_of' (Const (@{const_name rterm_of}, _), _) =
error ("illegal number of arguments for " ^ quote @{const_name rterm_of})
--- a/src/HOL/ex/predicate_compile.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/HOL/ex/predicate_compile.ML Fri May 15 15:13:09 2009 -0700
@@ -7,25 +7,22 @@
signature PREDICATE_COMPILE =
sig
type mode = int list option list * int list
- val create_def_equation': string -> mode option -> theory -> theory
- val create_def_equation: string -> theory -> theory
+ val prove_equation: string -> mode option -> theory -> theory
val intro_rule: theory -> string -> mode -> thm
val elim_rule: theory -> string -> mode -> thm
- val strip_intro_concl : term -> int -> (term * (term list * term list))
- val code_ind_intros_attrib : attribute
- val code_ind_cases_attrib : attribute
+ val strip_intro_concl: term -> int -> term * (term list * term list)
val modename_of: theory -> string -> mode -> string
val modes_of: theory -> string -> mode list
- val setup : theory -> theory
- val code_pred : string -> Proof.context -> Proof.state
- val code_pred_cmd : string -> Proof.context -> Proof.state
- val print_alternative_rules : theory -> theory
+ val setup: theory -> theory
+ val code_pred: string -> Proof.context -> Proof.state
+ val code_pred_cmd: string -> Proof.context -> Proof.state
+ val print_alternative_rules: theory -> theory (*FIXME diagnostic command?*)
val do_proofs: bool ref
- val pred_intros : theory -> string -> thm list
- val get_nparams : theory -> string -> int
+ val pred_intros: theory -> string -> thm list
+ val get_nparams: theory -> string -> int
end;
-structure Predicate_Compile: PREDICATE_COMPILE =
+structure Predicate_Compile : PREDICATE_COMPILE =
struct
(** auxiliary **)
@@ -1224,7 +1221,7 @@
val Ts = binder_types T
val names = Name.variant_list []
(map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts)))
- val vs = map Free (names ~~ Ts)
+ val vs = map2 (curry Free) names Ts
val clausehd = HOLogic.mk_Trueprop (list_comb(Const (predname, T), vs))
val intro_t = Logic.mk_implies (@{prop False}, clausehd)
val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))
@@ -1242,9 +1239,9 @@
(* main function *********************************************************************)
(*************************************************************************************)
-fun create_def_equation' ind_name (mode : mode option) thy =
+fun prove_equation ind_name mode thy =
let
- val _ = tracing ("starting create_def_equation' with " ^ ind_name)
+ val _ = tracing ("starting prove_equation' with " ^ ind_name)
val (prednames, preds) =
case (try (InductivePackage.the_inductive (ProofContext.init thy)) ind_name) of
SOME info => let val preds = info |> snd |> #preds
@@ -1268,7 +1265,7 @@
fun rec_call name thy =
(*FIXME use member instead of infix mem*)
if not (name mem (Symtab.keys (#modes (IndCodegenData.get thy)))) then
- create_def_equation name thy else thy
+ prove_equation name NONE thy else thy
val thy'' = fold rec_call name_of_calls thy'
val _ = tracing "returning from recursive calls"
val _ = tracing "starting mode inference"
@@ -1322,12 +1319,11 @@
val _ = tracing "starting proof"
val result_thms = prove_preds thy''' all_vs param_vs (extra_modes @ modes) clauses (pred_mode ~~ (flat ts))
val (_, thy'''') = yield_singleton PureThy.add_thmss
- ((Binding.name (Long_Name.base_name ind_name ^ "_codegen" (*FIXME other suffix*)), result_thms),
+ ((Binding.qualify true (Long_Name.base_name ind_name) (Binding.name "equation"), result_thms),
[Attrib.attribute_i thy''' Code.add_default_eqn_attrib]) thy'''
in
thy''''
end
-and create_def_equation ind_name thy = create_def_equation' ind_name NONE thy
fun set_nparams (pred, nparams) thy = map_nparams (Symtab.update (pred, nparams)) thy
@@ -1345,22 +1341,12 @@
in () end
val _ = map print preds
in thy end;
-
-fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I)
-val code_ind_intros_attrib = attrib add_intro_thm
-
-val code_ind_cases_attrib = attrib add_elim_thm
-
-val setup =
- Attrib.setup @{binding code_ind_intros} (Scan.succeed code_ind_intros_attrib)
- "adding alternative introduction rules for code generation of inductive predicates" #>
- Attrib.setup @{binding code_ind_cases} (Scan.succeed code_ind_cases_attrib)
- "adding alternative elimination rules for code generation of inductive predicates";
(* generation of case rules from user-given introduction rules *)
- fun mk_casesrule introrules nparams ctxt = let
+fun mk_casesrule introrules nparams ctxt =
+ let
val intros = map prop_of introrules
val (pred, (params, args)) = strip_intro_concl (hd intros) nparams
val ([propname], ctxt1) = Variable.variant_fixes ["thesis"] ctxt
@@ -1368,6 +1354,7 @@
val (argnames, ctxt2) = Variable.variant_fixes
(map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt1
val argvs = map Free (argnames ~~ (map fastype_of args))
+ (*FIXME map2*)
fun mk_case intro = let
val (_, (_, args)) = strip_intro_concl intro nparams
val prems = Logic.strip_imp_prems intro
@@ -1384,32 +1371,59 @@
ctxt2
in (pred, prop, ctxt3) end;
-(* setup for user interface *)
+
+(** user interface **)
+
+local
+
+fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
+
+val add_elim_attrib = attrib add_elim_thm;
- fun generic_code_pred prep_const raw_const lthy =
- let
- val thy = (ProofContext.theory_of lthy)
- val const = prep_const thy raw_const
- val nparams = get_nparams thy const
- val intro_rules = pred_intros thy const
- val (((tfrees, frees), fact), lthy') =
- Variable.import_thms true intro_rules lthy;
- val (pred, prop, lthy'') = mk_casesrule fact nparams lthy'
- val (predname, _) = dest_Const pred
- fun after_qed [[th]] lthy'' =
- LocalTheory.note Thm.theoremK
- ((Binding.name (Long_Name.base_name predname ^ "_cases"), (* FIXME: other suffix *)
- [Attrib.internal (K (code_ind_cases_attrib))]) , [th]) lthy''
- |> snd
- |> LocalTheory.theory (create_def_equation predname)
- in
- Proof.theorem_i NONE after_qed [[(prop, [])]] lthy''
- end;
+fun generic_code_pred prep_const raw_const lthy =
+ let
+ val thy = ProofContext.theory_of lthy
+ val const = prep_const thy raw_const
+ val nparams = get_nparams thy const
+ val intro_rules = pred_intros thy const
+ val (((tfrees, frees), fact), lthy') =
+ Variable.import_thms true intro_rules lthy;
+ val (pred, prop, lthy'') = mk_casesrule fact nparams lthy'
+ val (predname, _) = dest_Const pred
+ fun after_qed [[th]] lthy'' =
+ lthy''
+ |> LocalTheory.note Thm.theoremK
+ ((Binding.empty, [Attrib.internal (K add_elim_attrib)]), [th])
+ |> snd
+ |> LocalTheory.theory (prove_equation predname NONE)
+ in
+ Proof.theorem_i NONE after_qed [[(prop, [])]] lthy''
+ end;
+
+structure P = OuterParse
- val code_pred = generic_code_pred (K I);
- val code_pred_cmd = generic_code_pred Code_Unit.read_const
+in
+
+val code_pred = generic_code_pred (K I);
+val code_pred_cmd = generic_code_pred Code.read_const
+
+val setup =
+ Attrib.setup @{binding code_ind_intros} (Scan.succeed (attrib add_intro_thm))
+ "adding alternative introduction rules for code generation of inductive predicates" #>
+ Attrib.setup @{binding code_ind_cases} (Scan.succeed add_elim_attrib)
+ "adding alternative elimination rules for code generation of inductive predicates";
+ (*FIXME name discrepancy in attribs and ML code*)
+ (*FIXME intros should be better named intro*)
+ (*FIXME why distinguished atribute for cases?*)
+
+val _ = OuterSyntax.local_theory_to_proof "code_pred"
+ "prove equations for predicate specified by intro/elim rules"
+ OuterKeyword.thy_goal (P.term_group >> code_pred_cmd)
+
+end
+
+(*FIXME
+- Naming of auxiliary rules necessary?
+*)
end;
-
-fun pred_compile name thy = Predicate_Compile.create_def_equation
- (Sign.intern_const thy name) thy;
--- a/src/Pure/IsaMakefile Fri May 15 15:12:23 2009 -0700
+++ b/src/Pure/IsaMakefile Fri May 15 15:13:09 2009 -0700
@@ -56,7 +56,7 @@
General/table.ML General/url.ML General/xml.ML General/yxml.ML \
Isar/ROOT.ML Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML \
Isar/calculation.ML Isar/class.ML Isar/class_target.ML Isar/code.ML \
- Isar/code_unit.ML Isar/constdefs.ML Isar/context_rules.ML \
+ Isar/constdefs.ML Isar/context_rules.ML \
Isar/element.ML Isar/expression.ML Isar/isar_cmd.ML \
Isar/isar_document.ML Isar/isar_syn.ML Isar/local_defs.ML \
Isar/local_syntax.ML Isar/local_theory.ML Isar/locale.ML \
--- a/src/Pure/Isar/ROOT.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/Pure/Isar/ROOT.ML Fri May 15 15:13:09 2009 -0700
@@ -61,7 +61,6 @@
use "../simplifier.ML";
(*executable theory content*)
-use "code_unit.ML";
use "code.ML";
(*specifications*)
--- a/src/Pure/Isar/code.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/Pure/Isar/code.ML Fri May 15 15:13:09 2009 -0700
@@ -7,6 +7,55 @@
signature CODE =
sig
+ (*constructor sets*)
+ val constrset_of_consts: theory -> (string * typ) list
+ -> string * ((string * sort) list * (string * typ list) list)
+
+ (*typ instantiations*)
+ val typscheme: theory -> string * typ -> (string * sort) list * typ
+ val inst_thm: theory -> sort Vartab.table -> thm -> thm
+
+ (*constants*)
+ val string_of_typ: theory -> typ -> string
+ val string_of_const: theory -> string -> string
+ val no_args: theory -> string -> int
+ val check_const: theory -> term -> string
+ val read_bare_const: theory -> string -> string * typ
+ val read_const: theory -> string -> string
+
+ (*constant aliasses*)
+ val add_const_alias: thm -> theory -> theory
+ val triv_classes: theory -> class list
+ val resubst_alias: theory -> string -> string
+
+ (*code equations*)
+ val mk_eqn: theory -> (string -> bool) -> thm * bool -> thm * bool
+ val mk_eqn_liberal: theory -> (string -> bool) -> thm -> (thm * bool) option
+ val assert_eqn: theory -> thm * bool -> thm * bool
+ val assert_eqns_const: theory -> string
+ -> (thm * bool) list -> (thm * bool) list
+ val const_typ_eqn: thm -> string * typ
+ val const_eqn: theory -> thm -> string
+ val typscheme_eqn: theory -> thm -> (string * sort) list * typ
+ val expand_eta: theory -> int -> thm -> thm
+ val rewrite_eqn: simpset -> thm -> thm
+ val rewrite_head: thm list -> thm -> thm
+ val norm_args: theory -> thm list -> thm list
+ val norm_varnames: theory -> thm list -> thm list
+
+ (*case certificates*)
+ val case_cert: thm -> string * (int * string list)
+
+ (*infrastructure*)
+ val add_attribute: string * attribute parser -> theory -> theory
+ val purge_data: theory -> theory
+
+ (*executable content*)
+ val add_datatype: (string * typ) list -> theory -> theory
+ val add_datatype_cmd: string list -> theory -> theory
+ val type_interpretation:
+ (string * ((string * sort) list * (string * typ list) list)
+ -> theory -> theory) -> theory -> theory
val add_eqn: thm -> theory -> theory
val add_nbe_eqn: thm -> theory -> theory
val add_default_eqn: thm -> theory -> theory
@@ -15,37 +64,16 @@
val del_eqn: thm -> theory -> theory
val del_eqns: string -> theory -> theory
val add_eqnl: string * (thm * bool) list lazy -> theory -> theory
- val map_pre: (simpset -> simpset) -> theory -> theory
- val map_post: (simpset -> simpset) -> theory -> theory
- val add_inline: thm -> theory -> theory
- val add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory
- val del_functrans: string -> theory -> theory
- val simple_functrans: (theory -> thm list -> thm list option)
- -> theory -> (thm * bool) list -> (thm * bool) list option
- val add_datatype: (string * typ) list -> theory -> theory
- val add_datatype_cmd: string list -> theory -> theory
- val type_interpretation:
- (string * ((string * sort) list * (string * typ list) list)
- -> theory -> theory) -> theory -> theory
val add_case: thm -> theory -> theory
val add_undefined: string -> theory -> theory
- val purge_data: theory -> theory
- val these_eqns: theory -> string -> (thm * bool) list
- val these_raw_eqns: theory -> string -> (thm * bool) list
+ (*data retrieval*)
val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list)
val get_datatype_of_constr: theory -> string -> string option
+ val default_typscheme: theory -> string -> (string * sort) list * typ
+ val these_eqns: theory -> string -> (thm * bool) list
val get_case_scheme: theory -> string -> (int * (int * string list)) option
val is_undefined: theory -> string -> bool
- val default_typscheme: theory -> string -> (string * sort) list * typ
-
- val preprocess_conv: theory -> cterm -> thm
- val preprocess_term: theory -> term -> term
- val postprocess_conv: theory -> cterm -> thm
- val postprocess_term: theory -> term -> term
-
- val add_attribute: string * attribute parser -> theory -> theory
-
val print_codesetup: theory -> unit
end;
@@ -80,6 +108,400 @@
structure Code : PRIVATE_CODE =
struct
+(* auxiliary *)
+
+fun string_of_typ thy = setmp show_sorts true (Syntax.string_of_typ_global thy);
+fun string_of_const thy c = case AxClass.inst_of_param thy c
+ of SOME (c, tyco) => Sign.extern_const thy c ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco)
+ | NONE => Sign.extern_const thy c;
+
+fun no_args thy = length o fst o strip_type o Sign.the_const_type thy;
+
+
+(* utilities *)
+
+fun typscheme thy (c, ty) =
+ let
+ val ty' = Logic.unvarifyT ty;
+ fun dest (TFree (v, sort)) = (v, sort)
+ | dest ty = error ("Illegal type parameter in type scheme: " ^ Syntax.string_of_typ_global thy ty);
+ val vs = map dest (Sign.const_typargs thy (c, ty'));
+ in (vs, Type.strip_sorts ty') end;
+
+fun inst_thm thy tvars' thm =
+ let
+ val tvars = (Term.add_tvars o Thm.prop_of) thm [];
+ val inter_sort = Sorts.inter_sort (Sign.classes_of thy);
+ fun mk_inst (tvar as (v, sort)) = case Vartab.lookup tvars' v
+ of SOME sort' => SOME (pairself (Thm.ctyp_of thy o TVar)
+ (tvar, (v, inter_sort (sort, sort'))))
+ | NONE => NONE;
+ val insts = map_filter mk_inst tvars;
+ in Thm.instantiate (insts, []) thm end;
+
+fun expand_eta thy k thm =
+ let
+ val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm;
+ val (head, args) = strip_comb lhs;
+ val l = if k = ~1
+ then (length o fst o strip_abs) rhs
+ else Int.max (0, k - length args);
+ val used = Name.make_context (map (fst o fst) (Term.add_vars lhs []));
+ fun get_name _ 0 = pair []
+ | get_name (Abs (v, ty, t)) k =
+ Name.variants [v]
+ ##>> get_name t (k - 1)
+ #>> (fn ([v'], vs') => (v', ty) :: vs')
+ | get_name t k =
+ let
+ val (tys, _) = (strip_type o fastype_of) t
+ in case tys
+ of [] => raise TERM ("expand_eta", [t])
+ | ty :: _ =>
+ Name.variants [""]
+ #-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1)
+ #>> (fn vs' => (v, ty) :: vs'))
+ end;
+ val (vs, _) = get_name rhs l used;
+ fun expand (v, ty) thm = Drule.fun_cong_rule thm
+ (Thm.cterm_of thy (Var ((v, 0), ty)));
+ in
+ thm
+ |> fold expand vs
+ |> Conv.fconv_rule Drule.beta_eta_conversion
+ end;
+
+fun eqn_conv conv =
+ let
+ fun lhs_conv ct = if can Thm.dest_comb ct
+ then (Conv.combination_conv lhs_conv conv) ct
+ else Conv.all_conv ct;
+ in Conv.combination_conv (Conv.arg_conv lhs_conv) conv end;
+
+fun head_conv conv =
+ let
+ fun lhs_conv ct = if can Thm.dest_comb ct
+ then (Conv.fun_conv lhs_conv) ct
+ else conv ct;
+ in Conv.fun_conv (Conv.arg_conv lhs_conv) end;
+
+val rewrite_eqn = Conv.fconv_rule o eqn_conv o Simplifier.rewrite;
+val rewrite_head = Conv.fconv_rule o head_conv o MetaSimplifier.rewrite false;
+
+fun norm_args thy thms =
+ let
+ val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals;
+ val k = fold (curry Int.max o num_args_of o Thm.prop_of) thms 0;
+ in
+ thms
+ |> map (expand_eta thy k)
+ |> map (Conv.fconv_rule Drule.beta_eta_conversion)
+ end;
+
+fun canonical_tvars thy thm =
+ let
+ val ctyp = Thm.ctyp_of thy;
+ val purify_tvar = unprefix "'" #> Name.desymbolize false #> prefix "'";
+ fun tvars_subst_for thm = (fold_types o fold_atyps)
+ (fn TVar (v_i as (v, _), sort) => let
+ val v' = purify_tvar v
+ in if v = v' then I
+ else insert (op =) (v_i, (v', sort)) end
+ | _ => I) (prop_of thm) [];
+ fun mk_inst (v_i, (v', sort)) (maxidx, acc) =
+ let
+ val ty = TVar (v_i, sort)
+ in
+ (maxidx + 1, (ctyp ty, ctyp (TVar ((v', maxidx), sort))) :: acc)
+ end;
+ val maxidx = Thm.maxidx_of thm + 1;
+ val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []);
+ in Thm.instantiate (inst, []) thm end;
+
+fun canonical_vars thy thm =
+ let
+ val cterm = Thm.cterm_of thy;
+ val purify_var = Name.desymbolize false;
+ fun vars_subst_for thm = fold_aterms
+ (fn Var (v_i as (v, _), ty) => let
+ val v' = purify_var v
+ in if v = v' then I
+ else insert (op =) (v_i, (v', ty)) end
+ | _ => I) (prop_of thm) [];
+ fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) =
+ let
+ val t = Var (v_i, ty)
+ in
+ (maxidx + 1, (cterm t, cterm (Var ((v', maxidx), ty))) :: acc)
+ end;
+ val maxidx = Thm.maxidx_of thm + 1;
+ val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []);
+ in Thm.instantiate ([], inst) thm end;
+
+fun canonical_absvars thm =
+ let
+ val t = Thm.plain_prop_of thm;
+ val purify_var = Name.desymbolize false;
+ val t' = Term.map_abs_vars purify_var t;
+ in Thm.rename_boundvars t t' thm end;
+
+fun norm_varnames thy thms =
+ let
+ fun burrow_thms f [] = []
+ | burrow_thms f thms =
+ thms
+ |> Conjunction.intr_balanced
+ |> f
+ |> Conjunction.elim_balanced (length thms)
+ in
+ thms
+ |> map (canonical_vars thy)
+ |> map canonical_absvars
+ |> map Drule.zero_var_indexes
+ |> burrow_thms (canonical_tvars thy)
+ |> Drule.zero_var_indexes_list
+ end;
+
+
+(* const aliasses *)
+
+structure ConstAlias = TheoryDataFun
+(
+ type T = ((string * string) * thm) list * class list;
+ val empty = ([], []);
+ val copy = I;
+ val extend = I;
+ fun merge _ ((alias1, classes1), (alias2, classes2)) : T =
+ (Library.merge (eq_snd Thm.eq_thm_prop) (alias1, alias2),
+ Library.merge (op =) (classes1, classes2));
+);
+
+fun add_const_alias thm thy =
+ let
+ val lhs_rhs = case try Logic.dest_equals (Thm.prop_of thm)
+ of SOME lhs_rhs => lhs_rhs
+ | _ => error ("Not an equation: " ^ Display.string_of_thm thm);
+ val c_c' = case try (pairself (AxClass.unoverload_const thy o dest_Const)) lhs_rhs
+ of SOME c_c' => c_c'
+ | _ => error ("Not an equation with two constants: " ^ Display.string_of_thm thm);
+ val some_class = the_list (AxClass.class_of_param thy (snd c_c'));
+ in thy |>
+ ConstAlias.map (fn (alias, classes) =>
+ ((c_c', thm) :: alias, fold (insert (op =)) some_class classes))
+ end;
+
+fun resubst_alias thy =
+ let
+ val alias = fst (ConstAlias.get thy);
+ val subst_inst_param = Option.map fst o AxClass.inst_of_param thy;
+ fun subst_alias c =
+ get_first (fn ((c', c''), _) => if c = c'' then SOME c' else NONE) alias;
+ in
+ perhaps subst_inst_param
+ #> perhaps subst_alias
+ end;
+
+val triv_classes = snd o ConstAlias.get;
+
+
+(* reading constants as terms *)
+
+fun check_bare_const thy t = case try dest_Const t
+ of SOME c_ty => c_ty
+ | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
+
+fun check_const thy = AxClass.unoverload_const thy o check_bare_const thy;
+
+fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy;
+
+fun read_const thy = AxClass.unoverload_const thy o read_bare_const thy;
+
+
+(* constructor sets *)
+
+fun constrset_of_consts thy cs =
+ let
+ val _ = map (fn (c, _) => if (is_some o AxClass.class_of_param thy) c
+ then error ("Is a class parameter: " ^ string_of_const thy c) else ()) cs;
+ fun no_constr (c, ty) = error ("Not a datatype constructor: " ^ string_of_const thy c
+ ^ " :: " ^ string_of_typ thy ty);
+ fun last_typ c_ty ty =
+ let
+ val frees = OldTerm.typ_tfrees ty;
+ val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) ty
+ handle TYPE _ => no_constr c_ty
+ val _ = if has_duplicates (eq_fst (op =)) vs then no_constr c_ty else ();
+ val _ = if length frees <> length vs then no_constr c_ty else ();
+ in (tyco, vs) end;
+ fun ty_sorts (c, ty) =
+ let
+ val ty_decl = (Logic.unvarifyT o Sign.the_const_type thy) c;
+ val (tyco, _) = last_typ (c, ty) ty_decl;
+ val (_, vs) = last_typ (c, ty) ty;
+ in ((tyco, map snd vs), (c, (map fst vs, ty))) end;
+ fun add ((tyco', sorts'), c) ((tyco, sorts), cs) =
+ let
+ val _ = if tyco' <> tyco
+ then error "Different type constructors in constructor set"
+ else ();
+ val sorts'' = map2 (curry (Sorts.inter_sort (Sign.classes_of thy))) sorts' sorts
+ in ((tyco, sorts), c :: cs) end;
+ fun inst vs' (c, (vs, ty)) =
+ let
+ val the_v = the o AList.lookup (op =) (vs ~~ vs');
+ val ty' = map_atyps (fn TFree (v, _) => TFree (the_v v)) ty;
+ in (c, (fst o strip_type) ty') end;
+ val c' :: cs' = map ty_sorts cs;
+ val ((tyco, sorts), cs'') = fold add cs' (apsnd single c');
+ val vs = Name.names Name.context Name.aT sorts;
+ val cs''' = map (inst vs) cs'';
+ in (tyco, (vs, rev cs''')) end;
+
+
+(* code equations *)
+
+exception BAD_THM of string;
+fun bad_thm msg = raise BAD_THM msg;
+fun error_thm f thm = f thm handle BAD_THM msg => error msg;
+fun try_thm f thm = SOME (f thm) handle BAD_THM _ => NONE;
+
+fun is_linear thm =
+ let val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm
+ in not (has_duplicates (op =) ((fold o fold_aterms)
+ (fn Var (v, _) => cons v | _ => I) args [])) end;
+
+fun gen_assert_eqn thy is_constr_head is_constr_pat (thm, proper) =
+ let
+ val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm
+ handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm)
+ | THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm);
+ fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v
+ | Free _ => bad_thm ("Illegal free variable in equation\n"
+ ^ Display.string_of_thm thm)
+ | _ => I) t [];
+ fun tvars_of t = fold_term_types (fn _ =>
+ fold_atyps (fn TVar (v, _) => insert (op =) v
+ | TFree _ => bad_thm
+ ("Illegal free type variable in equation\n" ^ Display.string_of_thm thm))) t [];
+ val lhs_vs = vars_of lhs;
+ val rhs_vs = vars_of rhs;
+ val lhs_tvs = tvars_of lhs;
+ val rhs_tvs = tvars_of rhs;
+ val _ = if null (subtract (op =) lhs_vs rhs_vs)
+ then ()
+ else bad_thm ("Free variables on right hand side of equation\n"
+ ^ Display.string_of_thm thm);
+ val _ = if null (subtract (op =) lhs_tvs rhs_tvs)
+ then ()
+ else bad_thm ("Free type variables on right hand side of equation\n"
+ ^ Display.string_of_thm thm) val (head, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm;
+ val (c, ty) = case head
+ of Const (c_ty as (_, ty)) => (AxClass.unoverload_const thy c_ty, ty)
+ | _ => bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm thm);
+ fun check _ (Abs _) = bad_thm
+ ("Abstraction on left hand side of equation\n"
+ ^ Display.string_of_thm thm)
+ | check 0 (Var _) = ()
+ | check _ (Var _) = bad_thm
+ ("Variable with application on left hand side of equation\n"
+ ^ Display.string_of_thm thm)
+ | check n (t1 $ t2) = (check (n+1) t1; check 0 t2)
+ | check n (Const (c_ty as (c, ty))) = if n = (length o fst o strip_type) ty
+ then if not proper orelse is_constr_pat (AxClass.unoverload_const thy c_ty)
+ then ()
+ else bad_thm (quote c ^ " is not a constructor, on left hand side of equation\n"
+ ^ Display.string_of_thm thm)
+ else bad_thm
+ ("Partially applied constant " ^ quote c ^ " on left hand side of equation\n"
+ ^ Display.string_of_thm thm);
+ val _ = map (check 0) args;
+ val _ = if not proper orelse is_linear thm then ()
+ else bad_thm ("Duplicate variables on left hand side of equation\n"
+ ^ Display.string_of_thm thm);
+ val _ = if (is_none o AxClass.class_of_param thy) c
+ then ()
+ else bad_thm ("Polymorphic constant as head in equation\n"
+ ^ Display.string_of_thm thm)
+ val _ = if not (is_constr_head c)
+ then ()
+ else bad_thm ("Constructor as head in equation\n"
+ ^ Display.string_of_thm thm)
+ val ty_decl = Sign.the_const_type thy c;
+ val _ = if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty)
+ then () else bad_thm ("Type\n" ^ string_of_typ thy ty
+ ^ "\nof equation\n"
+ ^ Display.string_of_thm thm
+ ^ "\nis incompatible with declared function type\n"
+ ^ string_of_typ thy ty_decl)
+ in (thm, proper) end;
+
+fun assert_eqn thy is_constr = error_thm (gen_assert_eqn thy is_constr is_constr);
+
+val const_typ_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
+
+
+(*those following are permissive wrt. to overloaded constants!*)
+
+fun mk_eqn thy is_constr_head = error_thm (gen_assert_eqn thy is_constr_head (K true)) o
+ apfst (LocalDefs.meta_rewrite_rule (ProofContext.init thy));
+
+fun mk_eqn_liberal thy is_constr_head = Option.map (fn (thm, _) => (thm, is_linear thm))
+ o try_thm (gen_assert_eqn thy is_constr_head (K true))
+ o rpair false o LocalDefs.meta_rewrite_rule (ProofContext.init thy);
+
+fun const_typ_eqn_unoverload thy thm =
+ let
+ val (c, ty) = const_typ_eqn thm;
+ val c' = AxClass.unoverload_const thy (c, ty);
+ in (c', ty) end;
+
+fun typscheme_eqn thy = typscheme thy o const_typ_eqn_unoverload thy;
+fun const_eqn thy = fst o const_typ_eqn_unoverload thy;
+
+
+(* case cerificates *)
+
+fun case_certificate thm =
+ let
+ val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals
+ o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.prop_of) thm;
+ val _ = case head of Free _ => true
+ | Var _ => true
+ | _ => raise TERM ("case_cert", []);
+ val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr;
+ val (Const (case_const, _), raw_params) = strip_comb case_expr;
+ val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params;
+ val _ = if n = ~1 then raise TERM ("case_cert", []) else ();
+ val params = map (fst o dest_Var) (nth_drop n raw_params);
+ fun dest_case t =
+ let
+ val (head' $ t_co, rhs) = Logic.dest_equals t;
+ val _ = if head' = head then () else raise TERM ("case_cert", []);
+ val (Const (co, _), args) = strip_comb t_co;
+ val (Var (param, _), args') = strip_comb rhs;
+ val _ = if args' = args then () else raise TERM ("case_cert", []);
+ in (param, co) end;
+ fun analyze_cases cases =
+ let
+ val co_list = fold (AList.update (op =) o dest_case) cases [];
+ in map (the o AList.lookup (op =) co_list) params end;
+ fun analyze_let t =
+ let
+ val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t;
+ val _ = if head' = head then () else raise TERM ("case_cert", []);
+ val _ = if arg' = arg then () else raise TERM ("case_cert", []);
+ val _ = if [param'] = params then () else raise TERM ("case_cert", []);
+ in [] end;
+ fun analyze (cases as [let_case]) =
+ (analyze_cases cases handle Bind => analyze_let let_case)
+ | analyze cases = analyze_cases cases;
+ in (case_const, (n, analyze cases)) end;
+
+fun case_cert thm = case_certificate thm
+ handle Bind => error "bad case certificate"
+ | TERM _ => error "bad case certificate";
+
+
(** code attributes **)
structure CodeAttr = TheoryDataFun (
@@ -126,7 +548,8 @@
fun add_drop_redundant thy (thm, proper) thms =
let
- val args_of = snd o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
+ val args_of = snd o strip_comb o map_types Type.strip_sorts
+ o fst o Logic.dest_equals o Thm.plain_prop_of;
val args = args_of thm;
val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1);
fun matches_args args' = length args <= length args' andalso
@@ -159,6 +582,8 @@
fun mk_spec ((concluded_history, eqns), (dtyps, cases)) =
Spec { concluded_history = concluded_history, eqns = eqns, dtyps = dtyps, cases = cases };
+val empty_spec =
+ mk_spec ((false, Symtab.empty), (Symtab.empty, (Symtab.empty, Symtab.empty)));
fun map_spec f (Spec { concluded_history = concluded_history, eqns = eqns,
dtyps = dtyps, cases = cases }) =
mk_spec (f ((concluded_history, eqns), (dtyps, cases)));
@@ -167,7 +592,8 @@
let
fun merge_eqns ((_, history1), (_, history2)) =
let
- val raw_history = AList.merge (op =) (K true) (history1, history2)
+ val raw_history = AList.merge (op = : serial * serial -> bool)
+ (K true) (history1, history2)
val filtered_history = filter_out (fst o snd) raw_history
val history = if null filtered_history
then raw_history else filtered_history;
@@ -179,57 +605,16 @@
in mk_spec ((false, eqns), (dtyps, cases)) end;
-(* pre- and postprocessor *)
-
-datatype thmproc = Thmproc of {
- pre: simpset,
- post: simpset,
- functrans: (string * (serial * (theory -> (thm * bool) list -> (thm * bool) list option))) list
-};
-
-fun mk_thmproc ((pre, post), functrans) =
- Thmproc { pre = pre, post = post, functrans = functrans };
-fun map_thmproc f (Thmproc { pre, post, functrans }) =
- mk_thmproc (f ((pre, post), functrans));
-fun merge_thmproc (Thmproc { pre = pre1, post = post1, functrans = functrans1 },
- Thmproc { pre = pre2, post = post2, functrans = functrans2 }) =
- let
- val pre = Simplifier.merge_ss (pre1, pre2);
- val post = Simplifier.merge_ss (post1, post2);
- val functrans = AList.merge (op =) (eq_fst (op =)) (functrans1, functrans2);
- in mk_thmproc ((pre, post), functrans) end;
-
-datatype exec = Exec of {
- thmproc: thmproc,
- spec: spec
-};
-
-
(* code setup data *)
-fun mk_exec (thmproc, spec) =
- Exec { thmproc = thmproc, spec = spec };
-fun map_exec f (Exec { thmproc = thmproc, spec = spec }) =
- mk_exec (f (thmproc, spec));
-fun merge_exec (Exec { thmproc = thmproc1, spec = spec1 },
- Exec { thmproc = thmproc2, spec = spec2 }) =
- let
- val thmproc = merge_thmproc (thmproc1, thmproc2);
- val spec = merge_spec (spec1, spec2);
- in mk_exec (thmproc, spec) end;
-val empty_exec = mk_exec (mk_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []),
- mk_spec ((false, Symtab.empty), (Symtab.empty, (Symtab.empty, Symtab.empty))));
-
-fun the_thmproc (Exec { thmproc = Thmproc x, ...}) = x;
-fun the_spec (Exec { spec = Spec x, ...}) = x;
+fun the_spec (Spec x) = x;
val the_eqns = #eqns o the_spec;
val the_dtyps = #dtyps o the_spec;
val the_cases = #cases o the_spec;
-val map_thmproc = map_exec o apfst o map_thmproc;
-val map_concluded_history = map_exec o apsnd o map_spec o apfst o apfst;
-val map_eqns = map_exec o apsnd o map_spec o apfst o apsnd;
-val map_dtyps = map_exec o apsnd o map_spec o apsnd o apfst;
-val map_cases = map_exec o apsnd o map_spec o apsnd o apsnd;
+val map_concluded_history = map_spec o apfst o apfst;
+val map_eqns = map_spec o apfst o apsnd;
+val map_dtyps = map_spec o apsnd o apfst;
+val map_cases = map_spec o apsnd o apsnd;
(* data slots dependent on executable content *)
@@ -277,17 +662,17 @@
type data = Object.T Datatab.table;
val empty_data = Datatab.empty : data;
-structure CodeData = TheoryDataFun
+structure Code_Data = TheoryDataFun
(
- type T = exec * data ref;
- val empty = (empty_exec, ref empty_data);
- fun copy (exec, data) = (exec, ref (! data));
+ type T = spec * data ref;
+ val empty = (empty_spec, ref empty_data);
+ fun copy (spec, data) = (spec, ref (! data));
val extend = copy;
- fun merge pp ((exec1, data1), (exec2, data2)) =
- (merge_exec (exec1, exec2), ref empty_data);
+ fun merge pp ((spec1, data1), (spec2, data2)) =
+ (merge_spec (spec1, spec2), ref empty_data);
);
-fun thy_data f thy = f ((snd o CodeData.get) thy);
+fun thy_data f thy = f ((snd o Code_Data.get) thy);
fun get_ensure_init kind data_ref =
case Datatab.lookup (! data_ref) kind
@@ -299,7 +684,7 @@
(* access to executable content *)
-val the_exec = fst o CodeData.get;
+val the_exec = fst o Code_Data.get;
fun complete_class_params thy cs =
fold (fn c => case AxClass.inst_of_param thy c
@@ -307,11 +692,11 @@
| SOME (c', _) => insert (op =) c' #> insert (op =) c) cs [];
fun map_exec_purge touched f thy =
- CodeData.map (fn (exec, data) => (f exec, ref (case touched
+ Code_Data.map (fn (exec, data) => (f exec, ref (case touched
of SOME cs => invoke_purge_all thy (complete_class_params thy cs) (! data)
| NONE => empty_data))) thy;
-val purge_data = (CodeData.map o apsnd) (K (ref empty_data));
+val purge_data = (Code_Data.map o apsnd) (K (ref empty_data));
(* tackling equation history *)
@@ -323,21 +708,21 @@
fun continue_history thy = if (#concluded_history o the_spec o the_exec) thy
then thy
- |> (CodeData.map o apfst o map_concluded_history) (K false)
+ |> (Code_Data.map o apfst o map_concluded_history) (K false)
|> SOME
else NONE;
fun conclude_history thy = if (#concluded_history o the_spec o the_exec) thy
then NONE
else thy
- |> (CodeData.map o apfst)
+ |> (Code_Data.map o apfst)
((map_eqns o Symtab.map) (fn ((changed, current), history) =>
((false, current),
if changed then (serial (), current) :: history else history))
#> map_concluded_history (K true))
|> SOME;
-val _ = Context.>> (Context.map_theory (CodeData.init
+val _ = Context.>> (Context.map_theory (Code_Data.init
#> Theory.at_begin continue_history
#> Theory.at_end conclude_history));
@@ -366,9 +751,6 @@
end; (*local*)
-
-(* print executable content *)
-
fun print_codesetup thy =
let
val ctxt = ProofContext.init thy;
@@ -383,19 +765,16 @@
(Pretty.block o Pretty.breaks) (
Pretty.str s
:: Pretty.str "="
- :: separate (Pretty.str "|") (map (fn (c, []) => Pretty.str (Code_Unit.string_of_const thy c)
+ :: separate (Pretty.str "|") (map (fn (c, []) => Pretty.str (string_of_const thy c)
| (c, tys) =>
(Pretty.block o Pretty.breaks)
- (Pretty.str (Code_Unit.string_of_const thy c)
+ (Pretty.str (string_of_const thy c)
:: Pretty.str "of"
:: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos)
);
- val pre = (#pre o the_thmproc) exec;
- val post = (#post o the_thmproc) exec;
- val functrans = (map fst o #functrans o the_thmproc) exec;
val eqns = the_eqns exec
|> Symtab.dest
- |> (map o apfst) (Code_Unit.string_of_const thy)
+ |> (map o apfst) (string_of_const thy)
|> (map o apsnd) (snd o fst)
|> sort (string_ord o pairself fst);
val dtyps = the_dtyps exec
@@ -410,21 +789,6 @@
:: Pretty.fbrk
:: (Pretty.fbreaks o map pretty_eqn) eqns
),
- Pretty.block [
- Pretty.str "preprocessing simpset:",
- Pretty.fbrk,
- Simplifier.pretty_ss ctxt pre
- ],
- Pretty.block [
- Pretty.str "postprocessing simpset:",
- Pretty.fbrk,
- Simplifier.pretty_ss ctxt post
- ],
- Pretty.block (
- Pretty.str "function transformers:"
- :: Pretty.fbrk
- :: (Pretty.fbreaks o map Pretty.str) functrans
- ),
Pretty.block (
Pretty.str "datatypes:"
:: Pretty.fbrk
@@ -446,13 +810,13 @@
val max' = Thm.maxidx_of thm' + 1;
in (thm', max') end;
val (thms', maxidx) = fold_map incr_thm thms 0;
- val ty1 :: tys = map (snd o Code_Unit.const_typ_eqn) thms';
+ val ty1 :: tys = map (snd o const_typ_eqn) thms';
fun unify ty env = Sign.typ_unify thy (ty1, ty) env
handle Type.TUNIFY =>
error ("Type unificaton failed, while unifying code equations\n"
^ (cat_lines o map Display.string_of_thm) thms
^ "\nwith types\n"
- ^ (cat_lines o map (Code_Unit.string_of_typ thy)) (ty1 :: tys));
+ ^ (cat_lines o map (string_of_typ thy)) (ty1 :: tys));
val (env, _) = fold unify tys (Vartab.empty, maxidx)
val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
@@ -461,10 +825,6 @@
(** interfaces and attributes **)
-fun delete_force msg key xs =
- if AList.defined (op =) xs key then AList.delete (op =) key xs
- else error ("No such " ^ msg ^ ": " ^ quote key);
-
fun get_datatype thy tyco =
case these (Symtab.lookup ((the_dtyps o the_exec) thy) tyco)
of (_, spec) :: _ => spec
@@ -481,13 +841,13 @@
fun is_constr thy = is_some o get_datatype_of_constr thy;
-fun assert_eqn thy = Code_Unit.assert_eqn thy (is_constr thy);
+val assert_eqn = fn thy => assert_eqn thy (is_constr thy);
fun assert_eqns_const thy c eqns =
let
- fun cert (eqn as (thm, _)) = if c = Code_Unit.const_eqn thy thm
+ fun cert (eqn as (thm, _)) = if c = const_eqn thy thm
then eqn else error ("Wrong head of code equation,\nexpected constant "
- ^ Code_Unit.string_of_const thy c ^ "\n" ^ Display.string_of_thm thm)
+ ^ string_of_const thy c ^ "\n" ^ Display.string_of_thm thm)
in map (cert o assert_eqn thy) eqns end;
fun change_eqns delete c f = (map_exec_purge (SOME [c]) o map_eqns
@@ -495,19 +855,19 @@
o apfst) (fn (_, eqns) => (true, f eqns));
fun gen_add_eqn default (eqn as (thm, _)) thy =
- let val c = Code_Unit.const_eqn thy thm
+ let val c = const_eqn thy thm
in change_eqns false c (add_thm thy default eqn) thy end;
fun add_eqn thm thy =
- gen_add_eqn false (Code_Unit.mk_eqn thy (is_constr thy) (thm, true)) thy;
+ gen_add_eqn false (mk_eqn thy (is_constr thy) (thm, true)) thy;
fun add_default_eqn thm thy =
- case Code_Unit.mk_eqn_liberal thy (is_constr thy) thm
+ case mk_eqn_liberal thy (is_constr thy) thm
of SOME eqn => gen_add_eqn true eqn thy
| NONE => thy;
fun add_nbe_eqn thm thy =
- gen_add_eqn false (Code_Unit.mk_eqn thy (is_constr thy) (thm, false)) thy;
+ gen_add_eqn false (mk_eqn thy (is_constr thy) (thm, false)) thy;
fun add_eqnl (c, lthms) thy =
let
@@ -518,8 +878,8 @@
(fn thm => Context.mapping (add_default_eqn thm) I);
val add_default_eqn_attrib = Attrib.internal (K add_default_eqn_attribute);
-fun del_eqn thm thy = case Code_Unit.mk_eqn_liberal thy (is_constr thy) thm
- of SOME (thm, _) => change_eqns true (Code_Unit.const_eqn thy thm) (del_thm thm) thy
+fun del_eqn thm thy = case mk_eqn_liberal thy (is_constr thy) thm
+ of SOME (thm, _) => change_eqns true (const_eqn thy thm) (del_thm thm) thy
| NONE => thy;
fun del_eqns c = change_eqns true c (K (false, Lazy.value []));
@@ -533,7 +893,7 @@
fun add_datatype raw_cs thy =
let
val cs = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) raw_cs;
- val (tyco, vs_cos) = Code_Unit.constrset_of_consts thy cs;
+ val (tyco, vs_cos) = constrset_of_consts thy cs;
val old_cs = (map fst o snd o get_datatype thy) tyco;
fun drop_outdated_cases cases = fold Symtab.delete_safe
(Symtab.fold (fn (c, (_, (_, cos))) =>
@@ -553,12 +913,12 @@
fun add_datatype_cmd raw_cs thy =
let
- val cs = map (Code_Unit.read_bare_const thy) raw_cs;
+ val cs = map (read_bare_const thy) raw_cs;
in add_datatype cs thy end;
fun add_case thm thy =
let
- val (c, (k, case_pats)) = Code_Unit.case_cert thm;
+ val (c, (k, case_pats)) = case_cert thm;
val _ = case filter_out (is_constr thy) case_pats
of [] => ()
| cs => error ("Non-constructor(s) in case certificate: " ^ commas (map quote cs));
@@ -568,26 +928,6 @@
fun add_undefined c thy =
(map_exec_purge (SOME [c]) o map_cases o apsnd) (Symtab.update (c, ())) thy;
-val map_pre = map_exec_purge NONE o map_thmproc o apfst o apfst;
-val map_post = map_exec_purge NONE o map_thmproc o apfst o apsnd;
-
-val add_inline = map_pre o MetaSimplifier.add_simp;
-val del_inline = map_pre o MetaSimplifier.del_simp;
-val add_post = map_post o MetaSimplifier.add_simp;
-val del_post = map_post o MetaSimplifier.del_simp;
-
-fun add_functrans (name, f) =
- (map_exec_purge NONE o map_thmproc o apsnd)
- (AList.update (op =) (name, (serial (), f)));
-
-fun del_functrans name =
- (map_exec_purge NONE o map_thmproc o apsnd)
- (delete_force "function transformer" name);
-
-fun simple_functrans f thy eqns = case f thy (map fst eqns)
- of SOME thms' => SOME (map (rpair (forall snd eqns)) thms')
- | NONE => NONE;
-
val _ = Context.>> (Context.map_theory
(let
fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
@@ -600,81 +940,16 @@
TypeInterpretation.init
#> add_del_attribute ("", (add_eqn, del_eqn))
#> add_simple_attribute ("nbe", add_nbe_eqn)
- #> add_del_attribute ("inline", (add_inline, del_inline))
- #> add_del_attribute ("post", (add_post, del_post))
end));
-
-(** post- and preprocessing **)
-
-local
-
-fun apply_functrans thy c _ [] = []
- | apply_functrans thy c [] eqns = eqns
- | apply_functrans thy c functrans eqns = eqns
- |> perhaps (perhaps_loop (perhaps_apply functrans))
- |> assert_eqns_const thy c;
-
-fun rhs_conv conv thm = Thm.transitive thm ((conv o Thm.rhs_of) thm);
-
-fun term_of_conv thy f =
- Thm.cterm_of thy
- #> f
- #> Thm.prop_of
- #> Logic.dest_equals
- #> snd;
-
-fun preprocess thy c eqns =
- let
- val pre = (Simplifier.theory_context thy o #pre o the_thmproc o the_exec) thy;
- val functrans = (map (fn (_, (_, f)) => f thy) o #functrans
- o the_thmproc o the_exec) thy;
- in
- eqns
- |> apply_functrans thy c functrans
- |> (map o apfst) (Code_Unit.rewrite_eqn pre)
- |> (map o apfst) (AxClass.unoverload thy)
- |> map (assert_eqn thy)
- |> burrow_fst (common_typ_eqns thy)
- end;
-
-in
-
-fun preprocess_conv thy ct =
- let
- val pre = (Simplifier.theory_context thy o #pre o the_thmproc o the_exec) thy;
- in
- ct
- |> Simplifier.rewrite pre
- |> rhs_conv (AxClass.unoverload_conv thy)
- end;
-
-fun preprocess_term thy = term_of_conv thy (preprocess_conv thy);
-
-fun postprocess_conv thy ct =
- let
- val post = (Simplifier.theory_context thy o #post o the_thmproc o the_exec) thy;
- in
- ct
- |> AxClass.overload_conv thy
- |> rhs_conv (Simplifier.rewrite post)
- end;
-
-fun postprocess_term thy = term_of_conv thy (postprocess_conv thy);
-
-fun these_raw_eqns thy c =
+fun these_eqns thy c =
get_eqns thy c
|> (map o apfst) (Thm.transfer thy)
|> burrow_fst (common_typ_eqns thy);
-fun these_eqns thy c =
- get_eqns thy c
- |> (map o apfst) (Thm.transfer thy)
- |> preprocess thy c;
-
fun default_typscheme thy c =
let
- fun the_const_typscheme c = (curry (Code_Unit.typscheme thy) c o snd o dest_Const
+ fun the_const_typscheme c = (curry (typscheme thy) c o snd o dest_Const
o TermSubst.zero_var_indexes o curry Const "" o Sign.the_const_type thy) c;
fun strip_sorts (vs, ty) = (map (fn (v, _) => (v, [])) vs, ty);
in case AxClass.class_of_param thy c
@@ -682,11 +957,9 @@
| NONE => if is_constr thy c
then strip_sorts (the_const_typscheme c)
else case get_eqns thy c
- of (thm, _) :: _ => (Code_Unit.typscheme_eqn thy o Drule.zero_var_indexes) thm
+ of (thm, _) :: _ => (typscheme_eqn thy o Drule.zero_var_indexes) thm
| [] => strip_sorts (the_const_typscheme c) end;
-end; (*local*)
-
end; (*struct*)
--- a/src/Pure/Isar/code_unit.ML Fri May 15 15:12:23 2009 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,447 +0,0 @@
-(* Title: Pure/Isar/code_unit.ML
- Author: Florian Haftmann, TU Muenchen
-
-Basic notions of code generation. Auxiliary.
-*)
-
-signature CODE_UNIT =
-sig
- (*typ instantiations*)
- val typscheme: theory -> string * typ -> (string * sort) list * typ
- val inst_thm: theory -> sort Vartab.table -> thm -> thm
- val constrain_thm: theory -> sort -> thm -> thm
-
- (*constant aliasses*)
- val add_const_alias: thm -> theory -> theory
- val triv_classes: theory -> class list
- val resubst_alias: theory -> string -> string
-
- (*constants*)
- val string_of_typ: theory -> typ -> string
- val string_of_const: theory -> string -> string
- val no_args: theory -> string -> int
- val check_const: theory -> term -> string
- val read_bare_const: theory -> string -> string * typ
- val read_const: theory -> string -> string
-
- (*constructor sets*)
- val constrset_of_consts: theory -> (string * typ) list
- -> string * ((string * sort) list * (string * typ list) list)
-
- (*code equations*)
- val mk_eqn: theory -> (string -> bool) -> thm * bool -> thm * bool
- val mk_eqn_liberal: theory -> (string -> bool) -> thm -> (thm * bool) option
- val assert_eqn: theory -> (string -> bool) -> thm * bool -> thm * bool
- val const_eqn: theory -> thm -> string
- val const_typ_eqn: thm -> string * typ
- val typscheme_eqn: theory -> thm -> (string * sort) list * typ
- val expand_eta: theory -> int -> thm -> thm
- val rewrite_eqn: simpset -> thm -> thm
- val rewrite_head: thm list -> thm -> thm
- val norm_args: theory -> thm list -> thm list
- val norm_varnames: theory -> thm list -> thm list
-
- (*case certificates*)
- val case_cert: thm -> string * (int * string list)
-end;
-
-structure Code_Unit: CODE_UNIT =
-struct
-
-
-(* auxiliary *)
-
-fun string_of_typ thy = setmp show_sorts true (Syntax.string_of_typ_global thy);
-fun string_of_const thy c = case AxClass.inst_of_param thy c
- of SOME (c, tyco) => Sign.extern_const thy c ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco)
- | NONE => Sign.extern_const thy c;
-
-fun no_args thy = length o fst o strip_type o Sign.the_const_type thy;
-
-
-(* utilities *)
-
-fun typscheme thy (c, ty) =
- let
- val ty' = Logic.unvarifyT ty;
- fun dest (TFree (v, sort)) = (v, sort)
- | dest ty = error ("Illegal type parameter in type scheme: " ^ Syntax.string_of_typ_global thy ty);
- val vs = map dest (Sign.const_typargs thy (c, ty'));
- in (vs, Type.strip_sorts ty') end;
-
-fun inst_thm thy tvars' thm =
- let
- val tvars = (Term.add_tvars o Thm.prop_of) thm [];
- val inter_sort = Sorts.inter_sort (Sign.classes_of thy);
- fun mk_inst (tvar as (v, sort)) = case Vartab.lookup tvars' v
- of SOME sort' => SOME (pairself (Thm.ctyp_of thy o TVar)
- (tvar, (v, inter_sort (sort, sort'))))
- | NONE => NONE;
- val insts = map_filter mk_inst tvars;
- in Thm.instantiate (insts, []) thm end;
-
-fun constrain_thm thy sort thm =
- let
- val constrain = curry (Sorts.inter_sort (Sign.classes_of thy)) sort
- val tvars = (Term.add_tvars o Thm.prop_of) thm [];
- fun mk_inst (tvar as (v, sort)) = pairself (Thm.ctyp_of thy o TVar o pair v)
- (sort, constrain sort)
- val insts = map mk_inst tvars;
- in Thm.instantiate (insts, []) thm end;
-
-fun expand_eta thy k thm =
- let
- val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm;
- val (head, args) = strip_comb lhs;
- val l = if k = ~1
- then (length o fst o strip_abs) rhs
- else Int.max (0, k - length args);
- val used = Name.make_context (map (fst o fst) (Term.add_vars lhs []));
- fun get_name _ 0 = pair []
- | get_name (Abs (v, ty, t)) k =
- Name.variants [v]
- ##>> get_name t (k - 1)
- #>> (fn ([v'], vs') => (v', ty) :: vs')
- | get_name t k =
- let
- val (tys, _) = (strip_type o fastype_of) t
- in case tys
- of [] => raise TERM ("expand_eta", [t])
- | ty :: _ =>
- Name.variants [""]
- #-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1)
- #>> (fn vs' => (v, ty) :: vs'))
- end;
- val (vs, _) = get_name rhs l used;
- fun expand (v, ty) thm = Drule.fun_cong_rule thm
- (Thm.cterm_of thy (Var ((v, 0), ty)));
- in
- thm
- |> fold expand vs
- |> Conv.fconv_rule Drule.beta_eta_conversion
- end;
-
-fun eqn_conv conv =
- let
- fun lhs_conv ct = if can Thm.dest_comb ct
- then (Conv.combination_conv lhs_conv conv) ct
- else Conv.all_conv ct;
- in Conv.combination_conv (Conv.arg_conv lhs_conv) conv end;
-
-fun head_conv conv =
- let
- fun lhs_conv ct = if can Thm.dest_comb ct
- then (Conv.fun_conv lhs_conv) ct
- else conv ct;
- in Conv.fun_conv (Conv.arg_conv lhs_conv) end;
-
-val rewrite_eqn = Conv.fconv_rule o eqn_conv o Simplifier.rewrite;
-val rewrite_head = Conv.fconv_rule o head_conv o MetaSimplifier.rewrite false;
-
-fun norm_args thy thms =
- let
- val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals;
- val k = fold (curry Int.max o num_args_of o Thm.prop_of) thms 0;
- in
- thms
- |> map (expand_eta thy k)
- |> map (Conv.fconv_rule Drule.beta_eta_conversion)
- end;
-
-fun canonical_tvars thy thm =
- let
- val ctyp = Thm.ctyp_of thy;
- val purify_tvar = unprefix "'" #> Name.desymbolize false #> prefix "'";
- fun tvars_subst_for thm = (fold_types o fold_atyps)
- (fn TVar (v_i as (v, _), sort) => let
- val v' = purify_tvar v
- in if v = v' then I
- else insert (op =) (v_i, (v', sort)) end
- | _ => I) (prop_of thm) [];
- fun mk_inst (v_i, (v', sort)) (maxidx, acc) =
- let
- val ty = TVar (v_i, sort)
- in
- (maxidx + 1, (ctyp ty, ctyp (TVar ((v', maxidx), sort))) :: acc)
- end;
- val maxidx = Thm.maxidx_of thm + 1;
- val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []);
- in Thm.instantiate (inst, []) thm end;
-
-fun canonical_vars thy thm =
- let
- val cterm = Thm.cterm_of thy;
- val purify_var = Name.desymbolize false;
- fun vars_subst_for thm = fold_aterms
- (fn Var (v_i as (v, _), ty) => let
- val v' = purify_var v
- in if v = v' then I
- else insert (op =) (v_i, (v', ty)) end
- | _ => I) (prop_of thm) [];
- fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) =
- let
- val t = Var (v_i, ty)
- in
- (maxidx + 1, (cterm t, cterm (Var ((v', maxidx), ty))) :: acc)
- end;
- val maxidx = Thm.maxidx_of thm + 1;
- val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []);
- in Thm.instantiate ([], inst) thm end;
-
-fun canonical_absvars thm =
- let
- val t = Thm.plain_prop_of thm;
- val purify_var = Name.desymbolize false;
- val t' = Term.map_abs_vars purify_var t;
- in Thm.rename_boundvars t t' thm end;
-
-fun norm_varnames thy thms =
- let
- fun burrow_thms f [] = []
- | burrow_thms f thms =
- thms
- |> Conjunction.intr_balanced
- |> f
- |> Conjunction.elim_balanced (length thms)
- in
- thms
- |> map (canonical_vars thy)
- |> map canonical_absvars
- |> map Drule.zero_var_indexes
- |> burrow_thms (canonical_tvars thy)
- |> Drule.zero_var_indexes_list
- end;
-
-
-(* const aliasses *)
-
-structure ConstAlias = TheoryDataFun
-(
- type T = ((string * string) * thm) list * class list;
- val empty = ([], []);
- val copy = I;
- val extend = I;
- fun merge _ ((alias1, classes1), (alias2, classes2)) : T =
- (Library.merge (eq_snd Thm.eq_thm_prop) (alias1, alias2),
- Library.merge (op =) (classes1, classes2));
-);
-
-fun add_const_alias thm thy =
- let
- val lhs_rhs = case try Logic.dest_equals (Thm.prop_of thm)
- of SOME lhs_rhs => lhs_rhs
- | _ => error ("Not an equation: " ^ Display.string_of_thm thm);
- val c_c' = case try (pairself (AxClass.unoverload_const thy o dest_Const)) lhs_rhs
- of SOME c_c' => c_c'
- | _ => error ("Not an equation with two constants: " ^ Display.string_of_thm thm);
- val some_class = the_list (AxClass.class_of_param thy (snd c_c'));
- in thy |>
- ConstAlias.map (fn (alias, classes) =>
- ((c_c', thm) :: alias, fold (insert (op =)) some_class classes))
- end;
-
-fun resubst_alias thy =
- let
- val alias = fst (ConstAlias.get thy);
- val subst_inst_param = Option.map fst o AxClass.inst_of_param thy;
- fun subst_alias c =
- get_first (fn ((c', c''), _) => if c = c'' then SOME c' else NONE) alias;
- in
- perhaps subst_inst_param
- #> perhaps subst_alias
- end;
-
-val triv_classes = snd o ConstAlias.get;
-
-
-(* reading constants as terms *)
-
-fun check_bare_const thy t = case try dest_Const t
- of SOME c_ty => c_ty
- | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
-
-fun check_const thy = AxClass.unoverload_const thy o check_bare_const thy;
-
-fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy;
-
-fun read_const thy = AxClass.unoverload_const thy o read_bare_const thy;
-
-
-(* constructor sets *)
-
-fun constrset_of_consts thy cs =
- let
- val _ = map (fn (c, _) => if (is_some o AxClass.class_of_param thy) c
- then error ("Is a class parameter: " ^ string_of_const thy c) else ()) cs;
- fun no_constr (c, ty) = error ("Not a datatype constructor: " ^ string_of_const thy c
- ^ " :: " ^ string_of_typ thy ty);
- fun last_typ c_ty ty =
- let
- val frees = OldTerm.typ_tfrees ty;
- val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) ty
- handle TYPE _ => no_constr c_ty
- val _ = if has_duplicates (eq_fst (op =)) vs then no_constr c_ty else ();
- val _ = if length frees <> length vs then no_constr c_ty else ();
- in (tyco, vs) end;
- fun ty_sorts (c, ty) =
- let
- val ty_decl = (Logic.unvarifyT o Sign.the_const_type thy) c;
- val (tyco, _) = last_typ (c, ty) ty_decl;
- val (_, vs) = last_typ (c, ty) ty;
- in ((tyco, map snd vs), (c, (map fst vs, ty))) end;
- fun add ((tyco', sorts'), c) ((tyco, sorts), cs) =
- let
- val _ = if tyco' <> tyco
- then error "Different type constructors in constructor set"
- else ();
- val sorts'' = map2 (curry (Sorts.inter_sort (Sign.classes_of thy))) sorts' sorts
- in ((tyco, sorts), c :: cs) end;
- fun inst vs' (c, (vs, ty)) =
- let
- val the_v = the o AList.lookup (op =) (vs ~~ vs');
- val ty' = map_atyps (fn TFree (v, _) => TFree (the_v v)) ty;
- in (c, (fst o strip_type) ty') end;
- val c' :: cs' = map ty_sorts cs;
- val ((tyco, sorts), cs'') = fold add cs' (apsnd single c');
- val vs = Name.names Name.context Name.aT sorts;
- val cs''' = map (inst vs) cs'';
- in (tyco, (vs, rev cs''')) end;
-
-
-(* code equations *)
-
-exception BAD_THM of string;
-fun bad_thm msg = raise BAD_THM msg;
-fun error_thm f thm = f thm handle BAD_THM msg => error msg;
-fun try_thm f thm = SOME (f thm) handle BAD_THM _ => NONE;
-
-fun is_linear thm =
- let val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm
- in not (has_duplicates (op =) ((fold o fold_aterms)
- (fn Var (v, _) => cons v | _ => I) args [])) end;
-
-fun gen_assert_eqn thy is_constr_head is_constr_pat (thm, proper) =
- let
- val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm
- handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm)
- | THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm);
- fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v
- | Free _ => bad_thm ("Illegal free variable in equation\n"
- ^ Display.string_of_thm thm)
- | _ => I) t [];
- fun tvars_of t = fold_term_types (fn _ =>
- fold_atyps (fn TVar (v, _) => insert (op =) v
- | TFree _ => bad_thm
- ("Illegal free type variable in equation\n" ^ Display.string_of_thm thm))) t [];
- val lhs_vs = vars_of lhs;
- val rhs_vs = vars_of rhs;
- val lhs_tvs = tvars_of lhs;
- val rhs_tvs = tvars_of rhs;
- val _ = if null (subtract (op =) lhs_vs rhs_vs)
- then ()
- else bad_thm ("Free variables on right hand side of equation\n"
- ^ Display.string_of_thm thm);
- val _ = if null (subtract (op =) lhs_tvs rhs_tvs)
- then ()
- else bad_thm ("Free type variables on right hand side of equation\n"
- ^ Display.string_of_thm thm) val (head, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm;
- val (c, ty) = case head
- of Const (c_ty as (_, ty)) => (AxClass.unoverload_const thy c_ty, ty)
- | _ => bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm thm);
- fun check _ (Abs _) = bad_thm
- ("Abstraction on left hand side of equation\n"
- ^ Display.string_of_thm thm)
- | check 0 (Var _) = ()
- | check _ (Var _) = bad_thm
- ("Variable with application on left hand side of equation\n"
- ^ Display.string_of_thm thm)
- | check n (t1 $ t2) = (check (n+1) t1; check 0 t2)
- | check n (Const (c_ty as (c, ty))) = if n = (length o fst o strip_type) ty
- then if not proper orelse is_constr_pat (AxClass.unoverload_const thy c_ty)
- then ()
- else bad_thm (quote c ^ " is not a constructor, on left hand side of equation\n"
- ^ Display.string_of_thm thm)
- else bad_thm
- ("Partially applied constant " ^ quote c ^ " on left hand side of equation\n"
- ^ Display.string_of_thm thm);
- val _ = map (check 0) args;
- val _ = if not proper orelse is_linear thm then ()
- else bad_thm ("Duplicate variables on left hand side of equation\n"
- ^ Display.string_of_thm thm);
- val _ = if (is_none o AxClass.class_of_param thy) c
- then ()
- else bad_thm ("Polymorphic constant as head in equation\n"
- ^ Display.string_of_thm thm)
- val _ = if not (is_constr_head c)
- then ()
- else bad_thm ("Constructor as head in equation\n"
- ^ Display.string_of_thm thm)
- val ty_decl = Sign.the_const_type thy c;
- val _ = if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty)
- then () else bad_thm ("Type\n" ^ string_of_typ thy ty
- ^ "\nof equation\n"
- ^ Display.string_of_thm thm
- ^ "\nis incompatible with declared function type\n"
- ^ string_of_typ thy ty_decl)
- in (thm, proper) end;
-
-fun assert_eqn thy is_constr = error_thm (gen_assert_eqn thy is_constr is_constr);
-
-val const_typ_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
-
-fun typscheme_eqn thy = typscheme thy o const_typ_eqn;
-
-(*these are permissive wrt. to overloaded constants!*)
-fun mk_eqn thy is_constr_head = error_thm (gen_assert_eqn thy is_constr_head (K true)) o
- apfst (LocalDefs.meta_rewrite_rule (ProofContext.init thy));
-
-fun mk_eqn_liberal thy is_constr_head = Option.map (fn (thm, _) => (thm, is_linear thm))
- o try_thm (gen_assert_eqn thy is_constr_head (K true))
- o rpair false o LocalDefs.meta_rewrite_rule (ProofContext.init thy);
-
-fun const_eqn thy = AxClass.unoverload_const thy o const_typ_eqn;
-
-
-(* case cerificates *)
-
-fun case_certificate thm =
- let
- val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals
- o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.prop_of) thm;
- val _ = case head of Free _ => true
- | Var _ => true
- | _ => raise TERM ("case_cert", []);
- val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr;
- val (Const (case_const, _), raw_params) = strip_comb case_expr;
- val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params;
- val _ = if n = ~1 then raise TERM ("case_cert", []) else ();
- val params = map (fst o dest_Var) (nth_drop n raw_params);
- fun dest_case t =
- let
- val (head' $ t_co, rhs) = Logic.dest_equals t;
- val _ = if head' = head then () else raise TERM ("case_cert", []);
- val (Const (co, _), args) = strip_comb t_co;
- val (Var (param, _), args') = strip_comb rhs;
- val _ = if args' = args then () else raise TERM ("case_cert", []);
- in (param, co) end;
- fun analyze_cases cases =
- let
- val co_list = fold (AList.update (op =) o dest_case) cases [];
- in map (the o AList.lookup (op =) co_list) params end;
- fun analyze_let t =
- let
- val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t;
- val _ = if head' = head then () else raise TERM ("case_cert", []);
- val _ = if arg' = arg then () else raise TERM ("case_cert", []);
- val _ = if [param'] = params then () else raise TERM ("case_cert", []);
- in [] end;
- fun analyze (cases as [let_case]) =
- (analyze_cases cases handle Bind => analyze_let let_case)
- | analyze cases = analyze_cases cases;
- in (case_const, (n, analyze cases)) end;
-
-fun case_cert thm = case_certificate thm
- handle Bind => error "bad case certificate"
- | TERM _ => error "bad case certificate";
-
-end;
--- a/src/Pure/Isar/isar_syn.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/Pure/Isar/isar_syn.ML Fri May 15 15:13:09 2009 -0700
@@ -881,7 +881,7 @@
(opt_modes -- P.typ >> (Toplevel.no_timing oo IsarCmd.print_type));
val _ =
- OuterSyntax.improper_command "print_codesetup" "print code generator setup of this theory" K.diag
+ OuterSyntax.improper_command "print_codesetup" "print code generator setup" K.diag
(Scan.succeed
(Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep
(Code.print_codesetup o Toplevel.theory_of)));
--- a/src/Pure/codegen.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/Pure/codegen.ML Fri May 15 15:13:09 2009 -0700
@@ -75,7 +75,7 @@
val mk_type: bool -> typ -> Pretty.T
val mk_term_of: codegr -> string -> bool -> typ -> Pretty.T
val mk_gen: codegr -> string -> bool -> string list -> string -> typ -> Pretty.T
- val test_fn: (int -> (string * term) list option) ref
+ val test_fn: (int -> term list option) ref
val test_term: Proof.context -> term -> int -> term list option
val eval_result: (unit -> term) ref
val eval_term: theory -> term -> term
@@ -329,7 +329,7 @@
end;
val assoc_const_i = gen_assoc_const (K I);
-val assoc_const = gen_assoc_const Code_Unit.read_bare_const;
+val assoc_const = gen_assoc_const Code.read_bare_const;
(**** associate types with target language types ****)
@@ -871,39 +871,35 @@
[mk_gen gr module true xs a T, mk_type true T]) Ts) @
(if member (op =) xs s then [str a] else []))));
-val test_fn : (int -> (string * term) list option) ref = ref (fn _ => NONE);
+val test_fn : (int -> term list option) ref = ref (fn _ => NONE);
fun test_term ctxt t =
let
val thy = ProofContext.theory_of ctxt;
val (code, gr) = setmp mode ["term_of", "test"]
(generate_code_i thy [] "Generated") [("testf", t)];
- val frees = Name.names Name.context "a" ((map snd o fst o strip_abs) t);
- val frees' = frees ~~
- map (fn i => "arg" ^ string_of_int i) (1 upto length frees);
+ val Ts = map snd (fst (strip_abs t));
+ val args = map_index (fn (i, T) => ("arg" ^ string_of_int i, T)) Ts;
val s = "structure TestTerm =\nstruct\n\n" ^
cat_lines (map snd code) ^
"\nopen Generated;\n\n" ^ string_of
(Pretty.block [str "val () = Codegen.test_fn :=",
Pretty.brk 1, str ("(fn i =>"), Pretty.brk 1,
- mk_let (map (fn ((s, T), s') =>
- (mk_tuple [str s', str (s' ^ "_t")],
+ mk_let (map (fn (s, T) =>
+ (mk_tuple [str s, str (s ^ "_t")],
Pretty.block [mk_gen gr "Generated" false [] "" T, Pretty.brk 1,
- str "(i + 1)"])) frees')
+ str "i"])) args)
(Pretty.block [str "if ",
- mk_app false (str "testf") (map (str o snd) frees'),
+ mk_app false (str "testf") (map (str o fst) args),
Pretty.brk 1, str "then NONE",
Pretty.brk 1, str "else ",
Pretty.block [str "SOME ", Pretty.block (str "[" ::
- flat (separate [str ",", Pretty.brk 1]
- (map (fn ((s, T), s') => [Pretty.block
- [str ("(" ^ quote (Symbol.escape s) ^ ","), Pretty.brk 1,
- str (s' ^ "_t ())")]]) frees')) @
+ Pretty.commas (map (fn (s, _) => str (s ^ "_t ()")) args) @
[str "]"])]]),
str ");"]) ^
"\n\nend;\n";
val _ = ML_Context.eval_in (SOME ctxt) false Position.none s;
- in ! test_fn #> (Option.map o map) snd end;
+ in ! test_fn end;
@@ -1024,8 +1020,6 @@
val setup = add_codegen "default" default_codegen
#> add_tycodegen "default" default_tycodegen
- #> Code.add_attribute ("unfold", Scan.succeed (Thm.declaration_attribute
- (fn thm => Context.mapping (add_unfold thm #> Code.add_inline thm) I)))
#> add_preprocessor unfold_preprocessor;
val _ =
--- a/src/Tools/Code_Generator.thy Fri May 15 15:12:23 2009 -0700
+++ b/src/Tools/Code_Generator.thy Fri May 15 15:13:09 2009 -0700
@@ -9,7 +9,7 @@
uses
"~~/src/Tools/value.ML"
"~~/src/Tools/quickcheck.ML"
- "~~/src/Tools/code/code_wellsorted.ML"
+ "~~/src/Tools/code/code_preproc.ML"
"~~/src/Tools/code/code_thingol.ML"
"~~/src/Tools/code/code_printer.ML"
"~~/src/Tools/code/code_target.ML"
@@ -19,7 +19,8 @@
begin
setup {*
- Code_ML.setup
+ Code_Preproc.setup
+ #> Code_ML.setup
#> Code_Haskell.setup
#> Nbe.setup
*}
--- a/src/Tools/code/code_haskell.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/Tools/code/code_haskell.ML Fri May 15 15:13:09 2009 -0700
@@ -480,7 +480,7 @@
fun add_monad target' raw_c_bind thy =
let
- val c_bind = Code_Unit.read_const thy raw_c_bind;
+ val c_bind = Code.read_const thy raw_c_bind;
in if target = target' then
thy
|> Code_Target.add_syntax_const target c_bind
--- a/src/Tools/code/code_ml.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/Tools/code/code_ml.ML Fri May 15 15:13:09 2009 -0700
@@ -161,20 +161,21 @@
:: map (pr "|") clauses
)
end
- | pr_case is_closure thm vars fxy ((_, []), _) = str "raise Fail \"empty case\"";
+ | pr_case is_closure thm vars fxy ((_, []), _) =
+ (concat o map str) ["raise", "Fail", "\"empty case\""];
fun pr_stmt (MLExc (name, n)) =
let
val exc_str =
(ML_Syntax.print_string o Long_Name.base_name o Long_Name.qualifier) name;
in
- concat (
- str (if n = 0 then "val" else "fun")
- :: (str o deresolve) name
- :: map str (replicate n "_")
- @ str "="
- :: str "raise"
- :: str "(Fail"
- @@ str (exc_str ^ ")")
+ (concat o map str) (
+ (if n = 0 then "val" else "fun")
+ :: deresolve name
+ :: replicate n "_"
+ @ "="
+ :: "raise"
+ :: "Fail"
+ @@ exc_str
)
end
| pr_stmt (MLVal (name, (((vs, ty), t), (thm, _)))) =
@@ -458,7 +459,8 @@
:: map (pr "|") clauses
)
end
- | pr_case is_closure thm vars fxy ((_, []), _) = str "failwith \"empty case\"";
+ | pr_case is_closure thm vars fxy ((_, []), _) =
+ (concat o map str) ["failwith", "\"empty case\""];
fun fish_params vars eqs =
let
fun fish_param _ (w as SOME _) = w
@@ -477,13 +479,13 @@
val exc_str =
(ML_Syntax.print_string o Long_Name.base_name o Long_Name.qualifier) name;
in
- concat (
- str "let"
- :: (str o deresolve) name
- :: map str (replicate n "_")
- @ str "="
- :: str "failwith"
- @@ str exc_str
+ (concat o map str) (
+ "let"
+ :: deresolve name
+ :: replicate n "_"
+ @ "="
+ :: "failwith"
+ @@ exc_str
)
end
| pr_stmt (MLVal (name, (((vs, ty), t), (thm, _)))) =
@@ -994,7 +996,7 @@
val (ml_code, target_names) = eval_code_of NONE thy naming program (consts' @ tycos');
val (consts'', tycos'') = chop (length consts') target_names;
val consts_map = map2 (fn const => fn NONE =>
- error ("Constant " ^ (quote o Code_Unit.string_of_const thy) const
+ error ("Constant " ^ (quote o Code.string_of_const thy) const
^ "\nhas a user-defined serialization")
| SOME const'' => (const, const'')) consts consts''
val tycos_map = map2 (fn tyco => fn NONE =>
@@ -1048,7 +1050,7 @@
fun ml_code_antiq raw_const {struct_name, background} =
let
- val const = Code_Unit.check_const (ProofContext.theory_of background) raw_const;
+ val const = Code.check_const (ProofContext.theory_of background) raw_const;
val is_first = is_first_occ background;
val background' = register_const const background;
in (print_code struct_name is_first (print_const const), background') end;
@@ -1057,7 +1059,7 @@
let
val thy = ProofContext.theory_of background;
val tyco = Sign.intern_type thy raw_tyco;
- val constrs = map (Code_Unit.check_const thy) raw_constrs;
+ val constrs = map (Code.check_const thy) raw_constrs;
val constrs' = (map fst o snd o Code.get_datatype thy) tyco;
val _ = if gen_eq_set (op =) (constrs, constrs') then ()
else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors")
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/code/code_preproc.ML Fri May 15 15:13:09 2009 -0700
@@ -0,0 +1,515 @@
+(* Title: Tools/code/code_preproc.ML
+ Author: Florian Haftmann, TU Muenchen
+
+Preprocessing code equations into a well-sorted system
+in a graph with explicit dependencies.
+*)
+
+signature CODE_PREPROC =
+sig
+ val map_pre: (simpset -> simpset) -> theory -> theory
+ val map_post: (simpset -> simpset) -> theory -> theory
+ val add_inline: thm -> theory -> theory
+ val add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory
+ val del_functrans: string -> theory -> theory
+ val simple_functrans: (theory -> thm list -> thm list option)
+ -> theory -> (thm * bool) list -> (thm * bool) list option
+ val print_codeproc: theory -> unit
+
+ type code_algebra
+ type code_graph
+ val eqns: code_graph -> string -> (thm * bool) list
+ val typ: code_graph -> string -> (string * sort) list * typ
+ val all: code_graph -> string list
+ val pretty: theory -> code_graph -> Pretty.T
+ val obtain: theory -> string list -> term list -> code_algebra * code_graph
+ val eval_conv: theory -> (sort -> sort)
+ -> (code_algebra -> code_graph -> (string * sort) list -> term -> cterm -> thm) -> cterm -> thm
+ val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a)
+ -> (code_algebra -> code_graph -> (string * sort) list -> term -> 'a) -> term -> 'a
+
+ val setup: theory -> theory
+end
+
+structure Code_Preproc : CODE_PREPROC =
+struct
+
+(** preprocessor administration **)
+
+(* theory data *)
+
+datatype thmproc = Thmproc of {
+ pre: simpset,
+ post: simpset,
+ functrans: (string * (serial * (theory -> (thm * bool) list -> (thm * bool) list option))) list
+};
+
+fun mk_thmproc ((pre, post), functrans) =
+ Thmproc { pre = pre, post = post, functrans = functrans };
+fun map_thmproc f (Thmproc { pre, post, functrans }) =
+ mk_thmproc (f ((pre, post), functrans));
+fun merge_thmproc (Thmproc { pre = pre1, post = post1, functrans = functrans1 },
+ Thmproc { pre = pre2, post = post2, functrans = functrans2 }) =
+ let
+ val pre = Simplifier.merge_ss (pre1, pre2);
+ val post = Simplifier.merge_ss (post1, post2);
+ val functrans = AList.merge (op =) (eq_fst (op =)) (functrans1, functrans2);
+ in mk_thmproc ((pre, post), functrans) end;
+
+structure Code_Preproc_Data = TheoryDataFun
+(
+ type T = thmproc;
+ val empty = mk_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []);
+ fun copy spec = spec;
+ val extend = copy;
+ fun merge pp = merge_thmproc;
+);
+
+fun the_thmproc thy = case Code_Preproc_Data.get thy
+ of Thmproc x => x;
+
+fun delete_force msg key xs =
+ if AList.defined (op =) xs key then AList.delete (op =) key xs
+ else error ("No such " ^ msg ^ ": " ^ quote key);
+
+fun map_data f thy =
+ thy
+ |> Code.purge_data
+ |> (Code_Preproc_Data.map o map_thmproc) f;
+
+val map_pre = map_data o apfst o apfst;
+val map_post = map_data o apfst o apsnd;
+
+val add_inline = map_pre o MetaSimplifier.add_simp;
+val del_inline = map_pre o MetaSimplifier.del_simp;
+val add_post = map_post o MetaSimplifier.add_simp;
+val del_post = map_post o MetaSimplifier.del_simp;
+
+fun add_functrans (name, f) = (map_data o apsnd)
+ (AList.update (op =) (name, (serial (), f)));
+
+fun del_functrans name = (map_data o apsnd)
+ (delete_force "function transformer" name);
+
+
+(* post- and preprocessing *)
+
+fun apply_functrans thy c _ [] = []
+ | apply_functrans thy c [] eqns = eqns
+ | apply_functrans thy c functrans eqns = eqns
+ |> perhaps (perhaps_loop (perhaps_apply functrans))
+ |> Code.assert_eqns_const thy c;
+
+fun rhs_conv conv thm = Thm.transitive thm ((conv o Thm.rhs_of) thm);
+
+fun term_of_conv thy f =
+ Thm.cterm_of thy
+ #> f
+ #> Thm.prop_of
+ #> Logic.dest_equals
+ #> snd;
+
+fun preprocess thy c eqns =
+ let
+ val pre = (Simplifier.theory_context thy o #pre o the_thmproc) thy;
+ val functrans = (map (fn (_, (_, f)) => f thy) o #functrans
+ o the_thmproc) thy;
+ in
+ eqns
+ |> apply_functrans thy c functrans
+ |> (map o apfst) (Code.rewrite_eqn pre)
+ |> (map o apfst) (AxClass.unoverload thy)
+ |> map (Code.assert_eqn thy)
+ |> burrow_fst (Code.norm_args thy)
+ |> burrow_fst (Code.norm_varnames thy)
+ end;
+
+fun preprocess_conv thy ct =
+ let
+ val pre = (Simplifier.theory_context thy o #pre o the_thmproc) thy;
+ in
+ ct
+ |> Simplifier.rewrite pre
+ |> rhs_conv (AxClass.unoverload_conv thy)
+ end;
+
+fun postprocess_conv thy ct =
+ let
+ val post = (Simplifier.theory_context thy o #post o the_thmproc) thy;
+ in
+ ct
+ |> AxClass.overload_conv thy
+ |> rhs_conv (Simplifier.rewrite post)
+ end;
+
+fun postprocess_term thy = term_of_conv thy (postprocess_conv thy);
+
+fun print_codeproc thy =
+ let
+ val ctxt = ProofContext.init thy;
+ val pre = (#pre o the_thmproc) thy;
+ val post = (#post o the_thmproc) thy;
+ val functrans = (map fst o #functrans o the_thmproc) thy;
+ in
+ (Pretty.writeln o Pretty.chunks) [
+ Pretty.block [
+ Pretty.str "preprocessing simpset:",
+ Pretty.fbrk,
+ Simplifier.pretty_ss ctxt pre
+ ],
+ Pretty.block [
+ Pretty.str "postprocessing simpset:",
+ Pretty.fbrk,
+ Simplifier.pretty_ss ctxt post
+ ],
+ Pretty.block (
+ Pretty.str "function transformers:"
+ :: Pretty.fbrk
+ :: (Pretty.fbreaks o map Pretty.str) functrans
+ )
+ ]
+ end;
+
+fun simple_functrans f thy eqns = case f thy (map fst eqns)
+ of SOME thms' => SOME (map (rpair (forall snd eqns)) thms')
+ | NONE => NONE;
+
+
+(** sort algebra and code equation graph types **)
+
+type code_algebra = (sort -> sort) * Sorts.algebra;
+type code_graph = (((string * sort) list * typ) * (thm * bool) list) Graph.T;
+
+fun eqns eqngr = these o Option.map snd o try (Graph.get_node eqngr);
+fun typ eqngr = fst o Graph.get_node eqngr;
+fun all eqngr = Graph.keys eqngr;
+
+fun pretty thy eqngr =
+ AList.make (snd o Graph.get_node eqngr) (Graph.keys eqngr)
+ |> (map o apfst) (Code.string_of_const thy)
+ |> sort (string_ord o pairself fst)
+ |> map (fn (s, thms) =>
+ (Pretty.block o Pretty.fbreaks) (
+ Pretty.str s
+ :: map (Display.pretty_thm o fst) thms
+ ))
+ |> Pretty.chunks;
+
+
+(** the Waisenhaus algorithm **)
+
+(* auxiliary *)
+
+fun is_proper_class thy = can (AxClass.get_info thy);
+
+fun complete_proper_sort thy =
+ Sign.complete_sort thy #> filter (is_proper_class thy);
+
+fun inst_params thy tyco =
+ map (fn (c, _) => AxClass.param_of_inst thy (c, tyco))
+ o maps (#params o AxClass.get_info thy);
+
+fun consts_of thy eqns = [] |> (fold o fold o fold_aterms)
+ (fn Const (c, ty) => insert (op =) (c, Sign.const_typargs thy (c, Logic.unvarifyT ty)) | _ => I)
+ (map (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of o fst) eqns);
+
+fun tyscm_rhss_of thy c eqns =
+ let
+ val tyscm = case eqns of [] => Code.default_typscheme thy c
+ | ((thm, _) :: _) => Code.typscheme_eqn thy thm;
+ val rhss = consts_of thy eqns;
+ in (tyscm, rhss) end;
+
+
+(* data structures *)
+
+datatype const = Fun of string | Inst of class * string;
+
+fun const_ord (Fun c1, Fun c2) = fast_string_ord (c1, c2)
+ | const_ord (Inst class_tyco1, Inst class_tyco2) =
+ prod_ord fast_string_ord fast_string_ord (class_tyco1, class_tyco2)
+ | const_ord (Fun _, Inst _) = LESS
+ | const_ord (Inst _, Fun _) = GREATER;
+
+type var = const * int;
+
+structure Vargraph =
+ GraphFun(type key = var val ord = prod_ord const_ord int_ord);
+
+datatype styp = Tyco of string * styp list | Var of var | Free;
+
+fun styp_of c_lhs (Type (tyco, tys)) = Tyco (tyco, map (styp_of c_lhs) tys)
+ | styp_of c_lhs (TFree (v, _)) = case c_lhs
+ of SOME (c, lhs) => Var (Fun c, find_index (fn (v', _) => v = v') lhs)
+ | NONE => Free;
+
+type vardeps_data = ((string * styp list) list * class list) Vargraph.T
+ * (((string * sort) list * (thm * bool) list) Symtab.table
+ * (class * string) list);
+
+val empty_vardeps_data : vardeps_data =
+ (Vargraph.empty, (Symtab.empty, []));
+
+
+(* retrieving equations and instances from the background context *)
+
+fun obtain_eqns thy eqngr c =
+ case try (Graph.get_node eqngr) c
+ of SOME ((lhs, _), eqns) => ((lhs, []), [])
+ | NONE => let
+ val eqns = Code.these_eqns thy c
+ |> preprocess thy c;
+ val ((lhs, _), rhss) = tyscm_rhss_of thy c eqns;
+ in ((lhs, rhss), eqns) end;
+
+fun obtain_instance thy arities (inst as (class, tyco)) =
+ case AList.lookup (op =) arities inst
+ of SOME classess => (classess, ([], []))
+ | NONE => let
+ val all_classes = complete_proper_sort thy [class];
+ val superclasses = remove (op =) class all_classes
+ val classess = map (complete_proper_sort thy)
+ (Sign.arity_sorts thy tyco [class]);
+ val inst_params = inst_params thy tyco all_classes;
+ in (classess, (superclasses, inst_params)) end;
+
+
+(* computing instantiations *)
+
+fun add_classes thy arities eqngr c_k new_classes vardeps_data =
+ let
+ val (styps, old_classes) = Vargraph.get_node (fst vardeps_data) c_k;
+ val diff_classes = new_classes |> subtract (op =) old_classes;
+ in if null diff_classes then vardeps_data
+ else let
+ val c_ks = Vargraph.imm_succs (fst vardeps_data) c_k |> insert (op =) c_k;
+ in
+ vardeps_data
+ |> (apfst o Vargraph.map_node c_k o apsnd) (append diff_classes)
+ |> fold (fn styp => fold (ensure_typmatch_inst thy arities eqngr styp) new_classes) styps
+ |> fold (fn c_k => add_classes thy arities eqngr c_k diff_classes) c_ks
+ end end
+and add_styp thy arities eqngr c_k tyco_styps vardeps_data =
+ let
+ val (old_styps, classes) = Vargraph.get_node (fst vardeps_data) c_k;
+ in if member (op =) old_styps tyco_styps then vardeps_data
+ else
+ vardeps_data
+ |> (apfst o Vargraph.map_node c_k o apfst) (cons tyco_styps)
+ |> fold (ensure_typmatch_inst thy arities eqngr tyco_styps) classes
+ end
+and add_dep thy arities eqngr c_k c_k' vardeps_data =
+ let
+ val (_, classes) = Vargraph.get_node (fst vardeps_data) c_k;
+ in
+ vardeps_data
+ |> add_classes thy arities eqngr c_k' classes
+ |> apfst (Vargraph.add_edge (c_k, c_k'))
+ end
+and ensure_typmatch_inst thy arities eqngr (tyco, styps) class vardeps_data =
+ if can (Sign.arity_sorts thy tyco) [class]
+ then vardeps_data
+ |> ensure_inst thy arities eqngr (class, tyco)
+ |> fold_index (fn (k, styp) =>
+ ensure_typmatch thy arities eqngr styp (Inst (class, tyco), k)) styps
+ else vardeps_data (*permissive!*)
+and ensure_inst thy arities eqngr (inst as (class, tyco)) (vardeps_data as (_, (_, insts))) =
+ if member (op =) insts inst then vardeps_data
+ else let
+ val (classess, (superclasses, inst_params)) =
+ obtain_instance thy arities inst;
+ in
+ vardeps_data
+ |> (apsnd o apsnd) (insert (op =) inst)
+ |> fold_index (fn (k, _) =>
+ apfst (Vargraph.new_node ((Inst (class, tyco), k), ([] ,[])))) classess
+ |> fold (fn superclass => ensure_inst thy arities eqngr (superclass, tyco)) superclasses
+ |> fold (ensure_fun thy arities eqngr) inst_params
+ |> fold_index (fn (k, classes) =>
+ add_classes thy arities eqngr (Inst (class, tyco), k) classes
+ #> fold (fn superclass =>
+ add_dep thy arities eqngr (Inst (superclass, tyco), k)
+ (Inst (class, tyco), k)) superclasses
+ #> fold (fn inst_param =>
+ add_dep thy arities eqngr (Fun inst_param, k)
+ (Inst (class, tyco), k)
+ ) inst_params
+ ) classess
+ end
+and ensure_typmatch thy arities eqngr (Tyco tyco_styps) c_k vardeps_data =
+ vardeps_data
+ |> add_styp thy arities eqngr c_k tyco_styps
+ | ensure_typmatch thy arities eqngr (Var c_k') c_k vardeps_data =
+ vardeps_data
+ |> add_dep thy arities eqngr c_k c_k'
+ | ensure_typmatch thy arities eqngr Free c_k vardeps_data =
+ vardeps_data
+and ensure_rhs thy arities eqngr (c', styps) vardeps_data =
+ vardeps_data
+ |> ensure_fun thy arities eqngr c'
+ |> fold_index (fn (k, styp) =>
+ ensure_typmatch thy arities eqngr styp (Fun c', k)) styps
+and ensure_fun thy arities eqngr c (vardeps_data as (_, (eqntab, _))) =
+ if Symtab.defined eqntab c then vardeps_data
+ else let
+ val ((lhs, rhss), eqns) = obtain_eqns thy eqngr c;
+ val rhss' = (map o apsnd o map) (styp_of (SOME (c, lhs))) rhss;
+ in
+ vardeps_data
+ |> (apsnd o apfst) (Symtab.update_new (c, (lhs, eqns)))
+ |> fold_index (fn (k, _) =>
+ apfst (Vargraph.new_node ((Fun c, k), ([] ,[])))) lhs
+ |> fold_index (fn (k, (_, sort)) =>
+ add_classes thy arities eqngr (Fun c, k) (complete_proper_sort thy sort)) lhs
+ |> fold (ensure_rhs thy arities eqngr) rhss'
+ end;
+
+
+(* applying instantiations *)
+
+fun dicts_of thy (proj_sort, algebra) (T, sort) =
+ let
+ fun class_relation (x, _) _ = x;
+ fun type_constructor tyco xs class =
+ inst_params thy tyco (Sorts.complete_sort algebra [class])
+ @ (maps o maps) fst xs;
+ fun type_variable (TFree (_, sort)) = map (pair []) (proj_sort sort);
+ in
+ flat (Sorts.of_sort_derivation (Syntax.pp_global thy) algebra
+ { class_relation = class_relation, type_constructor = type_constructor,
+ type_variable = type_variable } (T, proj_sort sort)
+ handle Sorts.CLASS_ERROR _ => [] (*permissive!*))
+ end;
+
+fun add_arity thy vardeps (class, tyco) =
+ AList.default (op =)
+ ((class, tyco), map (fn k => (snd o Vargraph.get_node vardeps) (Inst (class, tyco), k))
+ (0 upto Sign.arity_number thy tyco - 1));
+
+fun add_eqs thy vardeps (c, (proto_lhs, proto_eqns)) (rhss, eqngr) =
+ if can (Graph.get_node eqngr) c then (rhss, eqngr)
+ else let
+ val lhs = map_index (fn (k, (v, _)) =>
+ (v, snd (Vargraph.get_node vardeps (Fun c, k)))) proto_lhs;
+ val inst_tab = Vartab.empty |> fold (fn (v, sort) =>
+ Vartab.update ((v, 0), sort)) lhs;
+ val eqns = proto_eqns
+ |> (map o apfst) (Code.inst_thm thy inst_tab);
+ val (tyscm, rhss') = tyscm_rhss_of thy c eqns;
+ val eqngr' = Graph.new_node (c, (tyscm, eqns)) eqngr;
+ in (map (pair c) rhss' @ rhss, eqngr') end;
+
+fun extend_arities_eqngr thy cs ts (arities, eqngr) =
+ let
+ val cs_rhss = (fold o fold_aterms) (fn Const (c_ty as (c, _)) =>
+ insert (op =) (c, (map (styp_of NONE) o Sign.const_typargs thy) c_ty) | _ => I) ts [];
+ val (vardeps, (eqntab, insts)) = empty_vardeps_data
+ |> fold (ensure_fun thy arities eqngr) cs
+ |> fold (ensure_rhs thy arities eqngr) cs_rhss;
+ val arities' = fold (add_arity thy vardeps) insts arities;
+ val pp = Syntax.pp_global thy;
+ val algebra = Sorts.subalgebra pp (is_proper_class thy)
+ (AList.lookup (op =) arities') (Sign.classes_of thy);
+ val (rhss, eqngr') = Symtab.fold (add_eqs thy vardeps) eqntab ([], eqngr);
+ fun deps_of (c, rhs) = c :: maps (dicts_of thy algebra)
+ (rhs ~~ (map snd o fst o fst o Graph.get_node eqngr') c);
+ val eqngr'' = fold (fn (c, rhs) => fold
+ (curry Graph.add_edge c) (deps_of rhs)) rhss eqngr';
+ in (algebra, (arities', eqngr'')) end;
+
+
+(** store for preprocessed arities and code equations **)
+
+structure Wellsorted = CodeDataFun
+(
+ type T = ((string * class) * sort list) list * code_graph;
+ val empty = ([], Graph.empty);
+ fun purge thy cs (arities, eqngr) =
+ let
+ val del_cs = ((Graph.all_preds eqngr
+ o filter (can (Graph.get_node eqngr))) cs);
+ val del_arities = del_cs
+ |> map_filter (AxClass.inst_of_param thy)
+ |> maps (fn (c, tyco) =>
+ (map (rpair tyco) o Sign.complete_sort thy o the_list
+ o AxClass.class_of_param thy) c);
+ val arities' = fold (AList.delete (op =)) del_arities arities;
+ val eqngr' = Graph.del_nodes del_cs eqngr;
+ in (arities', eqngr') end;
+);
+
+
+(** retrieval and evaluation interfaces **)
+
+fun obtain thy cs ts = apsnd snd
+ (Wellsorted.change_yield thy (extend_arities_eqngr thy cs ts));
+
+fun prepare_sorts_typ prep_sort
+ = map_type_tfree (fn (v, sort) => TFree (v, prep_sort sort));
+
+fun prepare_sorts prep_sort (Const (c, ty)) =
+ Const (c, prepare_sorts_typ prep_sort ty)
+ | prepare_sorts prep_sort (t1 $ t2) =
+ prepare_sorts prep_sort t1 $ prepare_sorts prep_sort t2
+ | prepare_sorts prep_sort (Abs (v, ty, t)) =
+ Abs (v, prepare_sorts_typ prep_sort ty, prepare_sorts prep_sort t)
+ | prepare_sorts _ (t as Bound _) = t;
+
+fun gen_eval thy cterm_of conclude_evaluation prep_sort evaluator proto_ct =
+ let
+ val pp = Syntax.pp_global thy;
+ val ct = cterm_of proto_ct;
+ val _ = (Sign.no_frees pp o map_types (K dummyT) o Sign.no_vars pp)
+ (Thm.term_of ct);
+ val thm = preprocess_conv thy ct;
+ val ct' = Thm.rhs_of thm;
+ val t' = Thm.term_of ct';
+ val vs = Term.add_tfrees t' [];
+ val consts = fold_aterms
+ (fn Const (c, _) => insert (op =) c | _ => I) t' [];
+
+ val t'' = prepare_sorts prep_sort t';
+ val (algebra', eqngr') = obtain thy consts [t''];
+ in conclude_evaluation (evaluator algebra' eqngr' vs t'' ct') thm end;
+
+fun simple_evaluator evaluator algebra eqngr vs t ct =
+ evaluator algebra eqngr vs t;
+
+fun eval_conv thy =
+ let
+ fun conclude_evaluation thm2 thm1 =
+ let
+ val thm3 = postprocess_conv thy (Thm.rhs_of thm2);
+ in
+ Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ =>
+ error ("could not construct evaluation proof:\n"
+ ^ (cat_lines o map Display.string_of_thm) [thm1, thm2, thm3])
+ end;
+ in gen_eval thy I conclude_evaluation end;
+
+fun eval thy prep_sort postproc evaluator = gen_eval thy (Thm.cterm_of thy)
+ (K o postproc (postprocess_term thy)) prep_sort (simple_evaluator evaluator);
+
+
+(** setup **)
+
+val setup =
+ let
+ fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
+ fun add_del_attribute (name, (add, del)) =
+ Code.add_attribute (name, Args.del |-- Scan.succeed (mk_attribute del)
+ || Scan.succeed (mk_attribute add))
+ in
+ add_del_attribute ("inline", (add_inline, del_inline))
+ #> add_del_attribute ("post", (add_post, del_post))
+ #> Code.add_attribute ("unfold", Scan.succeed (Thm.declaration_attribute
+ (fn thm => Context.mapping (Codegen.add_unfold thm #> add_inline thm) I)))
+ end;
+
+val _ =
+ OuterSyntax.improper_command "print_codeproc" "print code preprocessor setup"
+ OuterKeyword.diag (Scan.succeed
+ (Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep
+ (print_codeproc o Toplevel.theory_of)));
+
+end; (*struct*)
--- a/src/Tools/code/code_target.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/Tools/code/code_target.ML Fri May 15 15:13:09 2009 -0700
@@ -286,7 +286,7 @@
fun gen_add_syntax_const prep_const target raw_c raw_syn thy =
let
val c = prep_const thy raw_c;
- fun check_args (syntax as (n, _)) = if n > Code_Unit.no_args thy c
+ fun check_args (syntax as (n, _)) = if n > Code.no_args thy c
then error ("Too many arguments in syntax for constant " ^ quote c)
else syntax;
in case raw_syn
@@ -319,8 +319,8 @@
| add (name, NONE) incls = Symtab.delete name incls;
in map_includes target (add args) thy end;
-val add_include = gen_add_include Code_Unit.check_const;
-val add_include_cmd = gen_add_include Code_Unit.read_const;
+val add_include = gen_add_include Code.check_const;
+val add_include_cmd = gen_add_include Code.read_const;
fun add_module_alias target (thyname, modlname) =
let
@@ -363,11 +363,11 @@
val allow_abort = gen_allow_abort (K I);
val add_reserved = add_reserved;
-val add_syntax_class_cmd = gen_add_syntax_class read_class Code_Unit.read_const;
+val add_syntax_class_cmd = gen_add_syntax_class read_class Code.read_const;
val add_syntax_inst_cmd = gen_add_syntax_inst read_class read_tyco;
val add_syntax_tyco_cmd = gen_add_syntax_tyco read_tyco;
-val add_syntax_const_cmd = gen_add_syntax_const Code_Unit.read_const;
-val allow_abort_cmd = gen_allow_abort Code_Unit.read_const;
+val add_syntax_const_cmd = gen_add_syntax_const Code.read_const;
+val allow_abort_cmd = gen_allow_abort Code.read_const;
fun the_literals thy =
let
@@ -387,15 +387,15 @@
local
fun labelled_name thy program name = case Graph.get_node program name
- of Code_Thingol.Fun (c, _) => quote (Code_Unit.string_of_const thy c)
+ of Code_Thingol.Fun (c, _) => quote (Code.string_of_const thy c)
| Code_Thingol.Datatype (tyco, _) => "type " ^ quote (Sign.extern_type thy tyco)
- | Code_Thingol.Datatypecons (c, _) => quote (Code_Unit.string_of_const thy c)
+ | Code_Thingol.Datatypecons (c, _) => quote (Code.string_of_const thy c)
| Code_Thingol.Class (class, _) => "class " ^ quote (Sign.extern_class thy class)
| Code_Thingol.Classrel (sub, super) => let
val Code_Thingol.Class (sub, _) = Graph.get_node program sub
val Code_Thingol.Class (super, _) = Graph.get_node program super
in quote (Sign.extern_class thy sub ^ " < " ^ Sign.extern_class thy super) end
- | Code_Thingol.Classparam (c, _) => quote (Code_Unit.string_of_const thy c)
+ | Code_Thingol.Classparam (c, _) => quote (Code.string_of_const thy c)
| Code_Thingol.Classinst ((class, (tyco, _)), _) => let
val Code_Thingol.Class (class, _) = Graph.get_node program class
val Code_Thingol.Datatype (tyco, _) = Graph.get_node program tyco
--- a/src/Tools/code/code_thingol.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/Tools/code/code_thingol.ML Fri May 15 15:13:09 2009 -0700
@@ -498,7 +498,7 @@
let
val thms = if null (Term.add_tfreesT ty []) orelse (null o fst o strip_type) ty
then raw_thms
- else (map o apfst) (Code_Unit.expand_eta thy 1) raw_thms;
+ else (map o apfst) (Code.expand_eta thy 1) raw_thms;
in
fold_map (translate_tyvar_sort thy algbr funcgr) vs
##>> translate_typ thy algbr funcgr ty
@@ -509,7 +509,7 @@
of SOME tyco => stmt_datatypecons tyco
| NONE => (case AxClass.class_of_param thy c
of SOME class => stmt_classparam class
- | NONE => stmt_fun (Code_Wellsorted.typ funcgr c, Code_Wellsorted.eqns funcgr c))
+ | NONE => stmt_fun (Code_Preproc.typ funcgr c, Code_Preproc.eqns funcgr c))
in ensure_stmt lookup_const (declare_const thy) stmt_const c end
and ensure_class thy (algbr as (_, algebra)) funcgr class =
let
@@ -603,7 +603,7 @@
and translate_const thy algbr funcgr thm (c, ty) =
let
val tys = Sign.const_typargs thy (c, ty);
- val sorts = (map snd o fst o Code_Wellsorted.typ funcgr) c;
+ val sorts = (map snd o fst o Code_Preproc.typ funcgr) c;
val tys_args = (fst o Term.strip_type) ty;
in
ensure_const thy algbr funcgr c
@@ -634,7 +634,7 @@
Term.strip_abs_eta 1 (the_single ts_clause)
in [(true, (Free v_ty, body))] end
else map (uncurry mk_clause)
- (AList.make (Code_Unit.no_args thy) case_pats ~~ ts_clause);
+ (AList.make (Code.no_args thy) case_pats ~~ ts_clause);
fun retermify ty (_, (IVar x, body)) =
(x, ty) `|-> body
| retermify _ (_, (pat, body)) =
@@ -748,7 +748,7 @@
fun generate_consts thy algebra funcgr =
fold_map (ensure_const thy algebra funcgr);
in
- invoke_generation thy (Code_Wellsorted.obtain thy cs []) generate_consts cs
+ invoke_generation thy (Code_Preproc.obtain thy cs []) generate_consts cs
|-> project_consts
end;
@@ -788,8 +788,8 @@
val vs'' = map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v)) vs';
in evaluator naming program ((vs'', (vs', ty')), t') deps end;
-fun eval_conv thy prep_sort = Code_Wellsorted.eval_conv thy prep_sort o base_evaluator thy;
-fun eval thy prep_sort postproc = Code_Wellsorted.eval thy prep_sort postproc o base_evaluator thy;
+fun eval_conv thy prep_sort = Code_Preproc.eval_conv thy prep_sort o base_evaluator thy;
+fun eval thy prep_sort postproc = Code_Preproc.eval thy prep_sort postproc o base_evaluator thy;
(** diagnostic commands **)
@@ -812,12 +812,12 @@
fun read_const_expr "*" = ([], consts_of NONE)
| read_const_expr s = if String.isSuffix ".*" s
then ([], consts_of (SOME (unsuffix ".*" s)))
- else ([Code_Unit.read_const thy s], []);
+ else ([Code.read_const thy s], []);
in pairself flat o split_list o map read_const_expr end;
fun code_depgr thy consts =
let
- val (_, eqngr) = Code_Wellsorted.obtain thy consts [];
+ val (_, eqngr) = Code_Preproc.obtain thy consts [];
val select = Graph.all_succs eqngr consts;
in
eqngr
@@ -825,7 +825,7 @@
|> Graph.map_nodes ((apsnd o map o apfst) (AxClass.overload thy))
end;
-fun code_thms thy = Pretty.writeln o Code_Wellsorted.pretty thy o code_depgr thy;
+fun code_thms thy = Pretty.writeln o Code_Preproc.pretty thy o code_depgr thy;
fun code_deps thy consts =
let
@@ -839,7 +839,7 @@
|> map (the o Symtab.lookup mapping)
|> distinct (op =);
val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss;
- fun namify consts = map (Code_Unit.string_of_const thy) consts
+ fun namify consts = map (Code.string_of_const thy) consts
|> commas;
val prgr = map (fn (consts, constss) =>
{ name = namify consts, ID = namify consts, dir = "", unfold = true,
--- a/src/Tools/code/code_wellsorted.ML Fri May 15 15:12:23 2009 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,342 +0,0 @@
-(* Title: Tools/code/code_wellsorted.ML
- Author: Florian Haftmann, TU Muenchen
-
-Producing well-sorted systems of code equations in a graph
-with explicit dependencies -- the Waisenhaus algorithm.
-*)
-
-signature CODE_WELLSORTED =
-sig
- type code_algebra
- type code_graph
- val eqns: code_graph -> string -> (thm * bool) list
- val typ: code_graph -> string -> (string * sort) list * typ
- val all: code_graph -> string list
- val pretty: theory -> code_graph -> Pretty.T
- val obtain: theory -> string list -> term list -> code_algebra * code_graph
- val eval_conv: theory -> (sort -> sort)
- -> (code_algebra -> code_graph -> (string * sort) list -> term -> cterm -> thm) -> cterm -> thm
- val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a)
- -> (code_algebra -> code_graph -> (string * sort) list -> term -> 'a) -> term -> 'a
-end
-
-structure Code_Wellsorted : CODE_WELLSORTED =
-struct
-
-(** the algebra and code equation graph types **)
-
-type code_algebra = (sort -> sort) * Sorts.algebra;
-type code_graph = (((string * sort) list * typ) * (thm * bool) list) Graph.T;
-
-fun eqns eqngr = these o Option.map snd o try (Graph.get_node eqngr);
-fun typ eqngr = fst o Graph.get_node eqngr;
-fun all eqngr = Graph.keys eqngr;
-
-fun pretty thy eqngr =
- AList.make (snd o Graph.get_node eqngr) (Graph.keys eqngr)
- |> (map o apfst) (Code_Unit.string_of_const thy)
- |> sort (string_ord o pairself fst)
- |> map (fn (s, thms) =>
- (Pretty.block o Pretty.fbreaks) (
- Pretty.str s
- :: map (Display.pretty_thm o fst) thms
- ))
- |> Pretty.chunks;
-
-
-(** the Waisenhaus algorithm **)
-
-(* auxiliary *)
-
-fun is_proper_class thy = can (AxClass.get_info thy);
-
-fun complete_proper_sort thy =
- Sign.complete_sort thy #> filter (is_proper_class thy);
-
-fun inst_params thy tyco =
- map (fn (c, _) => AxClass.param_of_inst thy (c, tyco))
- o maps (#params o AxClass.get_info thy);
-
-fun consts_of thy eqns = [] |> (fold o fold o fold_aterms)
- (fn Const (c, ty) => insert (op =) (c, Sign.const_typargs thy (c, Logic.unvarifyT ty)) | _ => I)
- (map (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of o fst) eqns);
-
-fun tyscm_rhss_of thy c eqns =
- let
- val tyscm = case eqns of [] => Code.default_typscheme thy c
- | ((thm, _) :: _) => Code_Unit.typscheme_eqn thy thm;
- val rhss = consts_of thy eqns;
- in (tyscm, rhss) end;
-
-
-(* data structures *)
-
-datatype const = Fun of string | Inst of class * string;
-
-fun const_ord (Fun c1, Fun c2) = fast_string_ord (c1, c2)
- | const_ord (Inst class_tyco1, Inst class_tyco2) =
- prod_ord fast_string_ord fast_string_ord (class_tyco1, class_tyco2)
- | const_ord (Fun _, Inst _) = LESS
- | const_ord (Inst _, Fun _) = GREATER;
-
-type var = const * int;
-
-structure Vargraph =
- GraphFun(type key = var val ord = prod_ord const_ord int_ord);
-
-datatype styp = Tyco of string * styp list | Var of var | Free;
-
-fun styp_of c_lhs (Type (tyco, tys)) = Tyco (tyco, map (styp_of c_lhs) tys)
- | styp_of c_lhs (TFree (v, _)) = case c_lhs
- of SOME (c, lhs) => Var (Fun c, find_index (fn (v', _) => v = v') lhs)
- | NONE => Free;
-
-type vardeps_data = ((string * styp list) list * class list) Vargraph.T
- * (((string * sort) list * (thm * bool) list) Symtab.table
- * (class * string) list);
-
-val empty_vardeps_data : vardeps_data =
- (Vargraph.empty, (Symtab.empty, []));
-
-
-(* retrieving equations and instances from the background context *)
-
-fun obtain_eqns thy eqngr c =
- case try (Graph.get_node eqngr) c
- of SOME ((lhs, _), eqns) => ((lhs, []), [])
- | NONE => let
- val eqns = Code.these_eqns thy c
- |> burrow_fst (Code_Unit.norm_args thy)
- |> burrow_fst (Code_Unit.norm_varnames thy);
- val ((lhs, _), rhss) = tyscm_rhss_of thy c eqns;
- in ((lhs, rhss), eqns) end;
-
-fun obtain_instance thy arities (inst as (class, tyco)) =
- case AList.lookup (op =) arities inst
- of SOME classess => (classess, ([], []))
- | NONE => let
- val all_classes = complete_proper_sort thy [class];
- val superclasses = remove (op =) class all_classes
- val classess = map (complete_proper_sort thy)
- (Sign.arity_sorts thy tyco [class]);
- val inst_params = inst_params thy tyco all_classes;
- in (classess, (superclasses, inst_params)) end;
-
-
-(* computing instantiations *)
-
-fun add_classes thy arities eqngr c_k new_classes vardeps_data =
- let
- val (styps, old_classes) = Vargraph.get_node (fst vardeps_data) c_k;
- val diff_classes = new_classes |> subtract (op =) old_classes;
- in if null diff_classes then vardeps_data
- else let
- val c_ks = Vargraph.imm_succs (fst vardeps_data) c_k |> insert (op =) c_k;
- in
- vardeps_data
- |> (apfst o Vargraph.map_node c_k o apsnd) (append diff_classes)
- |> fold (fn styp => fold (assert_typmatch_inst thy arities eqngr styp) new_classes) styps
- |> fold (fn c_k => add_classes thy arities eqngr c_k diff_classes) c_ks
- end end
-and add_styp thy arities eqngr c_k tyco_styps vardeps_data =
- let
- val (old_styps, classes) = Vargraph.get_node (fst vardeps_data) c_k;
- in if member (op =) old_styps tyco_styps then vardeps_data
- else
- vardeps_data
- |> (apfst o Vargraph.map_node c_k o apfst) (cons tyco_styps)
- |> fold (assert_typmatch_inst thy arities eqngr tyco_styps) classes
- end
-and add_dep thy arities eqngr c_k c_k' vardeps_data =
- let
- val (_, classes) = Vargraph.get_node (fst vardeps_data) c_k;
- in
- vardeps_data
- |> add_classes thy arities eqngr c_k' classes
- |> apfst (Vargraph.add_edge (c_k, c_k'))
- end
-and assert_typmatch_inst thy arities eqngr (tyco, styps) class vardeps_data =
- if can (Sign.arity_sorts thy tyco) [class]
- then vardeps_data
- |> assert_inst thy arities eqngr (class, tyco)
- |> fold_index (fn (k, styp) =>
- assert_typmatch thy arities eqngr styp (Inst (class, tyco), k)) styps
- else vardeps_data (*permissive!*)
-and assert_inst thy arities eqngr (inst as (class, tyco)) (vardeps_data as (_, (_, insts))) =
- if member (op =) insts inst then vardeps_data
- else let
- val (classess, (superclasses, inst_params)) =
- obtain_instance thy arities inst;
- in
- vardeps_data
- |> (apsnd o apsnd) (insert (op =) inst)
- |> fold_index (fn (k, _) =>
- apfst (Vargraph.new_node ((Inst (class, tyco), k), ([] ,[])))) classess
- |> fold (fn superclass => assert_inst thy arities eqngr (superclass, tyco)) superclasses
- |> fold (assert_fun thy arities eqngr) inst_params
- |> fold_index (fn (k, classes) =>
- add_classes thy arities eqngr (Inst (class, tyco), k) classes
- #> fold (fn superclass =>
- add_dep thy arities eqngr (Inst (superclass, tyco), k)
- (Inst (class, tyco), k)) superclasses
- #> fold (fn inst_param =>
- add_dep thy arities eqngr (Fun inst_param, k)
- (Inst (class, tyco), k)
- ) inst_params
- ) classess
- end
-and assert_typmatch thy arities eqngr (Tyco tyco_styps) c_k vardeps_data =
- vardeps_data
- |> add_styp thy arities eqngr c_k tyco_styps
- | assert_typmatch thy arities eqngr (Var c_k') c_k vardeps_data =
- vardeps_data
- |> add_dep thy arities eqngr c_k c_k'
- | assert_typmatch thy arities eqngr Free c_k vardeps_data =
- vardeps_data
-and assert_rhs thy arities eqngr (c', styps) vardeps_data =
- vardeps_data
- |> assert_fun thy arities eqngr c'
- |> fold_index (fn (k, styp) =>
- assert_typmatch thy arities eqngr styp (Fun c', k)) styps
-and assert_fun thy arities eqngr c (vardeps_data as (_, (eqntab, _))) =
- if Symtab.defined eqntab c then vardeps_data
- else let
- val ((lhs, rhss), eqns) = obtain_eqns thy eqngr c;
- val rhss' = (map o apsnd o map) (styp_of (SOME (c, lhs))) rhss;
- in
- vardeps_data
- |> (apsnd o apfst) (Symtab.update_new (c, (lhs, eqns)))
- |> fold_index (fn (k, _) =>
- apfst (Vargraph.new_node ((Fun c, k), ([] ,[])))) lhs
- |> fold_index (fn (k, (_, sort)) =>
- add_classes thy arities eqngr (Fun c, k) (complete_proper_sort thy sort)) lhs
- |> fold (assert_rhs thy arities eqngr) rhss'
- end;
-
-
-(* applying instantiations *)
-
-fun dicts_of thy (proj_sort, algebra) (T, sort) =
- let
- fun class_relation (x, _) _ = x;
- fun type_constructor tyco xs class =
- inst_params thy tyco (Sorts.complete_sort algebra [class])
- @ (maps o maps) fst xs;
- fun type_variable (TFree (_, sort)) = map (pair []) (proj_sort sort);
- in
- flat (Sorts.of_sort_derivation (Syntax.pp_global thy) algebra
- { class_relation = class_relation, type_constructor = type_constructor,
- type_variable = type_variable } (T, proj_sort sort)
- handle Sorts.CLASS_ERROR _ => [] (*permissive!*))
- end;
-
-fun add_arity thy vardeps (class, tyco) =
- AList.default (op =)
- ((class, tyco), map (fn k => (snd o Vargraph.get_node vardeps) (Inst (class, tyco), k))
- (0 upto Sign.arity_number thy tyco - 1));
-
-fun add_eqs thy vardeps (c, (proto_lhs, proto_eqns)) (rhss, eqngr) =
- if can (Graph.get_node eqngr) c then (rhss, eqngr)
- else let
- val lhs = map_index (fn (k, (v, _)) =>
- (v, snd (Vargraph.get_node vardeps (Fun c, k)))) proto_lhs;
- val inst_tab = Vartab.empty |> fold (fn (v, sort) =>
- Vartab.update ((v, 0), sort)) lhs;
- val eqns = proto_eqns
- |> (map o apfst) (Code_Unit.inst_thm thy inst_tab);
- val (tyscm, rhss') = tyscm_rhss_of thy c eqns;
- val eqngr' = Graph.new_node (c, (tyscm, eqns)) eqngr;
- in (map (pair c) rhss' @ rhss, eqngr') end;
-
-fun extend_arities_eqngr thy cs ts (arities, eqngr) =
- let
- val cs_rhss = (fold o fold_aterms) (fn Const (c_ty as (c, _)) =>
- insert (op =) (c, (map (styp_of NONE) o Sign.const_typargs thy) c_ty) | _ => I) ts [];
- val (vardeps, (eqntab, insts)) = empty_vardeps_data
- |> fold (assert_fun thy arities eqngr) cs
- |> fold (assert_rhs thy arities eqngr) cs_rhss;
- val arities' = fold (add_arity thy vardeps) insts arities;
- val pp = Syntax.pp_global thy;
- val algebra = Sorts.subalgebra pp (is_proper_class thy)
- (AList.lookup (op =) arities') (Sign.classes_of thy);
- val (rhss, eqngr') = Symtab.fold (add_eqs thy vardeps) eqntab ([], eqngr);
- fun deps_of (c, rhs) = c :: maps (dicts_of thy algebra)
- (rhs ~~ (map snd o fst o fst o Graph.get_node eqngr') c);
- val eqngr'' = fold (fn (c, rhs) => fold
- (curry Graph.add_edge c) (deps_of rhs)) rhss eqngr';
- in (algebra, (arities', eqngr'')) end;
-
-
-(** store **)
-
-structure Wellsorted = CodeDataFun
-(
- type T = ((string * class) * sort list) list * code_graph;
- val empty = ([], Graph.empty);
- fun purge thy cs (arities, eqngr) =
- let
- val del_cs = ((Graph.all_preds eqngr
- o filter (can (Graph.get_node eqngr))) cs);
- val del_arities = del_cs
- |> map_filter (AxClass.inst_of_param thy)
- |> maps (fn (c, tyco) =>
- (map (rpair tyco) o Sign.complete_sort thy o the_list
- o AxClass.class_of_param thy) c);
- val arities' = fold (AList.delete (op =)) del_arities arities;
- val eqngr' = Graph.del_nodes del_cs eqngr;
- in (arities', eqngr') end;
-);
-
-
-(** retrieval interfaces **)
-
-fun obtain thy cs ts = apsnd snd
- (Wellsorted.change_yield thy (extend_arities_eqngr thy cs ts));
-
-fun prepare_sorts_typ prep_sort
- = map_type_tfree (fn (v, sort) => TFree (v, prep_sort sort));
-
-fun prepare_sorts prep_sort (Const (c, ty)) =
- Const (c, prepare_sorts_typ prep_sort ty)
- | prepare_sorts prep_sort (t1 $ t2) =
- prepare_sorts prep_sort t1 $ prepare_sorts prep_sort t2
- | prepare_sorts prep_sort (Abs (v, ty, t)) =
- Abs (v, prepare_sorts_typ prep_sort ty, prepare_sorts prep_sort t)
- | prepare_sorts _ (t as Bound _) = t;
-
-fun gen_eval thy cterm_of conclude_evaluation prep_sort evaluator proto_ct =
- let
- val pp = Syntax.pp_global thy;
- val ct = cterm_of proto_ct;
- val _ = (Sign.no_frees pp o map_types (K dummyT) o Sign.no_vars pp)
- (Thm.term_of ct);
- val thm = Code.preprocess_conv thy ct;
- val ct' = Thm.rhs_of thm;
- val t' = Thm.term_of ct';
- val vs = Term.add_tfrees t' [];
- val consts = fold_aterms
- (fn Const (c, _) => insert (op =) c | _ => I) t' [];
-
- val t'' = prepare_sorts prep_sort t';
- val (algebra', eqngr') = obtain thy consts [t''];
- in conclude_evaluation (evaluator algebra' eqngr' vs t'' ct') thm end;
-
-fun simple_evaluator evaluator algebra eqngr vs t ct =
- evaluator algebra eqngr vs t;
-
-fun eval_conv thy =
- let
- fun conclude_evaluation thm2 thm1 =
- let
- val thm3 = Code.postprocess_conv thy (Thm.rhs_of thm2);
- in
- Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ =>
- error ("could not construct evaluation proof:\n"
- ^ (cat_lines o map Display.string_of_thm) [thm1, thm2, thm3])
- end;
- in gen_eval thy I conclude_evaluation end;
-
-fun eval thy prep_sort postproc evaluator = gen_eval thy (Thm.cterm_of thy)
- (K o postproc (Code.postprocess_term thy)) prep_sort (simple_evaluator evaluator);
-
-end; (*struct*)
--- a/src/Tools/nbe.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/Tools/nbe.ML Fri May 15 15:13:09 2009 -0700
@@ -436,7 +436,7 @@
let
fun subst_const f = map_aterms (fn t as Term.Const (c, ty) => Term.Const (f c, ty)
| t => t);
- val resubst_triv_consts = subst_const (Code_Unit.resubst_alias thy);
+ val resubst_triv_consts = subst_const (Code.resubst_alias thy);
val ty' = typ_of_itype program vs0 ty;
fun type_infer t =
singleton (TypeInfer.infer_types (Syntax.pp_global thy) (Sign.tsig_of thy) I
@@ -459,7 +459,7 @@
(* evaluation oracle *)
fun add_triv_classes thy = curry (Sorts.inter_sort (Sign.classes_of thy))
- (Code_Unit.triv_classes thy);
+ (Code.triv_classes thy);
fun mk_equals thy lhs raw_rhs =
let
--- a/src/Tools/quickcheck.ML Fri May 15 15:12:23 2009 -0700
+++ b/src/Tools/quickcheck.ML Fri May 15 15:13:09 2009 -0700
@@ -94,7 +94,7 @@
error "Term to be tested contains type variables";
val _ = null (Term.add_vars t []) orelse
error "Term to be tested contains schematic variables";
- val frees = map dest_Free (OldTerm.term_frees t);
+ val frees = Term.add_frees t [];
in (map fst frees, list_abs_free (frees, t)) end
fun test_term ctxt quiet generator_name size i t =
@@ -104,12 +104,12 @@
of NONE => if quiet then mk_testers ctxt t' else mk_testers_strict ctxt t'
| SOME name => [mk_tester_select name ctxt t'];
fun iterate f 0 = NONE
- | iterate f k = case f () handle Match => (if quiet then ()
+ | iterate f j = case f () handle Match => (if quiet then ()
else warning "Exception Match raised during quickcheck"; NONE)
- of NONE => iterate f (k - 1) | SOME q => SOME q;
+ of NONE => iterate f (j - 1) | SOME q => SOME q;
fun with_testers k [] = NONE
| with_testers k (tester :: testers) =
- case iterate (fn () => tester k) i
+ case iterate (fn () => tester (k - 1)) i
of NONE => with_testers k testers
| SOME q => SOME q;
fun with_size k = if k > size then NONE