--- a/Admin/isatest/isatest-statistics Wed Oct 28 11:42:31 2009 +0000
+++ b/Admin/isatest/isatest-statistics Wed Oct 28 11:43:06 2009 +0000
@@ -51,7 +51,7 @@
SESSIONS="$@"
case "$PLATFORM" in
- *para* | *-M*)
+ *para* | *-M* | sun-poly)
PARALLEL=true
;;
*)
--- a/Admin/isatest/isatest-stats Wed Oct 28 11:42:31 2009 +0000
+++ b/Admin/isatest/isatest-stats Wed Oct 28 11:43:06 2009 +0000
@@ -6,7 +6,7 @@
THIS=$(cd "$(dirname "$0")"; pwd -P)
-PLATFORMS="at-poly at64-poly mac-poly-M4 mac-poly64-M4 mac-poly-M8 mac-poly64-M8 at-poly-5.1-para-e at64-poly-5.1-para at-mac-poly-5.1-para afp at-sml-dev"
+PLATFORMS="at-poly at64-poly mac-poly-M4 mac-poly64-M4 mac-poly-M8 mac-poly64-M8 at-poly-5.1-para-e at64-poly-5.1-para at-mac-poly-5.1-para afp at-sml-dev sun-poly"
ISABELLE_SESSIONS="\
HOL-Plain \
--- a/doc-src/Nitpick/nitpick.tex Wed Oct 28 11:42:31 2009 +0000
+++ b/doc-src/Nitpick/nitpick.tex Wed Oct 28 11:43:06 2009 +0000
@@ -2091,11 +2091,11 @@
\url{http://cs.technion.ac.il/~gershman/HaifaSat.htm}.
\item[$\bullet$] \textbf{\textit{smart}}: If \textit{sat\_solver} is set to
-\textit{smart}, Nitpick selects the first solver among MiniSat, PicoSAT, zChaff,
-RSat, BerkMin, BerkMinAlloy, and Jerusat that is recognized by Isabelle. If none
-is found, it falls back on SAT4J, which should always be available. If
-\textit{verbose} is enabled, Nitpick displays which SAT solver was chosen.
-
+\textit{smart}, Nitpick selects the first solver among MiniSatJNI, MiniSat,
+PicoSAT, zChaffJNI, zChaff, RSat, BerkMin, BerkMinAlloy, and Jerusat that is
+recognized by Isabelle. If none is found, it falls back on SAT4J, which should
+always be available. If \textit{verbose} is enabled, Nitpick displays which SAT
+solver was chosen.
\end{enum}
\fussy
@@ -2381,7 +2381,7 @@
(``genuine'', ``likely\_genuine'', ``potential'', ``none'', or ``unknown''). The
\textit{params} type is a large record that lets you set Nitpick's options. The
current default options can be retrieved by calling the following function
-defined in the \textit{NitpickIsar} structure:
+defined in the \textit{Nitpick\_Isar} structure:
\prew
$\textbf{val}\,~\textit{default\_params} :\,
@@ -2392,7 +2392,7 @@
put into a \textit{params} record. Here is an example:
\prew
-$\textbf{val}\,~\textit{params} = \textit{NitpickIsar.default\_params}~\textit{thy}~[(\textrm{``}\textrm{timeout}\textrm{''},\, \textrm{``}\textrm{none}\textrm{''})]$ \\
+$\textbf{val}\,~\textit{params} = \textit{Nitpick\_Isar.default\_params}~\textit{thy}~[(\textrm{``}\textrm{timeout}\textrm{''},\, \textrm{``}\textrm{none}\textrm{''})]$ \\
$\textbf{val}\,~(\textit{outcome},\, \textit{state}') = \textit{Nitpick.pick\_nits\_in\_subgoal}~\begin{aligned}[t]
& \textit{state}~\textit{params}~\textit{false} \\[-2pt]
& \textit{subgoal}\end{aligned}$
--- a/src/FOLP/simp.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/FOLP/simp.ML Wed Oct 28 11:43:06 2009 +0000
@@ -254,13 +254,13 @@
val insert_thms = fold_rev insert_thm_warn;
-fun addrew(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thm) =
+fun addrew thm (SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}) =
let val thms = mk_simps thm
in SS{auto_tac=auto_tac,congs=congs, cong_net=cong_net, mk_simps=mk_simps,
simps = (thm,thms)::simps, simp_net = insert_thms thms simp_net}
end;
-val op addrews = Library.foldl addrew;
+fun ss addrews thms = fold addrew thms ss;
fun op addcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thms) =
let val congs' = thms @ congs;
@@ -287,7 +287,7 @@
mk_simps= normed_rews congs', simps=simps, simp_net=simp_net}
end;
-fun delrew(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thm) =
+fun delrew thm (SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}) =
let fun find((p as (th,ths))::ps',ps) =
if Thm.eq_thm_prop(thm,th) then (ths,ps@ps') else find(ps',p::ps)
| find([],simps') =
@@ -298,7 +298,7 @@
simps = simps', simp_net = delete_thms thms simp_net }
end;
-val op delrews = Library.foldl delrew;
+fun ss delrews thms = fold delrew thms ss;
fun op setauto(SS{congs,cong_net,mk_simps,simps,simp_net,...}, auto_tac) =
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Wed Oct 28 11:43:06 2009 +0000
@@ -35,7 +35,7 @@
"Itm vs bs (Add a b) = Itm vs bs a + Itm vs bs b"
"Itm vs bs (Sub a b) = Itm vs bs a - Itm vs bs b"
"Itm vs bs (Mul c a) = (Ipoly vs c) * Itm vs bs a"
- "Itm vs bs (CNP n c t) = (Ipoly vs c)*(bs!n) + Itm vs bs t"
+ "Itm vs bs (CNP n c t) = (Ipoly vs c)*(bs!n) + Itm vs bs t"
fun allpolys:: "(poly \<Rightarrow> bool) \<Rightarrow> tm \<Rightarrow> bool" where
@@ -153,14 +153,14 @@
{assume mxs: "m \<le> length (x#xs)" hence ?case using prems by (cases m, auto)}
moreover
{assume mxs: "\<not> (m \<le> length (x#xs))"
- have th: "length (removen n (x#xs)) = length xs"
- using removen_length[where n="n" and xs="x#xs"] nxs by simp
- with mxs have mxs':"m \<ge> length (removen n (x#xs))" by auto
- hence "(removen n (x#xs))!m = [] ! (m - length xs)"
- using th nth_length_exceeds[OF mxs'] by auto
- hence th: "(removen n (x#xs))!m = [] ! (m - (length (x#xs) - 1))"
- by auto
- hence ?case using nxs mln mxs by auto }
+ have th: "length (removen n (x#xs)) = length xs"
+ using removen_length[where n="n" and xs="x#xs"] nxs by simp
+ with mxs have mxs':"m \<ge> length (removen n (x#xs))" by auto
+ hence "(removen n (x#xs))!m = [] ! (m - length xs)"
+ using th nth_length_exceeds[OF mxs'] by auto
+ hence th: "(removen n (x#xs))!m = [] ! (m - (length (x#xs) - 1))"
+ by auto
+ hence ?case using nxs mln mxs by auto }
ultimately have ?case by blast
}
ultimately have ?case by blast
@@ -443,7 +443,7 @@
"fmsize (A p) = 4+ fmsize p"
"fmsize p = 1"
(* several lemmas about fmsize *)
-lemma fmsize_pos: "fmsize p > 0"
+lemma fmsize_pos: "fmsize p > 0"
by (induct p rule: fmsize.induct) simp_all
(* Semantics of formulae (fm) *)
@@ -1393,16 +1393,16 @@
using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto}
moreover {assume cp: "?c > 0"
{fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
- using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e < 0" by simp
hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))"
- using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto} hence ?case by auto}
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto} hence ?case by auto}
moreover {assume cp: "?c < 0"
{fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
- using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e > 0" by simp
hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))"
- using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto} hence ?case by auto}
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto} hence ?case by auto}
ultimately show ?case by blast
next
case (4 c e) hence nbe: "tmbound0 e" by simp
@@ -1414,16 +1414,16 @@
moreover {assume "?c = 0" hence ?case using eqs by auto}
moreover {assume cp: "?c > 0"
{fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
- using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e < 0" by simp
hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))"
- using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
moreover {assume cp: "?c < 0"
{fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
- using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e > 0" by simp
hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))"
- using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
ultimately show ?case by blast
next
case (5 c e) hence nbe: "tmbound0 e" by simp
@@ -1436,16 +1436,16 @@
moreover {assume "?c = 0" hence ?case using eqs by auto}
moreover {assume cp: "?c > 0"
{fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
- using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e < 0" by simp
hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))"
- using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
moreover {assume cp: "?c < 0"
{fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
- using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e > 0" by simp
hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))"
- using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] cp by auto} hence ?case by auto}
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] cp by auto} hence ?case by auto}
ultimately show ?case by blast
next
case (6 c e) hence nbe: "tmbound0 e" by simp
@@ -1458,16 +1458,16 @@
moreover {assume "?c = 0" hence ?case using eqs by auto}
moreover {assume cp: "?c > 0"
{fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
- using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e < 0" by simp
hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))"
- using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
moreover {assume cp: "?c < 0"
{fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
- using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e > 0" by simp
hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))"
- using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
ultimately show ?case by blast
qed (auto)
@@ -1489,16 +1489,16 @@
using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto}
moreover {assume cp: "?c > 0"
{fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
- using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e > 0" by simp
hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))"
- using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto} hence ?case by auto}
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto} hence ?case by auto}
moreover {assume cp: "?c < 0"
{fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
- using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e < 0" by simp
hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))"
- using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto} hence ?case by auto}
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto} hence ?case by auto}
ultimately show ?case by blast
next
case (4 c e) hence nbe: "tmbound0 e" by simp
@@ -1510,16 +1510,16 @@
moreover {assume "?c = 0" hence ?case using eqs by auto}
moreover {assume cp: "?c > 0"
{fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
- using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e > 0" by simp
hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))"
- using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
moreover {assume cp: "?c < 0"
{fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
- using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e < 0" by simp
hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))"
- using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
ultimately show ?case by blast
next
case (5 c e) hence nbe: "tmbound0 e" by simp
@@ -1532,16 +1532,16 @@
moreover {assume "?c = 0" hence ?case using eqs by auto}
moreover {assume cp: "?c > 0"
{fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
- using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e > 0" by simp
hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))"
- using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
moreover {assume cp: "?c < 0"
{fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
- using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e < 0" by simp
hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))"
- using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] cp by auto} hence ?case by auto}
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] cp by auto} hence ?case by auto}
ultimately show ?case by blast
next
case (6 c e) hence nbe: "tmbound0 e" by simp
@@ -1554,16 +1554,16 @@
moreover {assume "?c = 0" hence ?case using eqs by auto}
moreover {assume cp: "?c > 0"
{fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
- using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e > 0" by simp
hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))"
- using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
moreover {assume cp: "?c < 0"
{fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
- using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
hence "?c * x + ?e < 0" by simp
hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))"
- using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
ultimately show ?case by blast
qed (auto)
@@ -1698,10 +1698,10 @@
{assume c: "?N c > 0"
from px pos_less_divide_eq[OF c, where a="x" and b="-?Nt x s"]
have px': "x < - ?Nt x s / ?N c"
- by (auto simp add: not_less ring_simps)
+ by (auto simp add: not_less ring_simps)
{assume y: "y < - ?Nt x s / ?N c"
hence "y * ?N c < - ?Nt x s"
- by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+ by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
moreover
@@ -1715,10 +1715,10 @@
{assume c: "?N c < 0"
from px neg_divide_less_eq[OF c, where a="x" and b="-?Nt x s"]
have px': "x > - ?Nt x s / ?N c"
- by (auto simp add: not_less ring_simps)
+ by (auto simp add: not_less ring_simps)
{assume y: "y > - ?Nt x s / ?N c"
hence "y * ?N c < - ?Nt x s"
- by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+ by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
moreover
@@ -1746,7 +1746,7 @@
have px': "x <= - ?Nt x s / ?N c" by (simp add: not_less ring_simps)
{assume y: "y < - ?Nt x s / ?N c"
hence "y * ?N c < - ?Nt x s"
- by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+ by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
moreover
@@ -1762,7 +1762,7 @@
have px': "x >= - ?Nt x s / ?N c" by (simp add: ring_simps)
{assume y: "y > - ?Nt x s / ?N c"
hence "y * ?N c < - ?Nt x s"
- by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+ by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
moreover
@@ -1899,8 +1899,8 @@
moreover{
assume "\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p"
then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M"
- and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
- by blast
+ and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
+ by blast
from t1M have "\<exists> (t1n,t1u) \<in> ?U. t1 = - ?Nt a t1u / ?N t1n" by auto
then obtain "t1u" "t1n" where t1uU: "(t1n,t1u) \<in> ?U" and t1u: "t1 = - ?Nt a t1u / ?N t1n" by blast
from t2M have "\<exists> (t2n,t2u) \<in> ?U. t2 = - ?Nt a t2u / ?N t2n" by auto
@@ -2236,7 +2236,7 @@
using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_less_cancel_left_disj[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp
also have "\<dots> \<longleftrightarrow> ?a*?t - (1 + 1)*?c *?r < 0"
using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using c d nc nd
apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
@@ -2270,7 +2270,7 @@
using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_less_cancel_left_disj[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp
also have "\<dots> \<longleftrightarrow> ?a*?s - (1 + 1)*?d *?r < 0"
using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using c d nc nd
apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
@@ -2392,7 +2392,7 @@
using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_le_cancel_left[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp
also have "\<dots> \<longleftrightarrow> ?a*?t - (1 + 1)*?c *?r <= 0"
using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using c d nc nd
apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
@@ -2426,7 +2426,7 @@
using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_le_cancel_left[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp
also have "\<dots> \<longleftrightarrow> ?a*?s - (1 + 1)*?d *?r <= 0"
using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using c d nc nd
apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
@@ -2638,8 +2638,8 @@
moreover
{assume z: "?c \<noteq> 0" "?d \<noteq> 0"
from z have ?F using ct
- apply - apply (rule bexI[where x = "(c,t)"], simp_all)
- by (rule bexI[where x = "(d,s)"], simp_all)
+ apply - apply (rule bexI[where x = "(c,t)"], simp_all)
+ by (rule bexI[where x = "(d,s)"], simp_all)
hence ?D by blast}
ultimately have ?D by auto}
ultimately show "?D" by blast
@@ -2799,17 +2799,17 @@
from H(1) th have "isnpoly n" by blast
hence nn: "isnpoly (n *\<^sub>p (C (-2,1)))" by (simp_all add: polymul_norm n2)
have nn': "allpolys isnpoly (CP (~\<^sub>p (n *\<^sub>p C (-2, 1))))"
- by (simp add: polyneg_norm nn)
+ by (simp add: polyneg_norm nn)
hence nn2: "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>n \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" using H(2) nn' nn
- by (auto simp add: msubst2_def lt zero_less_mult_iff mult_less_0_iff)
+ by (auto simp add: msubst2_def lt zero_less_mult_iff mult_less_0_iff)
from msubst2[OF lp nn nn2(1), of x bs t]
have "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p"
- using H(2) nn2 by (simp add: of_int_minus2 del: minus_add_distrib)}
+ using H(2) nn2 by (simp add: of_int_minus2 del: minus_add_distrib)}
moreover
{fix n t assume H: "(n, t)\<in>set (uset p)" "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p"
from H(1) th have "isnpoly n" by blast
hence nn: "isnpoly (n *\<^sub>p (C (-2,1)))" "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
- using H(2) by (simp_all add: polymul_norm n2)
+ using H(2) by (simp_all add: polymul_norm n2)
from msubst2[OF lp nn, of x bs t] have "?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)" using H(2,3) by (simp add: of_int_minus2 del: minus_add_distrib)}
ultimately show ?thesis by blast
qed
@@ -2820,21 +2820,21 @@
"Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))"
from H(1,2) th have "isnpoly c" "isnpoly d" by blast+
hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)"
- by (simp_all add: polymul_norm n2)
+ by (simp_all add: polymul_norm n2)
have stupid: "allpolys isnpoly (CP (~\<^sub>p (C (-2, 1) *\<^sub>p c *\<^sub>p d)))" "allpolys isnpoly (CP ((C (-2, 1) *\<^sub>p c *\<^sub>p d)))"
- by (simp_all add: polyneg_norm nn)
+ by (simp_all add: polyneg_norm nn)
have nn': "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
- using H(3) by (auto simp add: msubst2_def lt[OF stupid(1)] lt[OF stupid(2)] zero_less_mult_iff mult_less_0_iff)
+ using H(3) by (auto simp add: msubst2_def lt[OF stupid(1)] lt[OF stupid(2)] zero_less_mult_iff mult_less_0_iff)
from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn'
have "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p"
- apply (simp add: add_divide_distrib of_int_minus2 del: minus_add_distrib)
- by (simp add: mult_commute)}
+ apply (simp add: add_divide_distrib of_int_minus2 del: minus_add_distrib)
+ by (simp add: mult_commute)}
moreover
{fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)"
"\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p"
from H(1,2) th have "isnpoly c" "isnpoly d" by blast+
hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
- using H(3,4) by (simp_all add: polymul_norm n2)
+ using H(3,4) by (simp_all add: polymul_norm n2)
from msubst2[OF lp nn, of x bs ] H(3,4,5)
have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))" apply (simp add: add_divide_distrib of_int_minus2 del: minus_add_distrib)by (simp add: mult_commute)}
ultimately show ?thesis by blast
@@ -2906,16 +2906,16 @@
proof
assume H: ?lhs
hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
- by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)] mult_less_0_iff zero_less_mult_iff)
+ by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)] mult_less_0_iff zero_less_mult_iff)
from msubst2[OF lq norm2(1) z(1), of x bs]
- msubst2[OF lq norm2(2) z(2), of x bs] H
+ msubst2[OF lq norm2(2) z(2), of x bs] H
show ?rhs by (simp add: ring_simps)
next
assume H: ?rhs
hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
- by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)] mult_less_0_iff zero_less_mult_iff)
+ by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)] mult_less_0_iff zero_less_mult_iff)
from msubst2[OF lq norm2(1) z(1), of x bs]
- msubst2[OF lq norm2(2) z(2), of x bs] H
+ msubst2[OF lq norm2(2) z(2), of x bs] H
show ?lhs by (simp add: ring_simps)
qed}
hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (?s (x, y)) \<longleftrightarrow> ?I (?s (y, x))"
--- a/src/HOL/Decision_Procs/Polynomial_List.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy Wed Oct 28 11:43:06 2009 +0000
@@ -1,8 +1,8 @@
-(* Title: HOL/Decision_Procs/Polynomial_List.thy
- Author: Amine Chaieb
+(* Title: HOL/Decision_Procs/Polynomial_List.thy
+ Author: Amine Chaieb
*)
-header{*Univariate Polynomials as Lists *}
+header {* Univariate Polynomials as Lists *}
theory Polynomial_List
imports Main
@@ -653,7 +653,7 @@
apply (simp add: divides_def fun_eq poly_mult)
apply (rule_tac x = "[]" in exI)
apply (auto dest!: order2 [where a=a]
- intro: poly_exp_divides simp del: pexp_Suc)
+ intro: poly_exp_divides simp del: pexp_Suc)
done
lemma order_decomp:
@@ -780,4 +780,5 @@
done
lemma poly_Sing: "poly [c] x = c" by simp
+
end
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Wed Oct 28 11:43:06 2009 +0000
@@ -391,7 +391,7 @@
{assume nn':"\<not> n' < n" hence "n < n' \<or> n = n'" by arith
moreover {assume "n < n'" with prems have ?case by simp }
moreover {assume eq: "n = n'" hence ?case using prems
- by (cases "p +\<^sub>p p' = 0\<^sub>p", auto simp add: Let_def) }
+ by (cases "p +\<^sub>p p' = 0\<^sub>p", auto simp add: Let_def) }
ultimately have ?case by blast}
ultimately show ?case by blast
qed simp_all
@@ -413,10 +413,10 @@
by simp_all
{assume "?c = 0\<^sub>N" hence ?case by auto}
moreover {assume cnz: "?c \<noteq> 0\<^sub>N"
- from "2.hyps"(1)[rule_format,where xb="n'", OF cnz n(1) n(3)]
- "2.hyps"(2)[rule_format, where x="Suc n'"
- and xa="Suc n'" and xb = "n'", OF cnz ] cnz n have ?case
- by (auto simp add: min_def)}
+ from "2.hyps"(1)[rule_format,where xb="n'", OF cnz n(1) n(3)]
+ "2.hyps"(2)[rule_format, where x="Suc n'"
+ and xa="Suc n'" and xb = "n'", OF cnz ] cnz n have ?case
+ by (auto simp add: min_def)}
ultimately show ?case by blast
next
case (2 n0 n1) thus ?case by auto
@@ -431,10 +431,10 @@
by simp_all
{assume "?c' = 0\<^sub>N" hence ?case by auto}
moreover {assume cnz: "?c' \<noteq> 0\<^sub>N"
- from "3.hyps"(1)[rule_format,where xb="n", OF cnz n(3) n(1)]
- "3.hyps"(2)[rule_format, where x="Suc n"
- and xa="Suc n" and xb = "n", OF cnz ] cnz n have ?case
- by (auto simp add: min_def)}
+ from "3.hyps"(1)[rule_format,where xb="n", OF cnz n(3) n(1)]
+ "3.hyps"(2)[rule_format, where x="Suc n"
+ and xa="Suc n" and xb = "n", OF cnz ] cnz n have ?case
+ by (auto simp add: min_def)}
ultimately show ?case by blast
next
case (2 n0 n1) thus ?case apply auto done
@@ -446,32 +446,32 @@
{fix n0 n1
assume "isnpolyh ?cnp n0" and "isnpolyh ?cnp' n1"
hence cnp: "isnpolyh ?cnp n" and cnp': "isnpolyh ?cnp' n'"
- and np: "isnpolyh p n" and nc: "isnpolyh c (Suc n)"
- and np': "isnpolyh p' n'" and nc': "isnpolyh c' (Suc n')"
- and nn0: "n \<ge> n0" and nn1:"n' \<ge> n1"
- by simp_all
+ and np: "isnpolyh p n" and nc: "isnpolyh c (Suc n)"
+ and np': "isnpolyh p' n'" and nc': "isnpolyh c' (Suc n')"
+ and nn0: "n \<ge> n0" and nn1:"n' \<ge> n1"
+ by simp_all
have "n < n' \<or> n' < n \<or> n' = n" by auto
moreover
{assume nn': "n < n'"
- with "4.hyps"(5)[rule_format, OF nn' np cnp', where xb ="n"]
- "4.hyps"(6)[rule_format, OF nn' nc cnp', where xb="n"] nn' nn0 nn1 cnp
- have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
- by (simp add: min_def) }
+ with "4.hyps"(5)[rule_format, OF nn' np cnp', where xb ="n"]
+ "4.hyps"(6)[rule_format, OF nn' nc cnp', where xb="n"] nn' nn0 nn1 cnp
+ have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
+ by (simp add: min_def) }
moreover
{assume nn': "n > n'" hence stupid: "n' < n \<and> \<not> n < n'" by arith
- with "4.hyps"(3)[rule_format, OF stupid cnp np', where xb="n'"]
- "4.hyps"(4)[rule_format, OF stupid cnp nc', where xb="Suc n'"]
- nn' nn0 nn1 cnp'
- have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
- by (cases "Suc n' = n", simp_all add: min_def)}
+ with "4.hyps"(3)[rule_format, OF stupid cnp np', where xb="n'"]
+ "4.hyps"(4)[rule_format, OF stupid cnp nc', where xb="Suc n'"]
+ nn' nn0 nn1 cnp'
+ have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
+ by (cases "Suc n' = n", simp_all add: min_def)}
moreover
{assume nn': "n' = n" hence stupid: "\<not> n' < n \<and> \<not> n < n'" by arith
- from "4.hyps"(1)[rule_format, OF stupid cnp np', where xb="n"]
- "4.hyps"(2)[rule_format, OF stupid cnp nc', where xb="n"] nn' cnp cnp' nn1
-
- have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
- by simp (rule polyadd_normh,simp_all add: min_def isnpolyh_mono[OF nn0]) }
+ from "4.hyps"(1)[rule_format, OF stupid cnp np', where xb="n"]
+ "4.hyps"(2)[rule_format, OF stupid cnp nc', where xb="n"] nn' cnp cnp' nn1
+
+ have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
+ by simp (rule polyadd_normh,simp_all add: min_def isnpolyh_mono[OF nn0]) }
ultimately show "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)" by blast }
note th = this
{fix n0 n1 m
@@ -484,86 +484,86 @@
have "n'<n \<or> n < n' \<or> n' = n" by auto
moreover
{assume "n' < n \<or> n < n'"
- with "4.hyps" np np' m
- have ?eq apply (cases "n' < n", simp_all)
- apply (erule allE[where x="n"],erule allE[where x="n"],auto)
- done }
+ with "4.hyps" np np' m
+ have ?eq apply (cases "n' < n", simp_all)
+ apply (erule allE[where x="n"],erule allE[where x="n"],auto)
+ done }
moreover
{assume nn': "n' = n" hence nn:"\<not> n' < n \<and> \<not> n < n'" by arith
- from "4.hyps"(1)[rule_format, OF nn, where x="n" and xa ="n'" and xb="n"]
- "4.hyps"(2)[rule_format, OF nn, where x="n" and xa ="Suc n'" and xb="n"]
- np np' nn'
- have norm: "isnpolyh ?cnp n" "isnpolyh c' (Suc n)" "isnpolyh (?cnp *\<^sub>p c') n"
- "isnpolyh p' n" "isnpolyh (?cnp *\<^sub>p p') n" "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n"
- "(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)"
- "?cnp *\<^sub>p p' \<noteq> 0\<^sub>p" by (auto simp add: min_def)
- {assume mn: "m = n"
- from "4.hyps"(1)[rule_format, OF nn norm(1,4), where xb="n"]
- "4.hyps"(2)[rule_format, OF nn norm(1,2), where xb="n"] norm nn' mn
- have degs: "degreen (?cnp *\<^sub>p c') n =
- (if c'=0\<^sub>p then 0 else ?d1 + degreen c' n)"
- "degreen (?cnp *\<^sub>p p') n = ?d1 + degreen p' n" by (simp_all add: min_def)
- from degs norm
- have th1: "degreen(?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n" by simp
- hence neq: "degreen (?cnp *\<^sub>p c') n \<noteq> degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n"
- by simp
- have nmin: "n \<le> min n n" by (simp add: min_def)
- from polyadd_different_degreen[OF norm(3,6) neq nmin] th1
- have deg: "degreen (CN c n p *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n = degreen (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp
- from "4.hyps"(1)[rule_format, OF nn norm(1,4), where xb="n"]
- "4.hyps"(2)[rule_format, OF nn norm(1,2), where xb="n"]
- mn norm m nn' deg
- have ?eq by simp}
- moreover
- {assume mn: "m \<noteq> n" hence mn': "m < n" using m np by auto
- from nn' m np have max1: "m \<le> max n n" by simp
- hence min1: "m \<le> min n n" by simp
- hence min2: "m \<le> min n (Suc n)" by simp
- {assume "c' = 0\<^sub>p"
- from `c' = 0\<^sub>p` have ?eq
- using "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
- "4.hyps"(2)[rule_format, OF nn norm(1,2) min2] mn nn'
- apply simp
- done}
- moreover
- {assume cnz: "c' \<noteq> 0\<^sub>p"
- from "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
- "4.hyps"(2)[rule_format, OF nn norm(1,2) min2]
- degreen_polyadd[OF norm(3,6) max1]
+ from "4.hyps"(1)[rule_format, OF nn, where x="n" and xa ="n'" and xb="n"]
+ "4.hyps"(2)[rule_format, OF nn, where x="n" and xa ="Suc n'" and xb="n"]
+ np np' nn'
+ have norm: "isnpolyh ?cnp n" "isnpolyh c' (Suc n)" "isnpolyh (?cnp *\<^sub>p c') n"
+ "isnpolyh p' n" "isnpolyh (?cnp *\<^sub>p p') n" "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n"
+ "(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)"
+ "?cnp *\<^sub>p p' \<noteq> 0\<^sub>p" by (auto simp add: min_def)
+ {assume mn: "m = n"
+ from "4.hyps"(1)[rule_format, OF nn norm(1,4), where xb="n"]
+ "4.hyps"(2)[rule_format, OF nn norm(1,2), where xb="n"] norm nn' mn
+ have degs: "degreen (?cnp *\<^sub>p c') n =
+ (if c'=0\<^sub>p then 0 else ?d1 + degreen c' n)"
+ "degreen (?cnp *\<^sub>p p') n = ?d1 + degreen p' n" by (simp_all add: min_def)
+ from degs norm
+ have th1: "degreen(?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n" by simp
+ hence neq: "degreen (?cnp *\<^sub>p c') n \<noteq> degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n"
+ by simp
+ have nmin: "n \<le> min n n" by (simp add: min_def)
+ from polyadd_different_degreen[OF norm(3,6) neq nmin] th1
+ have deg: "degreen (CN c n p *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n = degreen (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp
+ from "4.hyps"(1)[rule_format, OF nn norm(1,4), where xb="n"]
+ "4.hyps"(2)[rule_format, OF nn norm(1,2), where xb="n"]
+ mn norm m nn' deg
+ have ?eq by simp}
+ moreover
+ {assume mn: "m \<noteq> n" hence mn': "m < n" using m np by auto
+ from nn' m np have max1: "m \<le> max n n" by simp
+ hence min1: "m \<le> min n n" by simp
+ hence min2: "m \<le> min n (Suc n)" by simp
+ {assume "c' = 0\<^sub>p"
+ from `c' = 0\<^sub>p` have ?eq
+ using "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
+ "4.hyps"(2)[rule_format, OF nn norm(1,2) min2] mn nn'
+ apply simp
+ done}
+ moreover
+ {assume cnz: "c' \<noteq> 0\<^sub>p"
+ from "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
+ "4.hyps"(2)[rule_format, OF nn norm(1,2) min2]
+ degreen_polyadd[OF norm(3,6) max1]
- have "degreen (?cnp *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (?cnp *\<^sub>p p')) m
- \<le> max (degreen (?cnp *\<^sub>p c') m) (degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) m)"
- using mn nn' cnz np np' by simp
- with "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
- "4.hyps"(2)[rule_format, OF nn norm(1,2) min2]
- degreen_0[OF norm(3) mn'] have ?eq using nn' mn cnz np np' by clarsimp}
- ultimately have ?eq by blast }
- ultimately have ?eq by blast}
+ have "degreen (?cnp *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (?cnp *\<^sub>p p')) m
+ \<le> max (degreen (?cnp *\<^sub>p c') m) (degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) m)"
+ using mn nn' cnz np np' by simp
+ with "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
+ "4.hyps"(2)[rule_format, OF nn norm(1,2) min2]
+ degreen_0[OF norm(3) mn'] have ?eq using nn' mn cnz np np' by clarsimp}
+ ultimately have ?eq by blast }
+ ultimately have ?eq by blast}
ultimately show ?eq by blast}
note degth = this
{ case (2 n0 n1)
hence np: "isnpolyh ?cnp n0" and np': "isnpolyh ?cnp' n1"
- and m: "m \<le> min n0 n1" by simp_all
+ and m: "m \<le> min n0 n1" by simp_all
hence mn: "m \<le> n" by simp
let ?c0p = "CN 0\<^sub>p n (?cnp *\<^sub>p p')"
{assume C: "?cnp *\<^sub>p c' +\<^sub>p ?c0p = 0\<^sub>p" "n' = n"
- hence nn: "\<not>n' < n \<and> \<not> n<n'" by simp
- from "4.hyps"(1) [rule_format, OF nn, where x="n" and xa = "n" and xb="n"]
- "4.hyps"(2) [rule_format, OF nn, where x="n" and xa = "Suc n" and xb="n"]
- np np' C(2) mn
- have norm: "isnpolyh ?cnp n" "isnpolyh c' (Suc n)" "isnpolyh (?cnp *\<^sub>p c') n"
- "isnpolyh p' n" "isnpolyh (?cnp *\<^sub>p p') n" "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n"
- "(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)"
- "?cnp *\<^sub>p p' \<noteq> 0\<^sub>p"
- "degreen (?cnp *\<^sub>p c') n = (if c'=0\<^sub>p then 0 else degreen ?cnp n + degreen c' n)"
- "degreen (?cnp *\<^sub>p p') n = degreen ?cnp n + degreen p' n"
- by (simp_all add: min_def)
-
- from norm have cn: "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp
- have degneq: "degreen (?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n"
- using norm by simp
- from polyadd_eq_const_degreen[OF norm(3) cn C(1), where m="n"] degneq
- have "False" by simp }
+ hence nn: "\<not>n' < n \<and> \<not> n<n'" by simp
+ from "4.hyps"(1) [rule_format, OF nn, where x="n" and xa = "n" and xb="n"]
+ "4.hyps"(2) [rule_format, OF nn, where x="n" and xa = "Suc n" and xb="n"]
+ np np' C(2) mn
+ have norm: "isnpolyh ?cnp n" "isnpolyh c' (Suc n)" "isnpolyh (?cnp *\<^sub>p c') n"
+ "isnpolyh p' n" "isnpolyh (?cnp *\<^sub>p p') n" "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n"
+ "(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)"
+ "?cnp *\<^sub>p p' \<noteq> 0\<^sub>p"
+ "degreen (?cnp *\<^sub>p c') n = (if c'=0\<^sub>p then 0 else degreen ?cnp n + degreen c' n)"
+ "degreen (?cnp *\<^sub>p p') n = degreen ?cnp n + degreen p' n"
+ by (simp_all add: min_def)
+
+ from norm have cn: "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp
+ have degneq: "degreen (?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n"
+ using norm by simp
+ from polyadd_eq_const_degreen[OF norm(3) cn C(1), where m="n"] degneq
+ have "False" by simp }
thus ?case using "4.hyps" by clarsimp}
qed auto
@@ -1022,8 +1022,8 @@
have "\<forall>x. poly (polypoly' (?ts @ xs) p) x = poly [] x" by simp
hence "poly (polypoly' (?ts @ xs) p) = poly []" by (auto intro: ext)
hence "\<forall> c \<in> set (coefficients p). Ipoly (?ts @ xs) (decrpoly c) = 0"
- thm poly_zero
- using poly_zero[where ?'a='a] by (simp add: polypoly'_def list_all_iff)
+ thm poly_zero
+ using poly_zero[where ?'a='a] by (simp add: polypoly'_def list_all_iff)
with coefficients_head[of p, symmetric]
have th0: "Ipoly (?ts @ xs) ?hd = 0" by simp
from bs have wf'': "wf_bs ?ts ?hd" unfolding wf_bs_def by simp
@@ -1272,18 +1272,18 @@
by (simp add: min_def)
{assume np: "n > 0"
with nn' head_isnpolyh_Suc'[OF np nth]
- head_isnpolyh_Suc'[OF np norm(5)] head_isnpolyh_Suc'[OF np norm(6)[simplified nn']]
+ head_isnpolyh_Suc'[OF np norm(5)] head_isnpolyh_Suc'[OF np norm(6)[simplified nn']]
have ?case by simp}
moreover
{moreover assume nz: "n = 0"
from polymul_degreen[OF norm(5,4), where m="0"]
- polymul_degreen[OF norm(5,3), where m="0"] nn' nz degree_eq_degreen0
+ polymul_degreen[OF norm(5,3), where m="0"] nn' nz degree_eq_degreen0
norm(5,6) degree_npolyhCN[OF norm(6)]
have dth:"degree (CN c 0 p *\<^sub>p c') < degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))" by simp
hence dth':"degree (CN c 0 p *\<^sub>p c') \<noteq> degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))" by simp
from polyadd_head[OF ncnpc'[simplified nz] ncnpp0'[simplified nz] dth'] dth
have ?case using norm prems(2)[rule_format, OF stupid norm(5,3)]
- prems(3)[rule_format, OF stupid norm(5,4)] nn' nz by simp }
+ prems(3)[rule_format, OF stupid norm(5,4)] nn' nz by simp }
ultimately have ?case by (cases n) auto}
ultimately show ?case by blast
qed simp_all
@@ -1399,182 +1399,182 @@
moreover
{assume dn': "\<not> d < n" hence dn: "d \<ge> n" by arith
have degsp': "degree s = degree ?p'"
- using ds dn ndp funpow_shift1_degree[where k = "d - n" and p="p"] by simp
+ using ds dn ndp funpow_shift1_degree[where k = "d - n" and p="p"] by simp
{assume ba: "?b = a"
- hence headsp': "head s = head ?p'" using ap headp' by simp
- have nr: "isnpolyh (s -\<^sub>p ?p') 0" using polysub_normh[OF ns np'] by simp
- from ds degree_polysub_samehead[OF ns np' headsp' degsp']
- have "degree (s -\<^sub>p ?p') < d \<or> s -\<^sub>p ?p' = 0\<^sub>p" by simp
- moreover
- {assume deglt:"degree (s -\<^sub>p ?p') < d"
- from H[rule_format, OF deglt nr,simplified]
- have domsp: "?D (a, n, p, k, s -\<^sub>p ?p')" by blast
- have dom: ?dths apply (rule polydivide_aux_real_domintros)
- using ba ds dn' domsp by simp_all
- from polydivide_aux.psimps[OF dom] sz dn' ba ds
- have eq: "polydivide_aux (a,n,p,k,s) = polydivide_aux (a,n,p,k, s -\<^sub>p ?p')"
- by (simp add: Let_def)
- {assume h1: "polydivide_aux (a, n, p, k, s) = (k', r)"
- from H[rule_format, OF deglt nr, where xa = "k" and xb="k'" and xc="r", simplified]
- trans[OF eq[symmetric] h1]
- have kk': "k \<le> k'" and nr:"\<exists>nr. isnpolyh r nr" and dr: "degree r = 0 \<or> degree r < degree p"
- and q1:"\<exists>q nq. isnpolyh q nq \<and> (a ^\<^sub>p k' - k *\<^sub>p (s -\<^sub>p ?p') = p *\<^sub>p q +\<^sub>p r)" by auto
- from q1 obtain q n1 where nq: "isnpolyh q n1"
- and asp:"a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p') = p *\<^sub>p q +\<^sub>p r" by blast
- from nr obtain nr where nr': "isnpolyh r nr" by blast
- from polymul_normh[OF nakk' ns] have nakks': "isnpolyh (a ^\<^sub>p (k' - k) *\<^sub>p s) 0" by simp
- from polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]
- have nq': "isnpolyh (?akk' *\<^sub>p ?xdn +\<^sub>p q) 0" by simp
- from polyadd_normh[OF polymul_normh[OF np
- polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]] nr']
- have nqr': "isnpolyh (p *\<^sub>p (?akk' *\<^sub>p ?xdn +\<^sub>p q) +\<^sub>p r) 0" by simp
- from asp have "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p')) =
- Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
- hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k)*\<^sub>p s) =
- Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs ?p' + Ipoly bs p * Ipoly bs q + Ipoly bs r"
- by (simp add: ring_simps)
- hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
- Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (d - n) shift1 1\<^sub>p *\<^sub>p p)
- + Ipoly bs p * Ipoly bs q + Ipoly bs r"
- by (auto simp only: funpow_shift1_1)
- hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
- Ipoly bs p * (Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (d - n) shift1 1\<^sub>p)
- + Ipoly bs q) + Ipoly bs r" by (simp add: ring_simps)
- hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
- Ipoly bs (p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r)" by simp
- with isnpolyh_unique[OF nakks' nqr']
- have "a ^\<^sub>p (k' - k) *\<^sub>p s =
- p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r" by blast
- hence ?qths using nq'
- apply (rule_tac x="(a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q" in exI)
- apply (rule_tac x="0" in exI) by simp
- with kk' nr dr have "k \<le> k' \<and> (degree r = 0 \<or> degree r < degree p) \<and> (\<exists>nr. isnpolyh r nr) \<and> ?qths"
- by blast } hence ?qrths by blast
- with dom have ?ths by blast}
- moreover
- {assume spz:"s -\<^sub>p ?p' = 0\<^sub>p"
- hence domsp: "?D (a, n, p, k, s -\<^sub>p ?p')"
- apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
- have dom: ?dths apply (rule polydivide_aux_real_domintros)
- using ba ds dn' domsp by simp_all
- from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{ring_char_0,division_by_zero,field}"]
- have " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs ?p'" by simp
- hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)" using np nxdn apply simp
- by (simp only: funpow_shift1_1) simp
- hence sp': "s = ?xdn *\<^sub>p p" using isnpolyh_unique[OF ns polymul_normh[OF nxdn np]] by blast
- {assume h1: "polydivide_aux (a,n,p,k,s) = (k',r)"
- from polydivide_aux.psimps[OF dom] sz dn' ba ds
- have eq: "polydivide_aux (a,n,p,k,s) = polydivide_aux (a,n,p,k, s -\<^sub>p ?p')"
- by (simp add: Let_def)
- also have "\<dots> = (k,0\<^sub>p)" using polydivide_aux.psimps[OF domsp] spz by simp
- finally have "(k',r) = (k,0\<^sub>p)" using h1 by simp
- with sp' ns np nxdn polyadd_0(1)[OF polymul_normh[OF np nxdn]]
- polyadd_0(2)[OF polymul_normh[OF np nxdn]] have ?qrths
- apply auto
- apply (rule exI[where x="?xdn"])
- apply auto
- apply (rule polymul_commute)
- apply simp_all
- done}
- with dom have ?ths by blast}
- ultimately have ?ths by blast }
+ hence headsp': "head s = head ?p'" using ap headp' by simp
+ have nr: "isnpolyh (s -\<^sub>p ?p') 0" using polysub_normh[OF ns np'] by simp
+ from ds degree_polysub_samehead[OF ns np' headsp' degsp']
+ have "degree (s -\<^sub>p ?p') < d \<or> s -\<^sub>p ?p' = 0\<^sub>p" by simp
+ moreover
+ {assume deglt:"degree (s -\<^sub>p ?p') < d"
+ from H[rule_format, OF deglt nr,simplified]
+ have domsp: "?D (a, n, p, k, s -\<^sub>p ?p')" by blast
+ have dom: ?dths apply (rule polydivide_aux_real_domintros)
+ using ba ds dn' domsp by simp_all
+ from polydivide_aux.psimps[OF dom] sz dn' ba ds
+ have eq: "polydivide_aux (a,n,p,k,s) = polydivide_aux (a,n,p,k, s -\<^sub>p ?p')"
+ by (simp add: Let_def)
+ {assume h1: "polydivide_aux (a, n, p, k, s) = (k', r)"
+ from H[rule_format, OF deglt nr, where xa = "k" and xb="k'" and xc="r", simplified]
+ trans[OF eq[symmetric] h1]
+ have kk': "k \<le> k'" and nr:"\<exists>nr. isnpolyh r nr" and dr: "degree r = 0 \<or> degree r < degree p"
+ and q1:"\<exists>q nq. isnpolyh q nq \<and> (a ^\<^sub>p k' - k *\<^sub>p (s -\<^sub>p ?p') = p *\<^sub>p q +\<^sub>p r)" by auto
+ from q1 obtain q n1 where nq: "isnpolyh q n1"
+ and asp:"a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p') = p *\<^sub>p q +\<^sub>p r" by blast
+ from nr obtain nr where nr': "isnpolyh r nr" by blast
+ from polymul_normh[OF nakk' ns] have nakks': "isnpolyh (a ^\<^sub>p (k' - k) *\<^sub>p s) 0" by simp
+ from polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]
+ have nq': "isnpolyh (?akk' *\<^sub>p ?xdn +\<^sub>p q) 0" by simp
+ from polyadd_normh[OF polymul_normh[OF np
+ polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]] nr']
+ have nqr': "isnpolyh (p *\<^sub>p (?akk' *\<^sub>p ?xdn +\<^sub>p q) +\<^sub>p r) 0" by simp
+ from asp have "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p')) =
+ Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
+ hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k)*\<^sub>p s) =
+ Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs ?p' + Ipoly bs p * Ipoly bs q + Ipoly bs r"
+ by (simp add: ring_simps)
+ hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (d - n) shift1 1\<^sub>p *\<^sub>p p)
+ + Ipoly bs p * Ipoly bs q + Ipoly bs r"
+ by (auto simp only: funpow_shift1_1)
+ hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ Ipoly bs p * (Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (d - n) shift1 1\<^sub>p)
+ + Ipoly bs q) + Ipoly bs r" by (simp add: ring_simps)
+ hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ Ipoly bs (p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r)" by simp
+ with isnpolyh_unique[OF nakks' nqr']
+ have "a ^\<^sub>p (k' - k) *\<^sub>p s =
+ p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r" by blast
+ hence ?qths using nq'
+ apply (rule_tac x="(a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q" in exI)
+ apply (rule_tac x="0" in exI) by simp
+ with kk' nr dr have "k \<le> k' \<and> (degree r = 0 \<or> degree r < degree p) \<and> (\<exists>nr. isnpolyh r nr) \<and> ?qths"
+ by blast } hence ?qrths by blast
+ with dom have ?ths by blast}
+ moreover
+ {assume spz:"s -\<^sub>p ?p' = 0\<^sub>p"
+ hence domsp: "?D (a, n, p, k, s -\<^sub>p ?p')"
+ apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
+ have dom: ?dths apply (rule polydivide_aux_real_domintros)
+ using ba ds dn' domsp by simp_all
+ from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{ring_char_0,division_by_zero,field}"]
+ have " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs ?p'" by simp
+ hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)" using np nxdn apply simp
+ by (simp only: funpow_shift1_1) simp
+ hence sp': "s = ?xdn *\<^sub>p p" using isnpolyh_unique[OF ns polymul_normh[OF nxdn np]] by blast
+ {assume h1: "polydivide_aux (a,n,p,k,s) = (k',r)"
+ from polydivide_aux.psimps[OF dom] sz dn' ba ds
+ have eq: "polydivide_aux (a,n,p,k,s) = polydivide_aux (a,n,p,k, s -\<^sub>p ?p')"
+ by (simp add: Let_def)
+ also have "\<dots> = (k,0\<^sub>p)" using polydivide_aux.psimps[OF domsp] spz by simp
+ finally have "(k',r) = (k,0\<^sub>p)" using h1 by simp
+ with sp' ns np nxdn polyadd_0(1)[OF polymul_normh[OF np nxdn]]
+ polyadd_0(2)[OF polymul_normh[OF np nxdn]] have ?qrths
+ apply auto
+ apply (rule exI[where x="?xdn"])
+ apply auto
+ apply (rule polymul_commute)
+ apply simp_all
+ done}
+ with dom have ?ths by blast}
+ ultimately have ?ths by blast }
moreover
{assume ba: "?b \<noteq> a"
- from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns]
- polymul_normh[OF head_isnpolyh[OF ns] np']]
- have nth: "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" by(simp add: min_def)
- have nzths: "a *\<^sub>p s \<noteq> 0\<^sub>p" "?b *\<^sub>p ?p' \<noteq> 0\<^sub>p"
- using polymul_eq0_iff[OF head_isnpolyh[OF np0, simplified ap] ns]
- polymul_eq0_iff[OF head_isnpolyh[OF ns] np']head_nz[OF np0] ap pnz sz head_nz[OF ns]
- funpow_shift1_nz[OF pnz] by simp_all
- from polymul_head_polyeq[OF head_isnpolyh[OF np] ns] head_nz[OF np] sz ap head_head[OF np] pnz
- polymul_head_polyeq[OF head_isnpolyh[OF ns] np'] head_nz [OF ns] sz funpow_shift1_nz[OF pnz, where n="d - n"]
- have hdth: "head (a *\<^sub>p s) = head (?b *\<^sub>p ?p')"
- using head_head[OF ns] funpow_shift1_head[OF np pnz]
- polymul_commute[OF head_isnpolyh[OF np] head_isnpolyh[OF ns]]
- by (simp add: ap)
- from polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"]
- head_nz[OF np] pnz sz ap[symmetric]
- funpow_shift1_nz[OF pnz, where n="d - n"]
- polymul_degreen[OF head_isnpolyh[OF ns] np', where m="0"] head_nz[OF ns]
- ndp ds[symmetric] dn
- have degth: "degree (a *\<^sub>p s) = degree (?b *\<^sub>p ?p') "
- by (simp add: degree_eq_degreen0[symmetric] funpow_shift1_degree)
- {assume dth: "degree ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) < d"
- have th: "?D (a, n, p, Suc k, (a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))"
- using H[rule_format, OF dth nth, simplified] by blast
- have dom: ?dths
- using ba ds dn' th apply simp apply (rule polydivide_aux_real_domintros)
- using ba ds dn' by simp_all
- from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np] ns] polymul_normh[OF head_isnpolyh[OF ns]np']]
- ap have nasbp': "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" by simp
- {assume h1:"polydivide_aux (a,n,p,k,s) = (k', r)"
- from h1 polydivide_aux.psimps[OF dom] sz dn' ba ds
- have eq:"polydivide_aux (a,n,p,Suc k,(a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = (k',r)"
- by (simp add: Let_def)
- with H[rule_format, OF dth nasbp', simplified, where xa="Suc k" and xb="k'" and xc="r"]
- obtain q nq nr where kk': "Suc k \<le> k'" and nr: "isnpolyh r nr" and nq: "isnpolyh q nq"
- and dr: "degree r = 0 \<or> degree r < degree p"
- and qr: "a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = p *\<^sub>p q +\<^sub>p r" by auto
- from kk' have kk'':"Suc (k' - Suc k) = k' - k" by arith
- {fix bs:: "'a::{ring_char_0,division_by_zero,field} list"
-
- from qr isnpolyh_unique[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k", simplified ap] nasbp', symmetric]
- have "Ipoly bs (a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))) = Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
- hence "Ipoly bs a ^ (Suc (k' - Suc k)) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?p' + Ipoly bs r"
- by (simp add: ring_simps power_Suc)
- hence "Ipoly bs a ^ (k' - k) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn * Ipoly bs p + Ipoly bs r"
- by (simp add:kk'' funpow_shift1_1[where n="d - n" and p="p"])
- hence "Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = Ipoly bs p * (Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn) + Ipoly bs r"
- by (simp add: ring_simps)}
- hence ieq:"\<forall>(bs :: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
- Ipoly bs (p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r)" by auto
- let ?q = "q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)"
- from polyadd_normh[OF nq polymul_normh[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k"] head_isnpolyh[OF ns], simplified ap ] nxdn]]
- have nqw: "isnpolyh ?q 0" by simp
- from ieq isnpolyh_unique[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - k"] ns, simplified ap] polyadd_normh[OF polymul_normh[OF np nqw] nr]]
- have asth: "(a ^\<^sub>p (k' - k) *\<^sub>p s) = p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r" by blast
- from dr kk' nr h1 asth nqw have ?qrths apply simp
- apply (rule conjI)
- apply (rule exI[where x="nr"], simp)
- apply (rule exI[where x="(q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn))"], simp)
- apply (rule exI[where x="0"], simp)
- done}
- hence ?qrths by blast
- with dom have ?ths by blast}
- moreover
- {assume spz: "a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p') = 0\<^sub>p"
- hence domsp: "?D (a, n, p, Suc k, a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p'))"
- apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
- have dom: ?dths using sz ba dn' ds domsp
- by - (rule polydivide_aux_real_domintros, simp_all)
- {fix bs :: "'a::{ring_char_0,division_by_zero,field} list"
- from isnpolyh_unique[OF nth, where ?'a="'a" and q="0\<^sub>p",simplified,symmetric] spz
- have "Ipoly bs (a*\<^sub>p s) = Ipoly bs ?b * Ipoly bs ?p'" by simp
- hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (?b *\<^sub>p ?xdn) * Ipoly bs p"
- by (simp add: funpow_shift1_1[where n="d - n" and p="p"])
- hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" by simp
- }
- hence hth: "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" ..
- from hth
- have asq: "a *\<^sub>p s = p *\<^sub>p (?b *\<^sub>p ?xdn)"
- using isnpolyh_unique[where ?'a = "'a::{ring_char_0,division_by_zero,field}", OF polymul_normh[OF head_isnpolyh[OF np] ns]
+ from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns]
+ polymul_normh[OF head_isnpolyh[OF ns] np']]
+ have nth: "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" by(simp add: min_def)
+ have nzths: "a *\<^sub>p s \<noteq> 0\<^sub>p" "?b *\<^sub>p ?p' \<noteq> 0\<^sub>p"
+ using polymul_eq0_iff[OF head_isnpolyh[OF np0, simplified ap] ns]
+ polymul_eq0_iff[OF head_isnpolyh[OF ns] np']head_nz[OF np0] ap pnz sz head_nz[OF ns]
+ funpow_shift1_nz[OF pnz] by simp_all
+ from polymul_head_polyeq[OF head_isnpolyh[OF np] ns] head_nz[OF np] sz ap head_head[OF np] pnz
+ polymul_head_polyeq[OF head_isnpolyh[OF ns] np'] head_nz [OF ns] sz funpow_shift1_nz[OF pnz, where n="d - n"]
+ have hdth: "head (a *\<^sub>p s) = head (?b *\<^sub>p ?p')"
+ using head_head[OF ns] funpow_shift1_head[OF np pnz]
+ polymul_commute[OF head_isnpolyh[OF np] head_isnpolyh[OF ns]]
+ by (simp add: ap)
+ from polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"]
+ head_nz[OF np] pnz sz ap[symmetric]
+ funpow_shift1_nz[OF pnz, where n="d - n"]
+ polymul_degreen[OF head_isnpolyh[OF ns] np', where m="0"] head_nz[OF ns]
+ ndp ds[symmetric] dn
+ have degth: "degree (a *\<^sub>p s) = degree (?b *\<^sub>p ?p') "
+ by (simp add: degree_eq_degreen0[symmetric] funpow_shift1_degree)
+ {assume dth: "degree ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) < d"
+ have th: "?D (a, n, p, Suc k, (a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))"
+ using H[rule_format, OF dth nth, simplified] by blast
+ have dom: ?dths
+ using ba ds dn' th apply simp apply (rule polydivide_aux_real_domintros)
+ using ba ds dn' by simp_all
+ from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np] ns] polymul_normh[OF head_isnpolyh[OF ns]np']]
+ ap have nasbp': "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" by simp
+ {assume h1:"polydivide_aux (a,n,p,k,s) = (k', r)"
+ from h1 polydivide_aux.psimps[OF dom] sz dn' ba ds
+ have eq:"polydivide_aux (a,n,p,Suc k,(a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = (k',r)"
+ by (simp add: Let_def)
+ with H[rule_format, OF dth nasbp', simplified, where xa="Suc k" and xb="k'" and xc="r"]
+ obtain q nq nr where kk': "Suc k \<le> k'" and nr: "isnpolyh r nr" and nq: "isnpolyh q nq"
+ and dr: "degree r = 0 \<or> degree r < degree p"
+ and qr: "a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = p *\<^sub>p q +\<^sub>p r" by auto
+ from kk' have kk'':"Suc (k' - Suc k) = k' - k" by arith
+ {fix bs:: "'a::{ring_char_0,division_by_zero,field} list"
+
+ from qr isnpolyh_unique[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k", simplified ap] nasbp', symmetric]
+ have "Ipoly bs (a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))) = Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
+ hence "Ipoly bs a ^ (Suc (k' - Suc k)) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?p' + Ipoly bs r"
+ by (simp add: ring_simps power_Suc)
+ hence "Ipoly bs a ^ (k' - k) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn * Ipoly bs p + Ipoly bs r"
+ by (simp add:kk'' funpow_shift1_1[where n="d - n" and p="p"])
+ hence "Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = Ipoly bs p * (Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn) + Ipoly bs r"
+ by (simp add: ring_simps)}
+ hence ieq:"\<forall>(bs :: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ Ipoly bs (p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r)" by auto
+ let ?q = "q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)"
+ from polyadd_normh[OF nq polymul_normh[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k"] head_isnpolyh[OF ns], simplified ap ] nxdn]]
+ have nqw: "isnpolyh ?q 0" by simp
+ from ieq isnpolyh_unique[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - k"] ns, simplified ap] polyadd_normh[OF polymul_normh[OF np nqw] nr]]
+ have asth: "(a ^\<^sub>p (k' - k) *\<^sub>p s) = p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r" by blast
+ from dr kk' nr h1 asth nqw have ?qrths apply simp
+ apply (rule conjI)
+ apply (rule exI[where x="nr"], simp)
+ apply (rule exI[where x="(q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn))"], simp)
+ apply (rule exI[where x="0"], simp)
+ done}
+ hence ?qrths by blast
+ with dom have ?ths by blast}
+ moreover
+ {assume spz: "a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p') = 0\<^sub>p"
+ hence domsp: "?D (a, n, p, Suc k, a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p'))"
+ apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
+ have dom: ?dths using sz ba dn' ds domsp
+ by - (rule polydivide_aux_real_domintros, simp_all)
+ {fix bs :: "'a::{ring_char_0,division_by_zero,field} list"
+ from isnpolyh_unique[OF nth, where ?'a="'a" and q="0\<^sub>p",simplified,symmetric] spz
+ have "Ipoly bs (a*\<^sub>p s) = Ipoly bs ?b * Ipoly bs ?p'" by simp
+ hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (?b *\<^sub>p ?xdn) * Ipoly bs p"
+ by (simp add: funpow_shift1_1[where n="d - n" and p="p"])
+ hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" by simp
+ }
+ hence hth: "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" ..
+ from hth
+ have asq: "a *\<^sub>p s = p *\<^sub>p (?b *\<^sub>p ?xdn)"
+ using isnpolyh_unique[where ?'a = "'a::{ring_char_0,division_by_zero,field}", OF polymul_normh[OF head_isnpolyh[OF np] ns]
polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]],
- simplified ap] by simp
- {assume h1: "polydivide_aux (a,n,p,k,s) = (k', r)"
- from h1 sz ds ba dn' spz polydivide_aux.psimps[OF dom] polydivide_aux.psimps[OF domsp]
- have "(k',r) = (Suc k, 0\<^sub>p)" by (simp add: Let_def)
- with h1 np head_isnpolyh[OF np, simplified ap] ns polymul_normh[OF head_isnpolyh[OF ns] nxdn]
- polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]] asq
- have ?qrths apply (clarsimp simp add: Let_def)
- apply (rule exI[where x="?b *\<^sub>p ?xdn"]) apply simp
- apply (rule exI[where x="0"], simp)
- done}
- hence ?qrths by blast
- with dom have ?ths by blast}
- ultimately have ?ths using degree_polysub_samehead[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns] polymul_normh[OF head_isnpolyh[OF ns] np'] hdth degth] polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"]
- head_nz[OF np] pnz sz ap[symmetric] ds[symmetric]
- by (simp add: degree_eq_degreen0[symmetric]) blast }
+ simplified ap] by simp
+ {assume h1: "polydivide_aux (a,n,p,k,s) = (k', r)"
+ from h1 sz ds ba dn' spz polydivide_aux.psimps[OF dom] polydivide_aux.psimps[OF domsp]
+ have "(k',r) = (Suc k, 0\<^sub>p)" by (simp add: Let_def)
+ with h1 np head_isnpolyh[OF np, simplified ap] ns polymul_normh[OF head_isnpolyh[OF ns] nxdn]
+ polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]] asq
+ have ?qrths apply (clarsimp simp add: Let_def)
+ apply (rule exI[where x="?b *\<^sub>p ?xdn"]) apply simp
+ apply (rule exI[where x="0"], simp)
+ done}
+ hence ?qrths by blast
+ with dom have ?ths by blast}
+ ultimately have ?ths using degree_polysub_samehead[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns] polymul_normh[OF head_isnpolyh[OF ns] np'] hdth degth] polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"]
+ head_nz[OF np] pnz sz ap[symmetric] ds[symmetric]
+ by (simp add: degree_eq_degreen0[symmetric]) blast }
ultimately have ?ths by blast
}
ultimately have ?ths by blast}
@@ -1732,7 +1732,7 @@
recdef isweaknpoly "measure size"
"isweaknpoly (C c) = True"
"isweaknpoly (CN c n p) \<longleftrightarrow> isweaknpoly c \<and> isweaknpoly p"
- "isweaknpoly p = False"
+ "isweaknpoly p = False"
lemma isnpolyh_isweaknpoly: "isnpolyh p n0 \<Longrightarrow> isweaknpoly p"
by (induct p arbitrary: n0, auto)
--- a/src/HOL/Import/shuffler.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Import/shuffler.ML Wed Oct 28 11:43:06 2009 +0000
@@ -248,16 +248,16 @@
val tvars = OldTerm.term_tvars t
val tfree_names = OldTerm.add_term_tfree_names(t,[])
val (type_inst,_) =
- Library.foldl (fn ((inst,used),(w as (v,_),S)) =>
+ fold (fn (w as (v,_), S) => fn (inst, used) =>
let
val v' = Name.variant used v
in
((w,TFree(v',S))::inst,v'::used)
end)
- (([],tfree_names),tvars)
+ tvars ([], tfree_names)
val t' = subst_TVars type_inst t
in
- (t',map (fn (w,TFree(v,S)) => (v,TVar(w,S))
+ (t', map (fn (w,TFree(v,S)) => (v,TVar(w,S))
| _ => error "Internal error in Shuffler.freeze_thaw") type_inst)
end
--- a/src/HOL/IsaMakefile Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/IsaMakefile Wed Oct 28 11:43:06 2009 +0000
@@ -256,6 +256,7 @@
Map.thy \
Nat_Numeral.thy \
Presburger.thy \
+ Predicate_Compile.thy \
Quickcheck.thy \
Random.thy \
Recdef.thy \
@@ -284,6 +285,13 @@
Tools/numeral_simprocs.ML \
Tools/numeral_syntax.ML \
Tools/polyhash.ML \
+ Tools/Predicate_Compile/predicate_compile_aux.ML \
+ Tools/Predicate_Compile/predicate_compile_core.ML \
+ Tools/Predicate_Compile/predicate_compile_data.ML \
+ Tools/Predicate_Compile/predicate_compile_fun.ML \
+ Tools/Predicate_Compile/predicate_compile.ML \
+ Tools/Predicate_Compile/predicate_compile_pred.ML \
+ Tools/Predicate_Compile/predicate_compile_set.ML \
Tools/quickcheck_generators.ML \
Tools/Qelim/cooper_data.ML \
Tools/Qelim/cooper.ML \
@@ -947,7 +955,7 @@
ex/Intuitionistic.thy ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy \
ex/MergeSort.thy ex/Meson_Test.thy ex/MonoidGroup.thy \
ex/Multiquote.thy ex/NatSum.thy ex/Numeral.thy ex/PER.thy \
- ex/Predicate_Compile.thy ex/Predicate_Compile_ex.thy \
+ ex/Predicate_Compile_ex.thy ex/Predicate_Compile_Quickcheck.thy \
ex/PresburgerEx.thy ex/Primrec.thy ex/Quickcheck_Examples.thy \
ex/ROOT.ML ex/Recdefs.thy ex/Records.thy ex/ReflectionEx.thy \
ex/Refute_Examples.thy ex/SAT_Examples.thy ex/SVC_Oracle.thy \
@@ -1201,12 +1209,11 @@
$(OUT)/HOL-SMT: $(OUT)/HOL-Word SMT/SMT_Base.thy SMT/Z3.thy \
SMT/SMT.thy SMT/Tools/smt_normalize.ML SMT/Tools/smt_monomorph.ML \
- SMT/Tools/smt_translate.ML SMT/Tools/smt_builtin.ML \
- SMT/Tools/smtlib_interface.ML SMT/Tools/smt_solver.ML \
- SMT/Tools/cvc3_solver.ML SMT/Tools/yices_solver.ML \
- SMT/Tools/z3_proof_terms.ML SMT/Tools/z3_proof_rules.ML \
- SMT/Tools/z3_proof.ML SMT/Tools/z3_model.ML \
- SMT/Tools/z3_interface.ML SMT/Tools/z3_solver.ML
+ SMT/Tools/smt_translate.ML SMT/Tools/smtlib_interface.ML \
+ SMT/Tools/smt_solver.ML SMT/Tools/cvc3_solver.ML \
+ SMT/Tools/yices_solver.ML SMT/Tools/z3_proof_terms.ML \
+ SMT/Tools/z3_proof_rules.ML SMT/Tools/z3_proof.ML \
+ SMT/Tools/z3_model.ML SMT/Tools/z3_interface.ML SMT/Tools/z3_solver.ML
@cd SMT; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Word HOL-SMT
--- a/src/HOL/Library/Code_Char.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Library/Code_Char.thy Wed Oct 28 11:43:06 2009 +0000
@@ -35,4 +35,28 @@
code_const "Code_Evaluation.term_of \<Colon> char \<Rightarrow> term"
(Eval "HOLogic.mk'_char/ (IntInf.fromInt/ (Char.ord/ _))")
+
+definition implode :: "string \<Rightarrow> String.literal" where
+ "implode = STR"
+
+primrec explode :: "String.literal \<Rightarrow> string" where
+ "explode (STR s) = s"
+
+lemma [code]:
+ "literal_case f s = f (explode s)"
+ "literal_rec f s = f (explode s)"
+ by (cases s, simp)+
+
+code_reserved SML String
+
+code_const implode
+ (SML "String.implode")
+ (OCaml "failwith/ \"implode\"")
+ (Haskell "_")
+
+code_const explode
+ (SML "String.explode")
+ (OCaml "failwith/ \"explode\"")
+ (Haskell "_")
+
end
--- a/src/HOL/Main.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Main.thy Wed Oct 28 11:43:06 2009 +0000
@@ -1,7 +1,7 @@
header {* Main HOL *}
theory Main
-imports Plain Nitpick Quickcheck Recdef
+imports Plain Nitpick Predicate_Compile Recdef
begin
text {*
--- a/src/HOL/Matrix/cplex/CplexMatrixConverter.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Matrix/cplex/CplexMatrixConverter.ML Wed Oct 28 11:43:06 2009 +0000
@@ -75,7 +75,7 @@
set_elem vec (s2i v) (if positive then num else "-"^num)
| setprod _ _ _ = raise (Converter "term is not a normed product")
- fun sum2vec (cplexSum ts) = Library.foldl (fn (vec, t) => setprod vec true t) (empty_vector, ts)
+ fun sum2vec (cplexSum ts) = fold (fn t => fn vec => setprod vec true t) ts empty_vector
| sum2vec t = setprod empty_vector true t
fun constrs2Ab j A b [] = (A, b)
@@ -100,9 +100,9 @@
fun convert_results (cplex.Optimal (opt, entries)) name2index =
let
- fun setv (v, (name, value)) = (matrix_builder.set_elem v (name2index name) value)
+ fun setv (name, value) v = matrix_builder.set_elem v (name2index name) value
in
- (opt, Library.foldl setv (matrix_builder.empty_vector, entries))
+ (opt, fold setv entries (matrix_builder.empty_vector))
end
| convert_results _ _ = raise (Converter "No optimal result")
--- a/src/HOL/Mirabelle/Tools/mirabelle.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle.ML Wed Oct 28 11:43:06 2009 +0000
@@ -28,8 +28,8 @@
val goal_thm_of : Proof.state -> thm
val can_apply : Time.time -> (Proof.context -> int -> tactic) ->
Proof.state -> bool
- val theorems_in_proof_term : Thm.thm -> Thm.thm list
- val theorems_of_sucessful_proof : Toplevel.state option -> Thm.thm list
+ val theorems_in_proof_term : thm -> thm list
+ val theorems_of_sucessful_proof : Toplevel.state option -> thm list
val get_setting : (string * string) list -> string * string -> string
val get_int_setting : (string * string) list -> string * int -> int
val cpu_time : ('a -> 'b) -> 'a -> 'b * int
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Wed Oct 28 11:43:06 2009 +0000
@@ -302,8 +302,11 @@
fun run_sh prover hard_timeout timeout dir st =
let
val (ctxt, goal) = Proof.get_goal st
- val ctxt' = if is_none dir then ctxt
- else Config.put ATP_Wrapper.destdir (the dir) ctxt
+ val change_dir = (fn SOME d => Config.put ATP_Wrapper.destdir d | _ => I)
+ val ctxt' =
+ ctxt
+ |> change_dir dir
+ |> Config.put ATP_Wrapper.measure_runtime true
val problem = ATP_Wrapper.problem_of_goal (! ATP_Manager.full_types) 1 (ctxt', goal);
val time_limit =
--- a/src/HOL/Mirabelle/doc/options.txt Wed Oct 28 11:42:31 2009 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,11 +0,0 @@
-Options for sledgehammer:
-
- * prover=NAME: name of the external prover to call
- * prover_timeout=TIME: timeout for invoked ATP (seconds of process time)
- * prover_hard_timeout=TIME: timeout for invoked ATP (seconds of elapsed time)
- * keep=PATH: path where to keep temporary files created by sledgehammer
- * full_types: enable full-typed encoding
- * minimize: enable minimization of theorem set found by sledgehammer
- * minimize_timeout=TIME: timeout for each minimization step (seconds of
- process time)
- * metis: apply metis with the theorems found by sledgehammer
--- a/src/HOL/Mirabelle/lib/Tools/mirabelle Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Mirabelle/lib/Tools/mirabelle Wed Oct 28 11:43:06 2009 +0000
@@ -35,8 +35,17 @@
echo " either NAME or NAME[OPTION,...,OPTION]. Available actions are:"
print_action_names
echo
- echo " A list of available OPTIONs can be found here:"
- echo " $MIRABELLE_HOME/doc/options.txt"
+ echo " Available OPTIONs for the ACTION sledgehammer:"
+ echo " * prover=NAME: name of the external prover to call"
+ echo " * prover_timeout=TIME: timeout for invoked ATP (seconds of process time)"
+ echo " * prover_hard_timeout=TIME: timeout for invoked ATP (seconds of elapsed"
+ echo " time)"
+ echo " * keep=PATH: path where to keep temporary files created by sledgehammer"
+ echo " * full_types: enable fully-typed encoding"
+ echo " * minimize: enable minimization of theorem set found by sledgehammer"
+ echo " * minimize_timeout=TIME: timeout for each minimization step (seconds of"
+ echo " process time)"
+ echo " * metis: apply metis to the theorems found by sledgehammer"
echo
echo " FILES is a list of theory files, where each file is either NAME.thy"
echo " or NAME.thy[START:END] and START and END are numbers indicating the"
--- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Wed Oct 28 11:43:06 2009 +0000
@@ -1,6 +1,5 @@
-theory Multivariate_Analysis imports
- Convex_Euclidean_Space
- Determinants
+theory Multivariate_Analysis
+imports Convex_Euclidean_Space Determinants
begin
end
--- a/src/HOL/Multivariate_Analysis/ROOT.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Multivariate_Analysis/ROOT.ML Wed Oct 28 11:43:06 2009 +0000
@@ -1,6 +1,1 @@
-(*
- no_document use_thy "ThisTheory";
- use_thy "ThatTheory";
-*)
-
use_thy "Multivariate_Analysis";
--- a/src/HOL/Nitpick.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Nitpick.thy Wed Oct 28 11:43:06 2009 +0000
@@ -28,7 +28,6 @@
typedecl bisim_iterator
-(* FIXME: use axiomatization (here and elsewhere) *)
axiomatization unknown :: 'a
and undefined_fast_The :: 'a
and undefined_fast_Eps :: 'a
--- a/src/HOL/Nitpick_Examples/Mono_Nits.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy Wed Oct 28 11:43:06 2009 +0000
@@ -14,8 +14,8 @@
ML {*
exception FAIL
-val defs = NitpickHOL.all_axioms_of @{theory} |> #1
-val def_table = NitpickHOL.const_def_table @{context} defs
+val defs = Nitpick_HOL.all_axioms_of @{theory} |> #1
+val def_table = Nitpick_HOL.const_def_table @{context} defs
val ext_ctxt =
{thy = @{theory}, ctxt = @{context}, max_bisim_depth = ~1, boxes = [],
user_axioms = NONE, debug = false, wfs = [], destroy_constrs = false,
--- a/src/HOL/Nitpick_Examples/Tests_Nits.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Nitpick_Examples/Tests_Nits.thy Wed Oct 28 11:43:06 2009 +0000
@@ -11,6 +11,6 @@
imports Main
begin
-ML {* NitpickTests.run_all_tests () *}
+ML {* Nitpick_Tests.run_all_tests () *}
end
--- a/src/HOL/Nominal/Examples/Pattern.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Nominal/Examples/Pattern.thy Wed Oct 28 11:43:06 2009 +0000
@@ -669,17 +669,17 @@
proof
assume "x = x' \<and> t = u"
with PVar PVar' have "PVar x T = ([]::name prm) \<bullet> q \<and>
- t = ([]::name prm) \<bullet> u \<and>
- set ([]::name prm) \<subseteq> (supp (PVar x T) \<union> supp q) \<times>
+ t = ([]::name prm) \<bullet> u \<and>
+ set ([]::name prm) \<subseteq> (supp (PVar x T) \<union> supp q) \<times>
(supp (PVar x T) \<union> supp q)" by simp
then show ?thesis ..
next
assume "x \<noteq> x' \<and> t = [(x, x')] \<bullet> u \<and> x \<sharp> u"
with PVar PVar' have "PVar x T = [(x, x')] \<bullet> q \<and>
- t = [(x, x')] \<bullet> u \<and>
- set [(x, x')] \<subseteq> (supp (PVar x T) \<union> supp q) \<times>
+ t = [(x, x')] \<bullet> u \<and>
+ set [(x, x')] \<subseteq> (supp (PVar x T) \<union> supp q) \<times>
(supp (PVar x T) \<union> supp q)"
- by (simp add: perm_swap swap_simps supp_atm perm_type)
+ by (simp add: perm_swap swap_simps supp_atm perm_type)
then show ?thesis ..
qed
next
--- a/src/HOL/Nominal/nominal_datatype.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Nominal/nominal_datatype.ML Wed Oct 28 11:43:06 2009 +0000
@@ -202,15 +202,15 @@
val atoms = atoms_of thy;
- fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
- let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs)
+ fun prep_constr (cname, cargs, mx) (constrs, sorts) =
+ let val (cargs', sorts') = fold_map (prep_typ tmp_thy) cargs sorts
in (constrs @ [(cname, cargs', mx)], sorts') end
- fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
- let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
+ fun prep_dt_spec (tvs, tname, mx, constrs) (dts, sorts) =
+ let val (constrs', sorts') = fold prep_constr constrs ([], sorts)
in (dts @ [(tvs, tname, mx, constrs')], sorts') end
- val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
+ val (dts', sorts) = fold prep_dt_spec dts ([], []);
val tyvars = map (map (fn s =>
(s, the (AList.lookup (op =) sorts s))) o #1) dts';
@@ -770,8 +770,8 @@
val full_new_type_names = map (Sign.full_bname thy) new_type_names;
- fun make_constr_def tname T T' ((thy, defs, eqns),
- (((cname_rep, _), (cname, cargs)), (cname', mx))) =
+ fun make_constr_def tname T T' (((cname_rep, _), (cname, cargs)), (cname', mx))
+ (thy, defs, eqns) =
let
fun constr_arg ((dts, dt), (j, l_args, r_args)) =
let
@@ -805,22 +805,24 @@
(PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]
in (thy', defs @ [def_thm], eqns @ [eqn]) end;
- fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)),
- (_, (_, _, constrs'))), tname), T), T'), constr_syntax)) =
+ fun dt_constr_defs ((((((_, (_, _, constrs)),
+ (_, (_, _, constrs'))), tname), T), T'), constr_syntax) (thy, defs, eqns, dist_lemmas) =
let
val rep_const = cterm_of thy
(Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
- val dist = Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
- val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T')
- ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax)
+ val dist =
+ Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
+ val (thy', defs', eqns') = fold (make_constr_def tname T T')
+ (constrs ~~ constrs' ~~ constr_syntax) (Sign.add_path tname thy, defs, [])
in
(Sign.parent_path thy', defs', eqns @ [eqns'], dist_lemmas @ [dist])
end;
- val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs
- ((thy7, [], [], []), List.take (descr, length new_type_names) ~~
+ val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = fold dt_constr_defs
+ (List.take (descr, length new_type_names) ~~
List.take (pdescr, length new_type_names) ~~
- new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax);
+ new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax)
+ (thy7, [], [], []);
val abs_inject_thms = map (collect_simp o #Abs_inject o fst) typedefs
val rep_inject_thms = map (#Rep_inject o fst) typedefs
@@ -1031,7 +1033,7 @@
(**** weak induction theorem ****)
- fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) =
+ fun mk_indrule_lemma (((i, _), T), U) (prems, concls) =
let
val Rep_t = Const (List.nth (rep_names, i), T --> U) $
mk_Free "x" T i;
@@ -1045,7 +1047,7 @@
end;
val (indrule_lemma_prems, indrule_lemma_concls) =
- Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs'));
+ fold mk_indrule_lemma (descr'' ~~ recTs ~~ recTs') ([], []);
val indrule_lemma = Goal.prove_global thy8 [] []
(Logic.mk_implies
@@ -1455,8 +1457,8 @@
(* FIXME: avoid collisions with other variable names? *)
val rec_ctxt = Free ("z", fsT');
- fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems',
- rec_eq_prems, l), ((cname, cargs), idxs)) =
+ fun make_rec_intr T p rec_set ((cname, cargs), idxs)
+ (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, l) =
let
val Ts = map (typ_of_dtyp descr'' sorts) cargs;
val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
@@ -1500,9 +1502,10 @@
end;
val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) =
- Library.foldl (fn (x, ((((d, d'), T), p), rec_set)) =>
- Library.foldl (make_rec_intr T p rec_set) (x, #3 (snd d) ~~ snd d'))
- (([], [], [], [], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets');
+ fold (fn ((((d, d'), T), p), rec_set) =>
+ fold (make_rec_intr T p rec_set) (#3 (snd d) ~~ snd d'))
+ (descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets')
+ ([], [], [], [], 0);
val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) =
thy10 |>
--- a/src/HOL/Nominal/nominal_permeq.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Nominal/nominal_permeq.ML Wed Oct 28 11:43:06 2009 +0000
@@ -301,9 +301,9 @@
val ps = Logic.strip_params (term_of goal);
val Ts = rev (map snd ps);
val vs = collect_vars 0 x [];
- val s = Library.foldr (fn (v, s) =>
+ val s = fold_rev (fn v => fn s =>
HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s)
- (vs, HOLogic.unit);
+ vs HOLogic.unit;
val s' = list_abs (ps,
Const ("Nominal.supp", fastype_of1 (Ts, s) -->
snd (split_last (binder_types T)) --> HOLogic.boolT) $ s);
@@ -343,9 +343,9 @@
val ps = Logic.strip_params (term_of goal);
val Ts = rev (map snd ps);
val vs = collect_vars 0 t [];
- val s = Library.foldr (fn (v, s) =>
+ val s = fold_rev (fn v => fn s =>
HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s)
- (vs, HOLogic.unit);
+ vs HOLogic.unit;
val s' = list_abs (ps,
Const ("Nominal.supp", fastype_of1 (Ts, s) --> (HOLogic.mk_setT T)) $ s);
val supports_fresh_rule' = Thm.lift_rule goal supports_fresh_rule;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile.thy Wed Oct 28 11:43:06 2009 +0000
@@ -0,0 +1,21 @@
+(* Title: HOL/Predicate_Compile.thy
+ Author: Stefan Berghofer, Lukas Bulwahn, Florian Haftmann, TU Muenchen
+*)
+
+header {* A compiler for predicates defined by introduction rules *}
+
+theory Predicate_Compile
+imports Quickcheck
+uses
+ "Tools/Predicate_Compile/predicate_compile_aux.ML"
+ "Tools/Predicate_Compile/predicate_compile_core.ML"
+ "Tools/Predicate_Compile/predicate_compile_set.ML"
+ "Tools/Predicate_Compile/predicate_compile_data.ML"
+ "Tools/Predicate_Compile/predicate_compile_fun.ML"
+ "Tools/Predicate_Compile/predicate_compile_pred.ML"
+ "Tools/Predicate_Compile/predicate_compile.ML"
+begin
+
+setup Predicate_Compile.setup
+
+end
\ No newline at end of file
--- a/src/HOL/Quickcheck.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Quickcheck.thy Wed Oct 28 11:43:06 2009 +0000
@@ -126,6 +126,47 @@
shows "random_aux k = rhs k"
using assms by (rule code_numeral.induct)
+subsection {* the Random-Predicate Monad *}
+
+types 'a randompred = "Random.seed \<Rightarrow> ('a Predicate.pred \<times> Random.seed)"
+
+definition empty :: "'a randompred"
+ where "empty = Pair (bot_class.bot)"
+
+definition single :: "'a => 'a randompred"
+ where "single x = Pair (Predicate.single x)"
+
+definition bind :: "'a randompred \<Rightarrow> ('a \<Rightarrow> 'b randompred) \<Rightarrow> 'b randompred"
+ where
+ "bind R f = (\<lambda>s. let
+ (P, s') = R s;
+ (s1, s2) = Random.split_seed s'
+ in (Predicate.bind P (%a. fst (f a s1)), s2))"
+
+definition union :: "'a randompred \<Rightarrow> 'a randompred \<Rightarrow> 'a randompred"
+where
+ "union R1 R2 = (\<lambda>s. let
+ (P1, s') = R1 s; (P2, s'') = R2 s'
+ in (upper_semilattice_class.sup P1 P2, s''))"
+
+definition if_randompred :: "bool \<Rightarrow> unit randompred"
+where
+ "if_randompred b = (if b then single () else empty)"
+
+definition not_randompred :: "unit randompred \<Rightarrow> unit randompred"
+where
+ "not_randompred P = (\<lambda>s. let
+ (P', s') = P s
+ in if Predicate.eval P' () then (Orderings.bot, s') else (Predicate.single (), s'))"
+
+definition Random :: "(Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed) \<Rightarrow> 'a randompred"
+ where "Random g = scomp g (Pair o (Predicate.single o fst))"
+
+definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a randompred \<Rightarrow> 'b randompred)"
+ where "map f P = bind P (single o f)"
+
+subsection {* Code setup *}
+
use "Tools/quickcheck_generators.ML"
setup {* Quickcheck_Generators.setup *}
@@ -136,8 +177,10 @@
code_reserved Quickcheck Quickcheck_Generators
-
+hide (open) fact empty_def single_def bind_def union_def if_randompred_def not_randompred_def Random_def map_def
+hide (open) type randompred
hide (open) const random collapse beyond random_fun_aux random_fun_lift
+ empty single bind union if_randompred not_randompred Random map
no_notation fcomp (infixl "o>" 60)
no_notation scomp (infixl "o\<rightarrow>" 60)
--- a/src/HOL/Random.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Random.thy Wed Oct 28 11:43:06 2009 +0000
@@ -12,15 +12,15 @@
subsection {* Auxiliary functions *}
+fun log :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
+ "log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
+
definition inc_shift :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
"inc_shift v k = (if v = k then 1 else k + 1)"
definition minus_shift :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
"minus_shift r k l = (if k < l then r + k - l else k - l)"
-fun log :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
- "log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
-
subsection {* Random seeds *}
@@ -29,28 +29,19 @@
primrec "next" :: "seed \<Rightarrow> code_numeral \<times> seed" where
"next (v, w) = (let
k = v div 53668;
- v' = minus_shift 2147483563 (40014 * (v mod 53668)) (k * 12211);
+ v' = minus_shift 2147483563 ((v mod 53668) * 40014) (k * 12211);
l = w div 52774;
- w' = minus_shift 2147483399 (40692 * (w mod 52774)) (l * 3791);
+ w' = minus_shift 2147483399 ((w mod 52774) * 40692) (l * 3791);
z = minus_shift 2147483562 v' (w' + 1) + 1
in (z, (v', w')))"
-lemma next_not_0:
- "fst (next s) \<noteq> 0"
- by (cases s) (auto simp add: minus_shift_def Let_def)
-
-primrec seed_invariant :: "seed \<Rightarrow> bool" where
- "seed_invariant (v, w) \<longleftrightarrow> 0 < v \<and> v < 9438322952 \<and> 0 < w \<and> True"
-
definition split_seed :: "seed \<Rightarrow> seed \<times> seed" where
"split_seed s = (let
(v, w) = s;
(v', w') = snd (next s);
v'' = inc_shift 2147483562 v;
- s'' = (v'', w');
- w'' = inc_shift 2147483398 w;
- s''' = (v', w'')
- in (s'', s'''))"
+ w'' = inc_shift 2147483398 w
+ in ((v'', w'), (v', w'')))"
subsection {* Base selectors *}
@@ -175,7 +166,7 @@
*}
hide (open) type seed
-hide (open) const inc_shift minus_shift log "next" seed_invariant split_seed
+hide (open) const inc_shift minus_shift log "next" split_seed
iterate range select pick select_weight
no_notation fcomp (infixl "o>" 60)
--- a/src/HOL/SMT/SMT_Base.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/SMT/SMT_Base.thy Wed Oct 28 11:43:06 2009 +0000
@@ -29,7 +29,7 @@
definition pat :: "'a \<Rightarrow> pattern"
where "pat _ = Pattern"
-definition nopat :: "bool \<Rightarrow> pattern"
+definition nopat :: "'a \<Rightarrow> pattern"
where "nopat _ = Pattern"
definition andpat :: "pattern \<Rightarrow> 'a \<Rightarrow> pattern" (infixl "andpat" 60)
--- a/src/HOL/SMT/Tools/smt_builtin.ML Wed Oct 28 11:42:31 2009 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,78 +0,0 @@
-(* Title: HOL/SMT/Tools/smt_builtin.ML
- Author: Sascha Boehme, TU Muenchen
-
-Tables for built-in symbols.
-*)
-
-signature SMT_BUILTIN =
-sig
- type sterm = (SMT_Translate.sym, typ) sterm
- type builtin_fun = typ -> sterm list -> (string * sterm list) option
- type table = (typ * builtin_fun) list Symtab.table
-
- val make: (term * string) list -> table
- val add: term * builtin_fun -> table -> table
- val lookup: table -> theory -> string * typ -> sterm list ->
- (string * sterm list) option
-
- val bv_rotate: (int -> string) -> builtin_fun
- val bv_extend: (int -> string) -> builtin_fun
- val bv_extract: (int -> int -> string) -> builtin_fun
-end
-
-structure SMT_Builtin: SMT_BUILTIN =
-struct
-
-structure T = SMT_Translate
-
-type sterm = (SMT_Translate.sym, typ) sterm
-type builtin_fun = typ -> sterm list -> (string * sterm list) option
-type table = (typ * builtin_fun) list Symtab.table
-
-fun make entries =
- let
- fun dest (t, bn) =
- let val (n, T) = Term.dest_Const t
- in (n, (Logic.varifyT T, K (pair bn))) end
- in Symtab.make (AList.group (op =) (map dest entries)) end
-
-fun add (t, f) tab =
- let val (n, T) = apsnd Logic.varifyT (Term.dest_Const t)
- in Symtab.map_default (n, []) (AList.update (op =) (T, f)) tab end
-
-fun lookup tab thy (n, T) =
- AList.lookup (Sign.typ_instance thy) (Symtab.lookup_list tab n) T T
-
-
-fun dest_binT T =
- (case T of
- Type (@{type_name "Numeral_Type.num0"}, _) => 0
- | Type (@{type_name "Numeral_Type.num1"}, _) => 1
- | Type (@{type_name "Numeral_Type.bit0"}, [T]) => 2 * dest_binT T
- | Type (@{type_name "Numeral_Type.bit1"}, [T]) => 1 + 2 * dest_binT T
- | _ => raise TYPE ("dest_binT", [T], []))
-
-fun dest_wordT T =
- (case T of
- Type (@{type_name "word"}, [T]) => dest_binT T
- | _ => raise TYPE ("dest_wordT", [T], []))
-
-
-val dest_nat = (fn
- SApp (SConst (@{const_name nat}, _), [SApp (SNum (i, _), _)]) => SOME i
- | _ => NONE)
-
-fun bv_rotate mk_name T ts =
- dest_nat (hd ts) |> Option.map (fn i => (mk_name i, tl ts))
-
-fun bv_extend mk_name T ts =
- (case (try dest_wordT (domain_type T), try dest_wordT (range_type T)) of
- (SOME i, SOME j) => if j - i >= 0 then SOME (mk_name (j - i), ts) else NONE
- | _ => NONE)
-
-fun bv_extract mk_name T ts =
- (case (try dest_wordT (body_type T), dest_nat (hd ts)) of
- (SOME i, SOME lb) => SOME (mk_name (i + lb - 1) lb, tl ts)
- | _ => NONE)
-
-end
--- a/src/HOL/SMT/Tools/smt_normalize.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/SMT/Tools/smt_normalize.ML Wed Oct 28 11:43:06 2009 +0000
@@ -14,8 +14,8 @@
signature SMT_NORMALIZE =
sig
val normalize_rule: Proof.context -> thm -> thm
- val instantiate_free: Thm.cterm * Thm.cterm -> thm -> thm
- val discharge_definition: Thm.cterm -> thm -> thm
+ val instantiate_free: cterm * cterm -> thm -> thm
+ val discharge_definition: cterm -> thm -> thm
val trivial_let: Proof.context -> thm list -> thm list
val positive_numerals: Proof.context -> thm list -> thm list
@@ -33,8 +33,7 @@
AddFunUpdRules |
AddAbsMinMaxRules
- val normalize: config list -> Proof.context -> thm list ->
- Thm.cterm list * thm list
+ val normalize: config list -> Proof.context -> thm list -> cterm list * thm list
val setup: theory -> theory
end
--- a/src/HOL/SMT/Tools/z3_proof.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/SMT/Tools/z3_proof.ML Wed Oct 28 11:43:06 2009 +0000
@@ -48,7 +48,7 @@
datatype context = Context of {
Ttab: typ Symtab.table,
- ttab: Thm.cterm Symtab.table,
+ ttab: cterm Symtab.table,
etab: T.preterm Inttab.table,
ptab: R.proof Inttab.table,
nctxt: Name.context }
--- a/src/HOL/SMT/Tools/z3_proof_rules.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/SMT/Tools/z3_proof_rules.ML Wed Oct 28 11:43:06 2009 +0000
@@ -13,7 +13,7 @@
(*proof reconstruction*)
type proof
- val make_proof: rule -> int list -> Thm.cterm * Thm.cterm list -> proof
+ val make_proof: rule -> int list -> cterm * cterm list -> proof
val prove: Proof.context -> thm list option -> proof Inttab.table -> int ->
thm
@@ -103,11 +103,11 @@
Unproved of {
rule: rule,
subs: int list,
- prop: Thm.cterm,
- vars: Thm.cterm list } |
+ prop: cterm,
+ vars: cterm list } |
Sequent of {
- hyps: Thm.cterm list,
- vars: Thm.cterm list,
+ hyps: cterm list,
+ vars: cterm list,
thm: theorem }
fun make_proof r ps (ct, cvs) = Unproved {rule=r, subs=ps, prop=ct, vars=cvs}
--- a/src/HOL/SMT/Tools/z3_proof_terms.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/SMT/Tools/z3_proof_terms.ML Wed Oct 28 11:43:06 2009 +0000
@@ -6,15 +6,15 @@
signature Z3_PROOF_TERMS =
sig
- val mk_prop: Thm.cterm -> Thm.cterm
- val mk_meta_eq: Thm.cterm -> Thm.cterm -> Thm.cterm
+ val mk_prop: cterm -> cterm
+ val mk_meta_eq: cterm -> cterm -> cterm
type preterm
- val compile: theory -> Name.context -> preterm -> Thm.cterm * Thm.cterm list
+ val compile: theory -> Name.context -> preterm -> cterm * cterm list
val mk_bound: theory -> int -> typ -> preterm
- val mk_fun: Thm.cterm -> preterm list -> preterm
+ val mk_fun: cterm -> preterm list -> preterm
val mk_forall: theory -> string * typ -> preterm -> preterm
val mk_exists: theory -> string * typ -> preterm -> preterm
@@ -73,8 +73,8 @@
datatype preterm = Preterm of {
- cterm: Thm.cterm,
- vars: (int * Thm.cterm) list }
+ cterm: cterm,
+ vars: (int * cterm) list }
fun mk_preterm (ct, vs) = Preterm {cterm=ct, vars=vs}
fun dest_preterm (Preterm {cterm, vars}) = (cterm, vars)
--- a/src/HOL/Statespace/state_space.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Statespace/state_space.ML Wed Oct 28 11:43:06 2009 +0000
@@ -297,7 +297,7 @@
val tt' = tt |> fold upd all_names;
val activate_simproc =
- Output.no_warnings_CRITICAL
+ Output.no_warnings_CRITICAL (* FIXME !?! *)
(Simplifier.map_ss (fn ss => ss addsimprocs [distinct_simproc]));
val ctxt' =
ctxt
--- a/src/HOL/String.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/String.thy Wed Oct 28 11:43:06 2009 +0000
@@ -155,7 +155,7 @@
datatype literal = STR string
-lemmas [code del] = literal.recs literal.cases
+declare literal.cases [code del] literal.recs [code del]
lemma [code]: "size (s\<Colon>literal) = 0"
by (cases s) simp_all
@@ -168,6 +168,9 @@
use "Tools/string_code.ML"
+code_reserved SML string
+code_reserved OCaml string
+
code_type literal
(SML "string")
(OCaml "string")
@@ -185,9 +188,6 @@
(OCaml "!((_ : string) = _)")
(Haskell infixl 4 "==")
-code_reserved SML string
-code_reserved OCaml string
-
types_code
"char" ("string")
--- a/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Wed Oct 28 11:43:06 2009 +0000
@@ -9,6 +9,7 @@
(*hooks for problem files*)
val destdir: string Config.T
val problem_prefix: string Config.T
+ val measure_runtime: bool Config.T
val setup: theory -> theory
(*prover configuration, problem format, and prover result*)
@@ -61,7 +62,10 @@
val (problem_prefix, problem_prefix_setup) =
Attrib.config_string "atp_problem_prefix" "prob";
-val setup = destdir_setup #> problem_prefix_setup;
+val (measure_runtime, measure_runtime_setup) =
+ Attrib.config_bool "atp_measure_runtime" false;
+
+val setup = destdir_setup #> problem_prefix_setup #> measure_runtime_setup;
(* prover configuration, problem format, and prover result *)
@@ -140,9 +144,14 @@
end;
(* write out problem file and call prover *)
- fun cmd_line probfile = "TIMEFORMAT='%3U'; { time " ^ space_implode " "
- [File.shell_path cmd, args, File.shell_path probfile] ^ " 2> /dev/null" ^
- " ; } 2>&1"
+ fun cmd_line probfile =
+ if Config.get ctxt measure_runtime
+ then (* Warning: suppresses error messages of ATPs *)
+ "TIMEFORMAT='%3U'; { time " ^ space_implode " " [File.shell_path cmd,
+ args, File.shell_path probfile] ^ " 2> /dev/null" ^ " ; } 2>&1"
+ else
+ space_implode " " ["exec", File.shell_path cmd, args,
+ File.shell_path probfile];
fun split_time s =
let
val split = String.tokens (fn c => str c = "\n");
@@ -154,10 +163,12 @@
val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b);
val as_time = the_default 0 o Scan.read Symbol.stopper time o explode;
in (proof, as_time t) end;
+ fun split_time' s =
+ if Config.get ctxt measure_runtime then split_time s else (s, 0)
fun run_on probfile =
if File.exists cmd then
write probfile clauses
- |> pair (apfst split_time (system_out (cmd_line probfile)))
+ |> pair (apfst split_time' (system_out (cmd_line probfile)))
else error ("Bad executable: " ^ Path.implode cmd);
(* if problemfile has not been exported, delete problemfile; otherwise export proof, too *)
--- a/src/HOL/Tools/Datatype/datatype.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Datatype/datatype.ML Wed Oct 28 11:43:06 2009 +0000
@@ -29,8 +29,7 @@
val make_case : Proof.context -> DatatypeCase.config -> string list -> term ->
(term * term) list -> term * (term * (int * bool)) list
val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
- val read_typ: theory ->
- (typ list * (string * sort) list) * string -> typ list * (string * sort) list
+ val read_typ: theory -> string -> (string * sort) list -> typ * (string * sort) list
val setup: theory -> theory
end;
@@ -160,23 +159,24 @@
(* prepare datatype specifications *)
-fun read_typ thy ((Ts, sorts), str) =
+fun read_typ thy str sorts =
let
val ctxt = ProofContext.init thy
|> fold (Variable.declare_typ o TFree) sorts;
val T = Syntax.read_typ ctxt str;
- in (Ts @ [T], Term.add_tfreesT T sorts) end;
+ in (T, Term.add_tfreesT T sorts) end;
-fun cert_typ sign ((Ts, sorts), raw_T) =
+fun cert_typ sign raw_T sorts =
let
- val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
- TYPE (msg, _, _) => error msg;
+ val T = Type.no_tvars (Sign.certify_typ sign raw_T)
+ handle TYPE (msg, _, _) => error msg;
val sorts' = Term.add_tfreesT T sorts;
- in (Ts @ [T],
+ val _ =
case duplicates (op =) (map fst sorts') of
- [] => sorts'
- | dups => error ("Inconsistent sort constraints for " ^ commas dups))
- end;
+ [] => ()
+ | dups => error ("Inconsistent sort constraints for " ^ commas dups)
+ in (T, sorts') end;
+
(* case names *)
@@ -460,8 +460,9 @@
let
fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
let
- val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
- val _ = (case subtract (op =) tvs (fold (curry OldTerm.add_typ_tfree_names) cargs' []) of
+ val (cargs', sorts'') = fold_map (prep_typ tmp_thy) cargs sorts';
+ val _ =
+ (case subtract (op =) tvs (fold (curry OldTerm.add_typ_tfree_names) cargs' []) of
[] => ()
| vs => error ("Extra type variables on rhs: " ^ commas vs))
in (constrs @ [(Sign.full_name_path tmp_thy tname'
--- a/src/HOL/Tools/Datatype/datatype_aux.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Datatype/datatype_aux.ML Wed Oct 28 11:43:06 2009 +0000
@@ -74,7 +74,7 @@
val get_rec_types : descr -> (string * sort) list -> typ list
val interpret_construction : descr -> (string * sort) list
-> { atyp: typ -> 'a, dtyp: typ list -> int * bool -> string * typ list -> 'a }
- -> ((string * Term.typ list) * (string * 'a list) list) list
+ -> ((string * typ list) * (string * 'a list) list) list
val check_nonempty : descr list -> unit
val unfold_datatypes :
theory -> descr -> (string * sort) list -> info Symtab.table ->
@@ -255,9 +255,8 @@
(* find all non-recursive types in datatype description *)
fun get_nonrec_types descr sorts =
- map (typ_of_dtyp descr sorts) (Library.foldl (fn (Ts, (_, (_, _, constrs))) =>
- Library.foldl (fn (Ts', (_, cargs)) =>
- union (op =) (filter_out is_rec_type cargs) Ts') (Ts, constrs)) ([], descr));
+ map (typ_of_dtyp descr sorts) (fold (fn (_, (_, _, constrs)) =>
+ fold (fn (_, cargs) => union (op =) (filter_out is_rec_type cargs)) constrs) descr []);
(* get all recursive types in datatype description *)
@@ -335,7 +334,7 @@
(* unfold a single constructor argument *)
- fun unfold_arg ((i, Ts, descrs), T) =
+ fun unfold_arg T (i, Ts, descrs) =
if is_rec_type T then
let val (Us, U) = strip_dtyp T
in if exists is_rec_type Us then
@@ -353,19 +352,17 @@
(* unfold a constructor *)
- fun unfold_constr ((i, constrs, descrs), (cname, cargs)) =
- let val (i', cargs', descrs') = Library.foldl unfold_arg ((i, [], descrs), cargs)
+ fun unfold_constr (cname, cargs) (i, constrs, descrs) =
+ let val (i', cargs', descrs') = fold unfold_arg cargs (i, [], descrs)
in (i', constrs @ [(cname, cargs')], descrs') end;
(* unfold a single datatype *)
- fun unfold_datatype ((i, dtypes, descrs), (j, (tname, tvars, constrs))) =
- let val (i', constrs', descrs') =
- Library.foldl unfold_constr ((i, [], descrs), constrs)
- in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs')
- end;
+ fun unfold_datatype (j, (tname, tvars, constrs)) (i, dtypes, descrs) =
+ let val (i', constrs', descrs') = fold unfold_constr constrs (i, [], descrs)
+ in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs') end;
- val (i', descr', descrs) = Library.foldl unfold_datatype ((i, [],[]), descr);
+ val (i', descr', descrs) = fold unfold_datatype descr (i, [], []);
in (descr' :: descrs, i') end;
--- a/src/HOL/Tools/Datatype/datatype_prop.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML Wed Oct 28 11:43:06 2009 +0000
@@ -221,7 +221,7 @@
(Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
(reccomb_names ~~ recTs ~~ rec_result_Ts);
- fun make_primrec T comb_t ((ts, f::fs), (cname, cargs)) =
+ fun make_primrec T comb_t (cname, cargs) (ts, f::fs) =
let
val recs = List.filter is_rec_type cargs;
val Ts = map (typ_of_dtyp descr' sorts) cargs;
@@ -242,11 +242,12 @@
in (ts @ [HOLogic.mk_Trueprop (HOLogic.mk_eq
(comb_t $ list_comb (Const (cname, Ts ---> T), frees),
list_comb (f, frees @ reccombs')))], fs)
- end
+ end;
- in fst (Library.foldl (fn (x, ((dt, T), comb_t)) =>
- Library.foldl (make_primrec T comb_t) (x, #3 (snd dt)))
- (([], rec_fns), descr' ~~ recTs ~~ reccombs))
+ in
+ fold (fn ((dt, T), comb_t) => fold (make_primrec T comb_t) (#3 (snd dt)))
+ (descr' ~~ recTs ~~ reccombs) ([], rec_fns)
+ |> fst
end;
(****************** make terms of form t_case f1 ... fn *********************)
--- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Wed Oct 28 11:43:06 2009 +0000
@@ -207,7 +207,7 @@
(* constructor definitions *)
- fun make_constr_def tname T n ((thy, defs, eqns, i), ((cname, cargs), (cname', mx))) =
+ fun make_constr_def tname T n ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) =
let
fun constr_arg (dt, (j, l_args, r_args)) =
let val T = typ_of_dtyp descr' sorts dt;
@@ -238,8 +238,8 @@
(* constructor definitions for datatype *)
- fun dt_constr_defs ((thy, defs, eqns, rep_congs, dist_lemmas),
- ((((_, (_, _, constrs)), tname), T), constr_syntax)) =
+ fun dt_constr_defs ((((_, (_, _, constrs)), tname), T), constr_syntax)
+ (thy, defs, eqns, rep_congs, dist_lemmas) =
let
val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
val rep_const = cterm_of thy
@@ -248,16 +248,18 @@
Drule.standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
val dist =
Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
- val (thy', defs', eqns', _) = Library.foldl ((make_constr_def tname T) (length constrs))
- ((Sign.add_path tname thy, defs, [], 1), constrs ~~ constr_syntax)
+ val (thy', defs', eqns', _) = fold ((make_constr_def tname T) (length constrs))
+ (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1);
in
(Sign.parent_path thy', defs', eqns @ [eqns'],
rep_congs @ [cong'], dist_lemmas @ [dist])
end;
- val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = Library.foldl dt_constr_defs
- ((thy3 |> Sign.add_consts_i iso_decls |> Sign.parent_path, [], [], [], []),
- hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax);
+ val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) =
+ fold dt_constr_defs
+ (hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax)
+ (thy3 |> Sign.add_consts_i iso_decls |> Sign.parent_path, [], [], [], []);
+
(*********** isomorphisms for new types (introduced by typedef) ***********)
@@ -283,7 +285,7 @@
(* e.g. dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *)
(*---------------------------------------------------------------------*)
- fun make_iso_def k ks n ((fs, eqns, i), (cname, cargs)) =
+ fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) =
let
val argTs = map (typ_of_dtyp descr' sorts) cargs;
val T = nth recTs k;
@@ -291,7 +293,7 @@
val rep_const = Const (rep_name, T --> Univ_elT);
val constr = Const (cname, argTs ---> T);
- fun process_arg ks' ((i2, i2', ts, Ts), dt) =
+ fun process_arg ks' dt (i2, i2', ts, Ts) =
let
val T' = typ_of_dtyp descr' sorts dt;
val (Us, U) = strip_type T'
@@ -307,12 +309,12 @@
| _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts))
end;
- val (i2, i2', ts, Ts) = Library.foldl (process_arg ks) ((1, 1, [], []), cargs);
+ val (i2, i2', ts, Ts) = fold (process_arg ks) cargs (1, 1, [], []);
val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i);
- val (_, _, ts', _) = Library.foldl (process_arg []) ((1, 1, [], []), cargs);
+ val (_, _, ts', _) = fold (process_arg []) cargs (1, 1, [], []);
val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
(rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i))
@@ -320,20 +322,20 @@
(* define isomorphisms for all mutually recursive datatypes in list ds *)
- fun make_iso_defs (ds, (thy, char_thms)) =
+ fun make_iso_defs ds (thy, char_thms) =
let
val ks = map fst ds;
val (_, (tname, _, _)) = hd ds;
val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname);
- fun process_dt ((fs, eqns, isos), (k, (tname, _, constrs))) =
+ fun process_dt (k, (tname, _, constrs)) (fs, eqns, isos) =
let
- val (fs', eqns', _) = Library.foldl (make_iso_def k ks (length constrs))
- ((fs, eqns, 1), constrs);
+ val (fs', eqns', _) =
+ fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1);
val iso = (nth recTs k, nth all_rep_names k)
in (fs', eqns', isos @ [iso]) end;
- val (fs, eqns, isos) = Library.foldl process_dt (([], [], []), ds);
+ val (fs, eqns, isos) = fold process_dt ds ([], [], []);
val fTs = map fastype_of fs;
val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Long_Name.base_name iso_name ^ "_def"),
Logic.mk_equals (Const (iso_name, T --> Univ_elT),
@@ -349,8 +351,8 @@
in (thy', char_thms' @ char_thms) end;
- val (thy5, iso_char_thms) = apfst Theory.checkpoint (List.foldr make_iso_defs
- (Sign.add_path big_name thy4, []) (tl descr));
+ val (thy5, iso_char_thms) = apfst Theory.checkpoint (fold_rev make_iso_defs
+ (tl descr) (Sign.add_path big_name thy4, []));
(* prove isomorphism properties *)
@@ -563,7 +565,7 @@
(map (fn r => r RS @{thm the_inv_f_f} RS subst) iso_inj_thms_unfolded);
val Rep_inverse_thms' = map (fn r => r RS @{thm the_inv_f_f}) iso_inj_thms_unfolded;
- fun mk_indrule_lemma ((prems, concls), ((i, _), T)) =
+ fun mk_indrule_lemma ((i, _), T) (prems, concls) =
let
val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $
mk_Free "x" T i;
@@ -582,7 +584,7 @@
end;
val (indrule_lemma_prems, indrule_lemma_concls) =
- Library.foldl mk_indrule_lemma (([], []), (descr' ~~ recTs));
+ fold mk_indrule_lemma (descr' ~~ recTs) ([], []);
val cert = cterm_of thy6;
--- a/src/HOL/Tools/Nitpick/kodkod.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/kodkod.ML Wed Oct 28 11:43:06 2009 +0000
@@ -157,6 +157,7 @@
type raw_bound = n_ary_index * int list list
datatype outcome =
+ NotInstalled |
Normal of (int * raw_bound list) list * int list |
TimedOut of int list |
Interrupted of int list option |
@@ -301,6 +302,7 @@
type raw_bound = n_ary_index * int list list
datatype outcome =
+ NotInstalled |
Normal of (int * raw_bound list) list * int list |
TimedOut of int list |
Interrupted of int list option |
@@ -1043,6 +1045,10 @@
let
val code =
system ("env CLASSPATH=\"$KODKODI_CLASSPATH:$CLASSPATH\" \
+ \JAVA_LIBRARY_PATH=\"$KODKODI_JAVA_LIBRARY_PATH:\
+ \$JAVA_LIBRARY_PATH\" \
+ \LD_LIBRARY_PATH=\"$KODKODI_JAVA_LIBRARY_PATH:\
+ \$LD_LIBRARY_PATH\" \
\\"$ISABELLE_TOOL\" java \
\de.tum.in.isabelle.Kodkodi.Kodkodi" ^
(if ms >= 0 then " -max-msecs " ^ Int.toString ms
@@ -1065,6 +1071,10 @@
if null ps then
if code = 2 then
TimedOut js
+ else if first_error |> Substring.full
+ |> Substring.position "NoClassDefFoundError" |> snd
+ |> Substring.isEmpty |> not then
+ NotInstalled
else if first_error <> "" then
Error (first_error |> perhaps (try (unsuffix "."))
|> perhaps (try (unprefix "Error: ")), js)
--- a/src/HOL/Tools/Nitpick/kodkod_sat.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/kodkod_sat.ML Wed Oct 28 11:43:06 2009 +0000
@@ -12,13 +12,15 @@
val sat_solver_spec : string -> string * string list
end;
-structure KodkodSAT : KODKOD_SAT =
+structure Kodkod_SAT : KODKOD_SAT =
struct
datatype sink = ToStdout | ToFile
+datatype availability = Java | JNI
+datatype mode = Batch | Incremental
datatype sat_solver_info =
- Internal of bool * string list |
+ Internal of availability * mode * string list |
External of sink * string * string * string list |
ExternalV2 of sink * string * string * string list * string * string * string
@@ -26,9 +28,11 @@
(* (string * sat_solver_info) list *)
val static_list =
- [("MiniSat", ExternalV2 (ToFile, "MINISAT_HOME", "minisat", [], "SAT", "",
+ [("MiniSatJNI", Internal (JNI, Incremental, ["MiniSat"])),
+ ("MiniSat", ExternalV2 (ToFile, "MINISAT_HOME", "minisat", [], "SAT", "",
"UNSAT")),
("PicoSAT", External (ToStdout, "PICOSAT_HOME", "picosat", [])),
+ ("zChaffJNI", Internal (JNI, Batch, ["zChaff"])),
("zChaff", ExternalV2 (ToStdout, "ZCHAFF_HOME", "zchaff", [],
"Instance Satisfiable", "",
"Instance Unsatisfiable")),
@@ -40,10 +44,8 @@
"solution =", "UNSATISFIABLE !!")),
("BerkMinAlloy", External (ToStdout, "BERKMINALLOY_HOME", "berkmin", [])),
("Jerusat", External (ToStdout, "JERUSAT_HOME", "Jerusat1.3", [])),
- ("SAT4J", Internal (true, ["DefaultSAT4J"])),
- ("MiniSatJNI", Internal (true, ["MiniSat"])),
- ("zChaffJNI", Internal (false, ["zChaff"])),
- ("SAT4JLight", Internal (true, ["LightSAT4J"])),
+ ("SAT4J", Internal (Java, Incremental, ["DefaultSAT4J"])),
+ ("SAT4JLight", Internal (Java, Incremental, ["LightSAT4J"])),
("HaifaSat", ExternalV2 (ToStdout, "HAIFASAT_HOME", "HaifaSat", ["-p", "1"],
"s SATISFIABLE", "v ", "s UNSATISFIABLE"))]
@@ -72,13 +74,27 @@
end)
(* bool -> string * sat_solver_info
-> (string * (unit -> string list)) option *)
-fun dynamic_entry_for_info false (name, Internal (_, ss)) = SOME (name, K ss)
+fun dynamic_entry_for_info incremental (name, Internal (Java, mode, ss)) =
+ if incremental andalso mode = Batch then NONE else SOME (name, K ss)
+ | dynamic_entry_for_info incremental (name, Internal (JNI, mode, ss)) =
+ if incremental andalso mode = Batch then
+ NONE
+ else
+ let
+ val lib_paths = getenv "KODKODI_JAVA_LIBRARY_PATH"
+ |> space_explode ":"
+ in
+ if exists (fn path => File.exists (Path.explode (path ^ "/")))
+ lib_paths then
+ SOME (name, K ss)
+ else
+ NONE
+ end
| dynamic_entry_for_info false (name, External (dev, home, exec, args)) =
dynamic_entry_for_external name dev home exec args []
| dynamic_entry_for_info false (name, ExternalV2 (dev, home, exec, args,
m1, m2, m3)) =
dynamic_entry_for_external name dev home exec args [m1, m2, m3]
- | dynamic_entry_for_info true (name, Internal (true, ss)) = SOME (name, K ss)
| dynamic_entry_for_info true _ = NONE
(* bool -> (string * (unit -> string list)) list *)
fun dynamic_list incremental =
--- a/src/HOL/Tools/Nitpick/minipick.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/minipick.ML Wed Oct 28 11:43:06 2009 +0000
@@ -14,10 +14,10 @@
struct
open Kodkod
-open NitpickUtil
-open NitpickHOL
-open NitpickPeephole
-open NitpickKodkod
+open Nitpick_Util
+open Nitpick_HOL
+open Nitpick_Peephole
+open Nitpick_Kodkod
(* theory -> typ -> unit *)
fun check_type ctxt (Type ("fun", Ts)) = List.app (check_type ctxt) Ts
--- a/src/HOL/Tools/Nitpick/nitpick.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick.ML Wed Oct 28 11:43:06 2009 +0000
@@ -62,15 +62,15 @@
structure Nitpick : NITPICK =
struct
-open NitpickUtil
-open NitpickHOL
-open NitpickMono
-open NitpickScope
-open NitpickPeephole
-open NitpickRep
-open NitpickNut
-open NitpickKodkod
-open NitpickModel
+open Nitpick_Util
+open Nitpick_HOL
+open Nitpick_Mono
+open Nitpick_Scope
+open Nitpick_Peephole
+open Nitpick_Rep
+open Nitpick_Nut
+open Nitpick_Kodkod
+open Nitpick_Model
type params = {
cards_assigns: (typ option * int list) list,
@@ -355,21 +355,21 @@
val effective_sat_solver =
if sat_solver <> "smart" then
if need_incremental andalso
- not (sat_solver mem KodkodSAT.configured_sat_solvers true) then
+ not (sat_solver mem Kodkod_SAT.configured_sat_solvers true) then
(print_m (K ("An incremental SAT solver is required: \"SAT4J\" will \
\be used instead of " ^ quote sat_solver ^ "."));
"SAT4J")
else
sat_solver
else
- KodkodSAT.smart_sat_solver_name need_incremental
+ Kodkod_SAT.smart_sat_solver_name need_incremental
val _ =
if sat_solver = "smart" then
print_v (fn () => "Using SAT solver " ^ quote effective_sat_solver ^
". The following" ^
(if need_incremental then " incremental " else " ") ^
"solvers are configured: " ^
- commas (map quote (KodkodSAT.configured_sat_solvers
+ commas (map quote (Kodkod_SAT.configured_sat_solvers
need_incremental)) ^ ".")
else
()
@@ -439,7 +439,7 @@
val formula = fold (fold s_and) [def_fs, nondef_fs] core_f
val comment = (if liberal then "liberal" else "conservative") ^ "\n" ^
PrintMode.setmp [] multiline_string_for_scope scope
- val kodkod_sat_solver = KodkodSAT.sat_solver_spec effective_sat_solver
+ val kodkod_sat_solver = Kodkod_SAT.sat_solver_spec effective_sat_solver
|> snd
val delay = if liberal then
Option.map (fn time => Time.- (time, Time.now ()))
@@ -483,7 +483,7 @@
defs = nondef_fs @ def_fs @ declarative_axioms})
end
handle LIMIT (loc, msg) =>
- if loc = "NitpickKodkod.check_arity"
+ if loc = "Nitpick_Kodkod.check_arity"
andalso not (Typtab.is_empty ofs) then
problem_for_scope liberal
{ext_ctxt = ext_ctxt, card_assigns = card_assigns,
@@ -626,7 +626,18 @@
else
case Kodkod.solve_any_problem overlord deadline max_threads
max_solutions (map fst problems) of
- Kodkod.Normal ([], unsat_js) =>
+ Kodkod.NotInstalled =>
+ (print_m (fn () =>
+ "Nitpick requires the external Java program Kodkodi. \
+ \To install it, download the package from Isabelle's \
+ \web page and add the \"kodkodi-x.y.z\" directory's \
+ \full path to \"" ^
+ Path.implode (Path.expand (Path.appends
+ (Path.variable "ISABELLE_HOME" ::
+ (map Path.basic ["etc", "components"])))) ^
+ "\".");
+ (max_potential, max_genuine, donno + 1))
+ | Kodkod.Normal ([], unsat_js) =>
(update_checked_problems problems unsat_js;
(max_potential, max_genuine, donno))
| Kodkod.Normal (sat_ps, unsat_js) =>
@@ -785,7 +796,7 @@
(* int -> int -> scope list -> int * int * int -> Kodkod.outcome *)
fun run_batches _ _ [] (max_potential, max_genuine, donno) =
if donno > 0 andalso max_genuine > 0 then
- (print_m (fn () => excipit "ran out of resources"); "unknown")
+ (print_m (fn () => excipit "ran into difficulties"); "unknown")
else if max_genuine = original_max_genuine then
if max_potential = original_max_potential then
(print_m (K ("Nitpick found no " ^ word_model ^ ".")); "none")
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Wed Oct 28 11:43:06 2009 +0000
@@ -139,10 +139,10 @@
extended_context -> term -> ((term list * term list) * (bool * bool)) * term
end;
-structure NitpickHOL : NITPICK_HOL =
+structure Nitpick_HOL : NITPICK_HOL =
struct
-open NitpickUtil
+open Nitpick_Util
type const_table = term list Symtab.table
type special_fun = (styp * int list * term list) * styp
@@ -263,7 +263,7 @@
val after_name_sep = snd o strip_first_name_sep
(* When you add constants to these lists, make sure to handle them in
- "NitpickNut.nut_from_term", and perhaps in "NitpickMono.consider_term" as
+ "Nitpick_Nut.nut_from_term", and perhaps in "Nitpick_Mono.consider_term" as
well. *)
val built_in_consts =
[(@{const_name all}, 1),
@@ -405,7 +405,7 @@
(* typ -> styp *)
fun const_for_iterator_type (Type (s, Ts)) = (after_name_sep s, Ts ---> bool_T)
| const_for_iterator_type T =
- raise TYPE ("NitpickHOL.const_for_iterator_type", [T], [])
+ raise TYPE ("Nitpick_HOL.const_for_iterator_type", [T], [])
(* int -> typ -> typ * typ *)
fun strip_n_binders 0 T = ([], T)
@@ -413,7 +413,7 @@
strip_n_binders (n - 1) T2 |>> cons T1
| strip_n_binders n (Type (@{type_name fun_box}, Ts)) =
strip_n_binders n (Type ("fun", Ts))
- | strip_n_binders _ T = raise TYPE ("NitpickHOL.strip_n_binders", [T], [])
+ | strip_n_binders _ T = raise TYPE ("Nitpick_HOL.strip_n_binders", [T], [])
(* typ -> typ *)
val nth_range_type = snd oo strip_n_binders
@@ -432,7 +432,7 @@
fun mk_flat_tuple _ [t] = t
| mk_flat_tuple (Type ("*", [T1, T2])) (t :: ts) =
HOLogic.pair_const T1 T2 $ t $ (mk_flat_tuple T2 ts)
- | mk_flat_tuple T ts = raise TYPE ("NitpickHOL.mk_flat_tuple", [T], ts)
+ | mk_flat_tuple T ts = raise TYPE ("Nitpick_HOL.mk_flat_tuple", [T], ts)
(* int -> term -> term list *)
fun dest_n_tuple 1 t = [t]
| dest_n_tuple n t = HOLogic.dest_prod t ||> dest_n_tuple (n - 1) |> op ::
@@ -441,7 +441,8 @@
fun dest_n_tuple_type 1 T = [T]
| dest_n_tuple_type n (Type (_, [T1, T2])) =
T1 :: dest_n_tuple_type (n - 1) T2
- | dest_n_tuple_type _ T = raise TYPE ("NitpickHOL.dest_n_tuple_type", [T], [])
+ | dest_n_tuple_type _ T =
+ raise TYPE ("Nitpick_HOL.dest_n_tuple_type", [T], [])
(* (typ * typ) list -> typ -> typ *)
fun typ_subst [] T = T
@@ -460,7 +461,7 @@
(Sign.typ_match thy (Logic.varifyT T1, T1') Vartab.empty))
(Logic.varifyT T2)
handle Type.TYPE_MATCH =>
- raise TYPE ("NitpickHOL.instantiate_type", [T1, T1'], [])
+ raise TYPE ("Nitpick_HOL.instantiate_type", [T1, T1'], [])
(* theory -> typ -> typ -> styp *)
fun repair_constr_type thy body_T' T =
@@ -483,7 +484,7 @@
val (co_s, co_Ts) = dest_Type co_T
val _ =
if forall is_TFree co_Ts andalso not (has_duplicates (op =) co_Ts) then ()
- else raise TYPE ("NitpickHOL.register_codatatype", [co_T], [])
+ else raise TYPE ("Nitpick_HOL.register_codatatype", [co_T], [])
val codatatypes = AList.update (op =) (co_s, (case_name, constr_xs))
codatatypes
in TheoryData.put {frac_types = frac_types, codatatypes = codatatypes} thy end
@@ -586,8 +587,8 @@
fun mate_of_rep_fun thy (x as (_, Type ("fun", [T1 as Type (s', _), T2]))) =
(case typedef_info thy s' of
SOME {Abs_name, ...} => (Abs_name, Type ("fun", [T2, T1]))
- | NONE => raise TERM ("NitpickHOL.mate_of_rep_fun", [Const x]))
- | mate_of_rep_fun _ x = raise TERM ("NitpickHOL.mate_of_rep_fun", [Const x])
+ | NONE => raise TERM ("Nitpick_HOL.mate_of_rep_fun", [Const x]))
+ | mate_of_rep_fun _ x = raise TERM ("Nitpick_HOL.mate_of_rep_fun", [Const x])
(* theory -> styp -> bool *)
fun is_coconstr thy (s, T) =
@@ -874,7 +875,7 @@
case AList.lookup (op =) asgns T of
SOME k => k
| NONE => if T = @{typ bisim_iterator} then 0
- else raise TYPE ("NitpickHOL.card_of_type", [T], [])
+ else raise TYPE ("Nitpick_HOL.card_of_type", [T], [])
(* int -> (typ * int) list -> typ -> int *)
fun bounded_card_of_type max default_card asgns (Type ("fun", [T1, T2])) =
let
@@ -894,7 +895,7 @@
card_of_type asgns T
else
card_of_type asgns T
- handle TYPE ("NitpickHOL.card_of_type", _, _) =>
+ handle TYPE ("Nitpick_HOL.card_of_type", _, _) =>
default_card)
(* theory -> int -> (typ * int) list -> typ -> int *)
fun bounded_precise_card_of_type thy max default_card asgns T =
@@ -1110,13 +1111,13 @@
(* term -> term *)
fun aux (v as Var _) t = lambda v t
| aux (c as Const (@{const_name TYPE}, T)) t = lambda c t
- | aux _ _ = raise TERM ("NitpickHOL.normalized_rhs_of", [t])
+ | aux _ _ = raise TERM ("Nitpick_HOL.normalized_rhs_of", [t])
val (lhs, rhs) =
case t of
Const (@{const_name "=="}, _) $ t1 $ t2 => (t1, t2)
| @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t1 $ t2) =>
(t1, t2)
- | _ => raise TERM ("NitpickHOL.normalized_rhs_of", [t])
+ | _ => raise TERM ("Nitpick_HOL.normalized_rhs_of", [t])
val args = strip_comb lhs |> snd
in fold_rev aux args rhs end
@@ -1170,7 +1171,7 @@
case term_under_def t of
Const (s, _) => (s, t)
| Free _ => raise NOT_SUPPORTED "local definitions"
- | t' => raise TERM ("NitpickHOL.pair_for_prop", [t, t'])
+ | t' => raise TERM ("Nitpick_HOL.pair_for_prop", [t, t'])
(* (Proof.context -> term list) -> Proof.context -> const_table *)
fun table_for get ctxt =
@@ -1308,7 +1309,7 @@
|> fold_rev (curry absdummy) (func_Ts @ [dataT])
end
-val redefined_in_NitpickDefs_thy =
+val redefined_in_Nitpick_thy =
[@{const_name option_case}, @{const_name nat_case}, @{const_name list_case},
@{const_name list_size}]
@@ -1325,7 +1326,8 @@
select_nth_constr_arg thy constr_x t j res_T
|> optimized_record_get thy s rec_T' res_T
end
- | _ => raise TYPE ("NitpickHOL.optimized_record_get", [rec_T], []))
+ | _ => raise TYPE ("Nitpick_HOL.optimized_record_get", [rec_T],
+ []))
| j => select_nth_constr_arg thy constr_x t j res_T
end
(* theory -> string -> typ -> term -> term -> term *)
@@ -1380,7 +1382,7 @@
(is_real_equational_fun ext_ctxt orf is_real_inductive_pred ext_ctxt
orf (String.isPrefix ubfp_prefix orf String.isPrefix lbfp_prefix) o fst)
andf (not o has_trivial_definition thy def_table)
- andf (not o member (op =) redefined_in_NitpickDefs_thy o fst)
+ andf (not o member (op =) redefined_in_Nitpick_thy o fst)
(* term * term -> term *)
fun s_betapply (Const (@{const_name If}, _) $ @{const True} $ t, _) = t
@@ -1532,7 +1534,7 @@
else case def_of_const thy def_table x of
SOME def =>
if depth > unfold_max_depth then
- raise LIMIT ("NitpickHOL.unfold_defs_in_term",
+ raise LIMIT ("Nitpick_HOL.unfold_defs_in_term",
"too many nested definitions (" ^
string_of_int depth ^ ") while expanding " ^
quote s)
@@ -1640,7 +1642,8 @@
({thy, ctxt, debug, fast_descrs, tac_timeout, intro_table, ...}
: extended_context) (x as (_, T)) =
case def_props_for_const thy fast_descrs intro_table x of
- [] => raise TERM ("NitpickHOL.is_is_well_founded_inductive_pred", [Const x])
+ [] => raise TERM ("Nitpick_HOL.is_is_well_founded_inductive_pred",
+ [Const x])
| intro_ts =>
(case map (triple_for_intro_rule thy x) intro_ts
|> filter_out (null o #2) of
@@ -1772,7 +1775,7 @@
|> ap_split tuple_T bool_T))
end
| aux t =
- raise TERM ("NitpickHOL.linear_pred_base_and_step_rhss.aux", [t])
+ raise TERM ("Nitpick_HOL.linear_pred_base_and_step_rhss.aux", [t])
in aux end
(* extended_context -> styp -> term -> term *)
@@ -1834,8 +1837,8 @@
val rhs = case fp_app of
Const _ $ t =>
betapply (t, list_comb (Const x', next :: outer_bounds))
- | _ => raise TERM ("NitpickHOL.unrolled_inductive_pred_const",
- [fp_app])
+ | _ => raise TERM ("Nitpick_HOL.unrolled_inductive_pred_\
+ \const", [fp_app])
val (inner, naked_rhs) = strip_abs rhs
val all = outer @ inner
val bounds = map Bound (length all - 1 downto 0)
@@ -1854,7 +1857,7 @@
val outer_bounds = map Bound (length outer - 1 downto 0)
val rhs = case fp_app of
Const _ $ t => betapply (t, list_comb (Const x, outer_bounds))
- | _ => raise TERM ("NitpickHOL.raw_inductive_pred_axiom",
+ | _ => raise TERM ("Nitpick_HOL.raw_inductive_pred_axiom",
[fp_app])
val (inner, naked_rhs) = strip_abs rhs
val all = outer @ inner
@@ -1876,7 +1879,7 @@
(* extended_context -> styp -> term list *)
fun raw_equational_fun_axioms (ext_ctxt as {thy, fast_descrs, simp_table,
psimp_table, ...}) (x as (s, _)) =
- if s mem redefined_in_NitpickDefs_thy then
+ if s mem redefined_in_Nitpick_thy then
[]
else case def_props_for_const thy fast_descrs (!simp_table) x of
[] => (case def_props_for_const thy fast_descrs psimp_table x of
@@ -2329,7 +2332,7 @@
ts
(* (term * int list) list -> term *)
fun mk_connection [] =
- raise ARG ("NitpickHOL.push_quantifiers_inward.aux.\
+ raise ARG ("Nitpick_HOL.push_quantifiers_inward.aux.\
\mk_connection", "")
| mk_connection ts_cum_bounds =
ts_cum_bounds |> map fst
@@ -2720,7 +2723,7 @@
|> new_s <> "fun"
? construct_value thy (@{const_name FunBox},
Type ("fun", new_Ts) --> new_T) o single
- | t' => raise TERM ("NitpickHOL.box_fun_and_pair_in_term.\
+ | t' => raise TERM ("Nitpick_HOL.box_fun_and_pair_in_term.\
\coerce_term", [t']))
| (Type (new_s, new_Ts as [new_T1, new_T2]),
Type (old_s, old_Ts as [old_T1, old_T2])) =>
@@ -2740,7 +2743,7 @@
else @{const_name PairBox}, new_Ts ---> new_T)
[coerce_term Ts new_T1 old_T1 t1,
coerce_term Ts new_T2 old_T2 t2]
- | t' => raise TERM ("NitpickHOL.box_fun_and_pair_in_term.\
+ | t' => raise TERM ("Nitpick_HOL.box_fun_and_pair_in_term.\
\coerce_term", [t'])
else
raise TYPE ("coerce_term", [new_T, old_T], [t])
@@ -2753,7 +2756,7 @@
(case T' of
Type (_, [T1, T2]) =>
fold (add_boxed_types_for_var z) [(T1, t1), (T2, t2)]
- | _ => raise TYPE ("NitpickHOL.box_fun_and_pair_in_term.\
+ | _ => raise TYPE ("Nitpick_HOL.box_fun_and_pair_in_term.\
\add_boxed_types_for_var", [T'], []))
| _ => exists_subterm (equal (Var z)) t' ? insert (op =) T
(* typ list -> typ list -> term -> indexname * typ -> typ *)
@@ -3008,7 +3011,7 @@
else
let val accum as (xs, _) = (x :: xs, axs) in
if depth > axioms_max_depth then
- raise LIMIT ("NitpickHOL.axioms_for_term.add_axioms_for_term",
+ raise LIMIT ("Nitpick_HOL.axioms_for_term.add_axioms_for_term",
"too many nested axioms (" ^ string_of_int depth ^
")")
else if Refute.is_const_of_class thy x then
@@ -3081,7 +3084,7 @@
[] => t
| [(x, S)] =>
Refute.monomorphic_term (Vartab.make [(x, (S, T))]) t
- | _ => raise TERM ("NitpickHOL.axioms_for_term.\
+ | _ => raise TERM ("Nitpick_HOL.axioms_for_term.\
\add_axioms_for_sort", [t]))
class_axioms
in fold (add_nondef_axiom depth) monomorphic_class_axioms end
--- a/src/HOL/Tools/Nitpick/nitpick_isar.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_isar.ML Wed Oct 28 11:43:06 2009 +0000
@@ -13,13 +13,13 @@
val default_params : theory -> (string * string) list -> params
end
-structure NitpickIsar : NITPICK_ISAR =
+structure Nitpick_Isar : NITPICK_ISAR =
struct
-open NitpickUtil
-open NitpickHOL
-open NitpickRep
-open NitpickNut
+open Nitpick_Util
+open Nitpick_HOL
+open Nitpick_Rep
+open Nitpick_Nut
open Nitpick
type raw_param = string * string list
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Wed Oct 28 11:43:06 2009 +0000
@@ -7,10 +7,10 @@
signature NITPICK_KODKOD =
sig
- type extended_context = NitpickHOL.extended_context
- type dtype_spec = NitpickScope.dtype_spec
- type kodkod_constrs = NitpickPeephole.kodkod_constrs
- type nut = NitpickNut.nut
+ type extended_context = Nitpick_HOL.extended_context
+ type dtype_spec = Nitpick_Scope.dtype_spec
+ type kodkod_constrs = Nitpick_Peephole.kodkod_constrs
+ type nut = Nitpick_Nut.nut
type nfa_transition = Kodkod.rel_expr * typ
type nfa_entry = typ * nfa_transition list
type nfa_table = nfa_entry list
@@ -37,15 +37,15 @@
int Typtab.table -> bool -> kodkod_constrs -> nut -> Kodkod.formula
end;
-structure NitpickKodkod : NITPICK_KODKOD =
+structure Nitpick_Kodkod : NITPICK_KODKOD =
struct
-open NitpickUtil
-open NitpickHOL
-open NitpickScope
-open NitpickPeephole
-open NitpickRep
-open NitpickNut
+open Nitpick_Util
+open Nitpick_HOL
+open Nitpick_Scope
+open Nitpick_Peephole
+open Nitpick_Rep
+open Nitpick_Nut
type nfa_transition = Kodkod.rel_expr * typ
type nfa_entry = typ * nfa_transition list
@@ -89,7 +89,7 @@
(* int -> int -> unit *)
fun check_arity univ_card n =
if n > Kodkod.max_arity univ_card then
- raise LIMIT ("NitpickKodkod.check_arity",
+ raise LIMIT ("Nitpick_Kodkod.check_arity",
"arity " ^ string_of_int n ^ " too large for universe of \
\cardinality " ^ string_of_int univ_card)
else
@@ -132,7 +132,7 @@
(* int -> unit *)
fun check_table_size k =
if k > max_table_size then
- raise LIMIT ("NitpickKodkod.check_table_size",
+ raise LIMIT ("Nitpick_Kodkod.check_table_size",
"precomputed table too large (" ^ string_of_int k ^ ")")
else
()
@@ -245,7 +245,7 @@
("norm_frac", tabulate_int_op2_2 debug univ_card (int_card, j0)
isa_norm_frac)
else
- raise ARG ("NitpickKodkod.tabulate_built_in_rel", "unknown relation"))
+ raise ARG ("Nitpick_Kodkod.tabulate_built_in_rel", "unknown relation"))
(* bool -> int -> int -> int -> int -> int * int -> Kodkod.rel_expr
-> Kodkod.bound *)
@@ -266,11 +266,11 @@
if nick = @{const_name bisim_iterator_max} then
case R of
Atom (k, j0) => [Kodkod.TupleSet [Kodkod.Tuple [k - 1 + j0]]]
- | _ => raise NUT ("NitpickKodkod.bound_for_plain_rel", [u])
+ | _ => raise NUT ("Nitpick_Kodkod.bound_for_plain_rel", [u])
else
[Kodkod.TupleSet [], upper_bound_for_rep R])
| bound_for_plain_rel _ _ u =
- raise NUT ("NitpickKodkod.bound_for_plain_rel", [u])
+ raise NUT ("Nitpick_Kodkod.bound_for_plain_rel", [u])
(* Proof.context -> bool -> dtype_spec list -> nut -> Kodkod.bound *)
fun bound_for_sel_rel ctxt debug dtypes
@@ -293,7 +293,7 @@
end)
end
| bound_for_sel_rel _ _ _ u =
- raise NUT ("NitpickKodkod.bound_for_sel_rel", [u])
+ raise NUT ("Nitpick_Kodkod.bound_for_sel_rel", [u])
(* Kodkod.bound list -> Kodkod.bound list *)
fun merge_bounds bs =
@@ -320,7 +320,7 @@
if is_lone_rep R then
all_combinations_for_rep R |> map singleton_from_combination
else
- raise REP ("NitpickKodkod.all_singletons_for_rep", [R])
+ raise REP ("Nitpick_Kodkod.all_singletons_for_rep", [R])
(* Kodkod.rel_expr -> Kodkod.rel_expr list *)
fun unpack_products (Kodkod.Product (r1, r2)) =
@@ -333,7 +333,7 @@
val empty_rel_for_rep = empty_n_ary_rel o arity_of_rep
fun full_rel_for_rep R =
case atom_schema_of_rep R of
- [] => raise REP ("NitpickKodkod.full_rel_for_rep", [R])
+ [] => raise REP ("Nitpick_Kodkod.full_rel_for_rep", [R])
| schema => foldl1 Kodkod.Product (map Kodkod.AtomSeq schema)
(* int -> int list -> Kodkod.decl list *)
@@ -424,7 +424,7 @@
fun rel_expr_from_formula kk R f =
case unopt_rep R of
Atom (2, j0) => atom_from_formula kk j0 f
- | _ => raise REP ("NitpickKodkod.rel_expr_from_formula", [R])
+ | _ => raise REP ("Nitpick_Kodkod.rel_expr_from_formula", [R])
(* kodkod_cotrs -> int -> int -> Kodkod.rel_expr -> Kodkod.rel_expr list *)
fun unpack_vect_in_chunks ({kk_project_seq, ...} : kodkod_constrs) chunk_arity
@@ -471,13 +471,13 @@
if is_lone_rep old_R andalso is_lone_rep new_R
andalso card = card_of_rep new_R then
if card >= lone_rep_fallback_max_card then
- raise LIMIT ("NitpickKodkod.lone_rep_fallback",
+ raise LIMIT ("Nitpick_Kodkod.lone_rep_fallback",
"too high cardinality (" ^ string_of_int card ^ ")")
else
kk_case_switch kk old_R new_R r (all_singletons_for_rep old_R)
(all_singletons_for_rep new_R)
else
- raise REP ("NitpickKodkod.lone_rep_fallback", [old_R, new_R])
+ raise REP ("Nitpick_Kodkod.lone_rep_fallback", [old_R, new_R])
end
(* kodkod_constrs -> int * int -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *)
and atom_from_rel_expr kk (x as (k, j0)) old_R r =
@@ -490,7 +490,7 @@
atom_from_rel_expr kk x (Vect (dom_card, R2'))
(vect_from_rel_expr kk dom_card R2' old_R r)
end
- | Opt _ => raise REP ("NitpickKodkod.atom_from_rel_expr", [old_R])
+ | Opt _ => raise REP ("Nitpick_Kodkod.atom_from_rel_expr", [old_R])
| _ => lone_rep_fallback kk (Atom x) old_R r
(* kodkod_constrs -> rep list -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *)
and struct_from_rel_expr kk Rs old_R r =
@@ -515,7 +515,7 @@
else
lone_rep_fallback kk (Struct Rs) old_R r
end
- | _ => raise REP ("NitpickKodkod.struct_from_rel_expr", [old_R])
+ | _ => raise REP ("Nitpick_Kodkod.struct_from_rel_expr", [old_R])
(* kodkod_constrs -> int -> rep -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *)
and vect_from_rel_expr kk k R old_R r =
case old_R of
@@ -530,7 +530,7 @@
rel_expr_from_formula kk R (#kk_subset kk arg_r r))
(all_singletons_for_rep R1))
else
- raise REP ("NitpickKodkod.vect_from_rel_expr", [old_R])
+ raise REP ("Nitpick_Kodkod.vect_from_rel_expr", [old_R])
| Func (Unit, R2) => rel_expr_from_rel_expr kk R R2 r
| Func (R1, R2) =>
fold1 (#kk_product kk)
@@ -538,7 +538,7 @@
rel_expr_from_rel_expr kk R R2
(kk_n_fold_join kk true R1 R2 arg_r r))
(all_singletons_for_rep R1))
- | _ => raise REP ("NitpickKodkod.vect_from_rel_expr", [old_R])
+ | _ => raise REP ("Nitpick_Kodkod.vect_from_rel_expr", [old_R])
(* kodkod_constrs -> rep -> rep -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *)
and func_from_no_opt_rel_expr kk R1 R2 (Atom x) r =
let
@@ -555,12 +555,12 @@
| Func (Atom (1, _), Formula Neut) =>
(case unopt_rep R2 of
Atom (2, j0) => atom_from_formula kk j0 (#kk_some kk r)
- | _ => raise REP ("NitpickKodkod.func_from_no_opt_rel_expr",
+ | _ => raise REP ("Nitpick_Kodkod.func_from_no_opt_rel_expr",
[old_R, Func (Unit, R2)]))
| Func (R1', R2') =>
rel_expr_from_rel_expr kk R2 R2' (#kk_project_seq kk r (arity_of_rep R1')
(arity_of_rep R2'))
- | _ => raise REP ("NitpickKodkod.func_from_no_opt_rel_expr",
+ | _ => raise REP ("Nitpick_Kodkod.func_from_no_opt_rel_expr",
[old_R, Func (Unit, R2)]))
| func_from_no_opt_rel_expr kk R1 (Formula Neut) old_R r =
(case old_R of
@@ -592,7 +592,7 @@
| Func (R1', Atom (2, j0)) =>
func_from_no_opt_rel_expr kk R1 (Formula Neut)
(Func (R1', Formula Neut)) (#kk_join kk r (Kodkod.Atom (j0 + 1)))
- | _ => raise REP ("NitpickKodkod.func_from_no_opt_rel_expr",
+ | _ => raise REP ("Nitpick_Kodkod.func_from_no_opt_rel_expr",
[old_R, Func (R1, Formula Neut)]))
| func_from_no_opt_rel_expr kk R1 R2 old_R r =
case old_R of
@@ -621,7 +621,7 @@
(#kk_rel_eq kk r2 r3)
end
end
- | _ => raise REP ("NitpickKodkod.func_from_no_opt_rel_expr",
+ | _ => raise REP ("Nitpick_Kodkod.func_from_no_opt_rel_expr",
[old_R, Func (R1, R2)]))
| Func (Unit, R2') =>
let val j0 = some_j0 in
@@ -648,7 +648,7 @@
(dom_schema @ ran_schema))
(#kk_subset kk ran_prod app)
end
- | _ => raise REP ("NitpickKodkod.func_from_no_opt_rel_expr",
+ | _ => raise REP ("Nitpick_Kodkod.func_from_no_opt_rel_expr",
[old_R, Func (R1, R2)])
(* kodkod_constrs -> rep -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *)
and rel_expr_from_rel_expr kk new_R old_R r =
@@ -657,7 +657,7 @@
val unopt_new_R = unopt_rep new_R
in
if unopt_old_R <> old_R andalso unopt_new_R = new_R then
- raise REP ("NitpickKodkod.rel_expr_from_rel_expr", [old_R, new_R])
+ raise REP ("Nitpick_Kodkod.rel_expr_from_rel_expr", [old_R, new_R])
else if unopt_new_R = unopt_old_R then
r
else
@@ -666,7 +666,7 @@
| Struct Rs => struct_from_rel_expr kk Rs
| Vect (k, R') => vect_from_rel_expr kk k R'
| Func (R1, R2) => func_from_no_opt_rel_expr kk R1 R2
- | _ => raise REP ("NitpickKodkod.rel_expr_from_rel_expr",
+ | _ => raise REP ("Nitpick_Kodkod.rel_expr_from_rel_expr",
[old_R, new_R]))
unopt_old_R r
end
@@ -683,13 +683,13 @@
else if is_lone_rep R andalso card_of_rep R > 1 then kk_lone (Kodkod.Rel x)
else Kodkod.True
| declarative_axiom_for_plain_rel _ u =
- raise NUT ("NitpickKodkod.declarative_axiom_for_plain_rel", [u])
+ raise NUT ("Nitpick_Kodkod.declarative_axiom_for_plain_rel", [u])
(* nut NameTable.table -> styp -> Kodkod.rel_expr * rep * int *)
fun const_triple rel_table (x as (s, T)) =
case the_name rel_table (ConstName (s, T, Any)) of
FreeRel ((n, j), _, R, _) => (Kodkod.Rel (n, j), R, n)
- | _ => raise TERM ("NitpickKodkod.const_triple", [Const x])
+ | _ => raise TERM ("Nitpick_Kodkod.const_triple", [Const x])
(* nut NameTable.table -> styp -> Kodkod.rel_expr *)
fun discr_rel_expr rel_table = #1 o const_triple rel_table o discr_for_constr
@@ -849,7 +849,8 @@
(~1 upto num_sels - 1)
val j0 = case triples |> hd |> #2 of
Func (Atom (_, j0), _) => j0
- | R => raise REP ("NitpickKodkod.uniqueness_axiom_for_constr", [R])
+ | R => raise REP ("Nitpick_Kodkod.uniqueness_axiom_for_constr",
+ [R])
val set_r = triples |> hd |> #1
in
if num_sels = 0 then
@@ -960,7 +961,7 @@
let val opt1 = is_opt_rep (rep_of u1) in
case polar of
Neut => if opt1 then
- raise NUT ("NitpickKodkod.to_f (Finite)", [u])
+ raise NUT ("Nitpick_Kodkod.to_f (Finite)", [u])
else
Kodkod.True
| Pos => formula_for_bool (not opt1)
@@ -992,7 +993,7 @@
if not (is_opt_rep ran_R) then
to_set_bool_op kk_implies kk_subset u1 u2
else if polar = Neut then
- raise NUT ("NitpickKodkod.to_f (Subset)", [u])
+ raise NUT ("Nitpick_Kodkod.to_f (Subset)", [u])
else
let
(* bool -> nut -> Kodkod.rel_expr *)
@@ -1102,16 +1103,16 @@
| Op3 (If, _, _, u1, u2, u3) =>
kk_formula_if (to_f u1) (to_f u2) (to_f u3)
| FormulaReg (j, _, _) => Kodkod.FormulaReg j
- | _ => raise NUT ("NitpickKodkod.to_f", [u]))
+ | _ => raise NUT ("Nitpick_Kodkod.to_f", [u]))
| Atom (2, j0) => formula_from_atom j0 (to_r u)
- | _ => raise NUT ("NitpickKodkod.to_f", [u])
+ | _ => raise NUT ("Nitpick_Kodkod.to_f", [u])
(* polarity -> nut -> Kodkod.formula *)
and to_f_with_polarity polar u =
case rep_of u of
Formula _ => to_f u
| Atom (2, j0) => formula_from_atom j0 (to_r u)
| Opt (Atom (2, j0)) => formula_from_opt_atom polar j0 (to_r u)
- | _ => raise NUT ("NitpickKodkod.to_f_with_polarity", [u])
+ | _ => raise NUT ("Nitpick_Kodkod.to_f_with_polarity", [u])
(* nut -> Kodkod.rel_expr *)
and to_r u =
case u of
@@ -1142,12 +1143,12 @@
| Cst (Num j, @{typ int}, R) =>
(case atom_for_int (card_of_rep R, offset_of_type ofs int_T) j of
~1 => if is_opt_rep R then Kodkod.None
- else raise NUT ("NitpickKodkod.to_r (Num)", [u])
+ else raise NUT ("Nitpick_Kodkod.to_r (Num)", [u])
| j' => Kodkod.Atom j')
| Cst (Num j, T, R) =>
if j < card_of_rep R then Kodkod.Atom (j + offset_of_type ofs T)
else if is_opt_rep R then Kodkod.None
- else raise NUT ("NitpickKodkod.to_r", [u])
+ else raise NUT ("Nitpick_Kodkod.to_r", [u])
| Cst (Unknown, _, R) => empty_rel_for_rep R
| Cst (Unrep, _, R) => empty_rel_for_rep R
| Cst (Suc, T, Func (Atom x, _)) =>
@@ -1177,7 +1178,7 @@
(kk_product (Kodkod.AtomSeq (max_int_for_card int_k + 1, nat_j0))
Kodkod.Univ)
else
- raise BAD ("NitpickKodkod.to_r (NatToInt)", "\"nat_j0 <> int_j0\"")
+ raise BAD ("Nitpick_Kodkod.to_r (NatToInt)", "\"nat_j0 <> int_j0\"")
| Cst (IntToNat, _, Func (Atom (int_k, int_j0), nat_R)) =>
let
val abs_card = max_int_for_card int_k + 1
@@ -1192,7 +1193,7 @@
(kk_product (Kodkod.AtomSeq (overlap, int_j0))
Kodkod.Univ))
else
- raise BAD ("NitpickKodkod.to_r (IntToNat)", "\"nat_j0 <> int_j0\"")
+ raise BAD ("Nitpick_Kodkod.to_r (IntToNat)", "\"nat_j0 <> int_j0\"")
end
| Op1 (Not, _, R, u1) => kk_not3 (to_rep R u1)
| Op1 (Finite, _, Opt (Atom _), _) => Kodkod.None
@@ -1204,10 +1205,10 @@
Func (Struct [R1, R2], _) => (R1, R2)
| Func (R1, _) =>
if card_of_rep R1 <> 1 then
- raise REP ("NitpickKodkod.to_r (Converse)", [R])
+ raise REP ("Nitpick_Kodkod.to_r (Converse)", [R])
else
pairself (Atom o pair 1 o offset_of_type ofs) (b_T, a_T)
- | _ => raise REP ("NitpickKodkod.to_r (Converse)", [R])
+ | _ => raise REP ("Nitpick_Kodkod.to_r (Converse)", [R])
val body_R = body_rep R
val a_arity = arity_of_rep a_R
val b_arity = arity_of_rep b_R
@@ -1257,7 +1258,7 @@
(rel_expr_to_func kk R1 bool_atom_R
(Func (R1, Formula Neut)) r))
(to_opt R1 u1)
- | _ => raise NUT ("NitpickKodkod.to_r (SingletonSet)", [u]))
+ | _ => raise NUT ("Nitpick_Kodkod.to_r (SingletonSet)", [u]))
| Op1 (Tha, T, R, u1) =>
if is_opt_rep R then
kk_join (to_rep (Func (unopt_rep R, Opt bool_atom_R)) u1) true_atom
@@ -1384,7 +1385,7 @@
kk_product (kk_join (do_nut r a_R b_R u2)
(do_nut r b_R c_R u1)) r
in kk_union (do_term true_atom) (do_term false_atom) end
- | _ => raise NUT ("NitpickKodkod.to_r (Composition)", [u]))
+ | _ => raise NUT ("Nitpick_Kodkod.to_r (Composition)", [u]))
|> rel_expr_from_rel_expr kk R (Func (Struct [a_R, c_R], body_R))
end
| Op2 (Product, T, R, u1, u2) =>
@@ -1408,7 +1409,7 @@
fun do_term r =
kk_product (kk_product (do_nut r a_R u1) (do_nut r b_R u2)) r
in kk_union (do_term true_atom) (do_term false_atom) end
- | _ => raise NUT ("NitpickKodkod.to_r (Product)", [u]))
+ | _ => raise NUT ("Nitpick_Kodkod.to_r (Product)", [u]))
|> rel_expr_from_rel_expr kk R (Func (Struct [a_R, b_R], body_R))
end
| Op2 (Image, T, R, u1, u2) =>
@@ -1437,8 +1438,8 @@
rel_expr_from_rel_expr kk R core_R core_r
end
else
- raise NUT ("NitpickKodkod.to_r (Image)", [u1, u2])
- | _ => raise NUT ("NitpickKodkod.to_r (Image)", [u1, u2]))
+ raise NUT ("Nitpick_Kodkod.to_r (Image)", [u1, u2])
+ | _ => raise NUT ("Nitpick_Kodkod.to_r (Image)", [u1, u2]))
| Op2 (Apply, @{typ nat}, _,
Op2 (Apply, _, _, Cst (Subtract, _, _), u1), u2) =>
if is_Cst Unrep u2 andalso not (is_opt_rep (rep_of u1)) then
@@ -1492,7 +1493,7 @@
| us' =>
kk_rel_if (kk_some (fold1 kk_product (map to_r us')))
(Kodkod.Atom j0) Kodkod.None)
- | _ => raise NUT ("NitpickKodkod.to_r (Tuple)", [u]))
+ | _ => raise NUT ("Nitpick_Kodkod.to_r (Tuple)", [u]))
| Construct ([u'], _, _, []) => to_r u'
| Construct (_ :: sel_us, T, R, arg_us) =>
let
@@ -1516,23 +1517,23 @@
| BoundRel (x, _, _, _) => Kodkod.Var x
| FreeRel (x, _, _, _) => Kodkod.Rel x
| RelReg (j, _, R) => Kodkod.RelReg (arity_of_rep R, j)
- | u => raise NUT ("NitpickKodkod.to_r", [u])
+ | u => raise NUT ("Nitpick_Kodkod.to_r", [u])
(* nut -> Kodkod.decl *)
and to_decl (BoundRel (x, _, R, _)) =
Kodkod.DeclOne (x, Kodkod.AtomSeq (the_single (atom_schema_of_rep R)))
- | to_decl u = raise NUT ("NitpickKodkod.to_decl", [u])
+ | to_decl u = raise NUT ("Nitpick_Kodkod.to_decl", [u])
(* nut -> Kodkod.expr_assign *)
and to_expr_assign (FormulaReg (j, _, R)) u =
Kodkod.AssignFormulaReg (j, to_f u)
| to_expr_assign (RelReg (j, _, R)) u =
Kodkod.AssignRelReg ((arity_of_rep R, j), to_r u)
- | to_expr_assign u1 _ = raise NUT ("NitpickKodkod.to_expr_assign", [u1])
+ | to_expr_assign u1 _ = raise NUT ("Nitpick_Kodkod.to_expr_assign", [u1])
(* int * int -> nut -> Kodkod.rel_expr *)
and to_atom (x as (k, j0)) u =
case rep_of u of
Formula _ => atom_from_formula kk j0 (to_f u)
| Unit => if k = 1 then Kodkod.Atom j0
- else raise NUT ("NitpickKodkod.to_atom", [u])
+ else raise NUT ("Nitpick_Kodkod.to_atom", [u])
| R => atom_from_rel_expr kk x R (to_r u)
(* rep list -> nut -> Kodkod.rel_expr *)
and to_struct Rs u =
@@ -1563,7 +1564,7 @@
| to_rep (Vect (k, R)) u = to_vect k R u
| to_rep (Func (R1, R2)) u = to_func R1 R2 u
| to_rep (Opt R) u = to_opt R u
- | to_rep R _ = raise REP ("NitpickKodkod.to_rep", [R])
+ | to_rep R _ = raise REP ("Nitpick_Kodkod.to_rep", [R])
(* nut -> Kodkod.rel_expr *)
and to_integer u = to_opt (one_rep ofs (type_of u) (rep_of u)) u
(* nut list -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *)
@@ -1593,7 +1594,7 @@
(* rep list -> nut list -> Kodkod.rel_expr *)
and to_product Rs us =
case map (uncurry to_opt) (filter (not_equal Unit o fst) (Rs ~~ us)) of
- [] => raise REP ("NitpickKodkod.to_product", Rs)
+ [] => raise REP ("Nitpick_Kodkod.to_product", Rs)
| rs => fold1 kk_product rs
(* int -> typ -> rep -> nut -> Kodkod.rel_expr *)
and to_nth_pair_sel n res_T res_R u =
@@ -1639,7 +1640,7 @@
connective (formula_from_atom j0 r1) (formula_from_atom j0 r2)
| Func (R1, Atom _) => set_oper (kk_join r1 true_atom)
(kk_join r2 true_atom)
- | _ => raise REP ("NitpickKodkod.to_set_bool_op", [min_R])
+ | _ => raise REP ("Nitpick_Kodkod.to_set_bool_op", [min_R])
end
(* (Kodkod.formula -> Kodkod.formula -> Kodkod.formula)
-> (Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.rel_expr)
@@ -1676,7 +1677,7 @@
neg_second)))
false_atom))
r1 r2
- | _ => raise REP ("NitpickKodkod.to_set_op", [min_R]))
+ | _ => raise REP ("Nitpick_Kodkod.to_set_op", [min_R]))
end
(* rep -> rep -> Kodkod.rel_expr -> nut -> Kodkod.rel_expr *)
and to_apply res_R func_u arg_u =
@@ -1713,7 +1714,7 @@
rel_expr_from_rel_expr kk res_R R2
(kk_n_fold_join kk true R1 R2 (to_opt R1 arg_u) (to_r func_u))
|> body_rep R2 = Formula Neut ? to_guard [arg_u] res_R
- | _ => raise NUT ("NitpickKodkod.to_apply", [func_u])
+ | _ => raise NUT ("Nitpick_Kodkod.to_apply", [func_u])
(* int -> rep -> rep -> Kodkod.rel_expr -> nut *)
and to_apply_vect k R' res_R func_r arg_u =
let
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML Wed Oct 28 11:43:06 2009 +0000
@@ -7,9 +7,9 @@
signature NITPICK_MODEL =
sig
- type scope = NitpickScope.scope
- type rep = NitpickRep.rep
- type nut = NitpickNut.nut
+ type scope = Nitpick_Scope.scope
+ type rep = Nitpick_Rep.rep
+ type nut = Nitpick_Nut.nut
type params = {
show_skolems: bool,
@@ -29,15 +29,15 @@
-> Kodkod.raw_bound list -> term -> bool option
end;
-structure NitpickModel : NITPICK_MODEL =
+structure Nitpick_Model : NITPICK_MODEL =
struct
-open NitpickUtil
-open NitpickHOL
-open NitpickScope
-open NitpickPeephole
-open NitpickRep
-open NitpickNut
+open Nitpick_Util
+open Nitpick_HOL
+open Nitpick_Scope
+open Nitpick_Peephole
+open Nitpick_Rep
+open Nitpick_Nut
type params = {
show_skolems: bool,
@@ -57,7 +57,7 @@
prefix ^ substring (short_name s, 0, 1) ^ nat_subscript (j + 1)
| atom_name prefix (T as TFree (s, _)) j =
prefix ^ perhaps (try (unprefix "'")) s ^ nat_subscript (j + 1)
- | atom_name _ T _ = raise TYPE ("NitpickModel.atom_name", [T], [])
+ | atom_name _ T _ = raise TYPE ("Nitpick_Model.atom_name", [T], [])
(* bool -> typ -> int -> term *)
fun atom for_auto T j =
if for_auto then
@@ -130,7 +130,7 @@
fun aux (Const (s, _)) = (s <> non_opt_name, ([], []))
| aux (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) =
let val (s, (ts1, ts2)) = aux t0 in (s, (t1 :: ts1, t2 :: ts2)) end
- | aux t = raise TERM ("NitpickModel.dest_plain_fun", [t])
+ | aux t = raise TERM ("Nitpick_Model.dest_plain_fun", [t])
in apsnd (pairself rev) o aux end
(* typ -> term -> term list * term *)
@@ -138,7 +138,7 @@
(Const (@{const_name Pair}, _) $ t1 $ t2) =
break_in_two T2 t2 |>> cons t1
| break_in_two _ (Const (@{const_name Pair}, _) $ t1 $ t2) = ([t1], t2)
- | break_in_two _ t = raise TERM ("NitpickModel.break_in_two", [t])
+ | break_in_two _ t = raise TERM ("Nitpick_Model.break_in_two", [t])
(* typ -> term -> term -> term *)
fun pair_up (Type ("*", [T1', T2']))
(t1 as Const (@{const_name Pair},
@@ -180,7 +180,7 @@
[T1' --> T2', T1', T2'] ---> T1' --> T2')
$ do_arrow T1' T2' T1 T2 t0 $ do_term T1' T1 t1 $ do_term T2' T2 t2
| do_arrow _ _ _ _ t =
- raise TERM ("NitpickModel.typecast_fun.do_arrow", [t])
+ raise TERM ("Nitpick_Model.typecast_fun.do_arrow", [t])
and do_fun T1' T2' T1 T2 t =
case factor_out_types T1' T1 of
((_, NONE), (_, NONE)) => t |> do_arrow T1' T2' T1 T2
@@ -188,7 +188,7 @@
t |> do_curry T1a T1b T2 |> do_arrow T1' T2' T1a (T1b --> T2)
| ((T1a', SOME T1b'), (_, NONE)) =>
t |> do_arrow T1a' (T1b' --> T2') T1 T2 |> do_uncurry T1' T2'
- | _ => raise TYPE ("NitpickModel.typecast_fun.do_fun", [T1, T1'], [])
+ | _ => raise TYPE ("Nitpick_Model.typecast_fun.do_fun", [T1, T1'], [])
(* typ -> typ -> term -> term *)
and do_term (Type ("fun", [T1', T2'])) (Type ("fun", [T1, T2])) t =
do_fun T1' T2' T1 T2 t
@@ -198,9 +198,9 @@
$ do_term T1' T1 t1 $ do_term T2' T2 t2
| do_term T' T t =
if T = T' then t
- else raise TYPE ("NitpickModel.typecast_fun.do_term", [T, T'], [])
+ else raise TYPE ("Nitpick_Model.typecast_fun.do_term", [T, T'], [])
in if T1' = T1 andalso T2' = T2 then t else do_fun T1' T2' T1 T2 t end
- | typecast_fun T' _ _ _ = raise TYPE ("NitpickModel.typecast_fun", [T'], [])
+ | typecast_fun T' _ _ _ = raise TYPE ("Nitpick_Model.typecast_fun", [T'], [])
(* term -> string *)
fun truth_const_sort_key @{const True} = "0"
@@ -302,7 +302,7 @@
term_for_rep seen T T' (Vect (k1, Atom (k2, 0)))
[nth_combination (replicate k1 (k2, 0)) j]
handle General.Subscript =>
- raise ARG ("NitpickModel.reconstruct_term.term_for_atom",
+ raise ARG ("Nitpick_Model.reconstruct_term.term_for_atom",
signed_string_of_int j ^ " for " ^
string_for_rep (Vect (k1, Atom (k2, 0))))
end
@@ -350,7 +350,7 @@
map (fn x => get_first
(fn ConstName (s', T', R) =>
if (s', T') = x then SOME R else NONE
- | u => raise NUT ("NitpickModel.reconstruct_\
+ | u => raise NUT ("Nitpick_Model.reconstruct_\
\term.term_for_atom", [u]))
sel_names |> the) sel_xs
val arg_Rs = map (snd o dest_Func) sel_Rs
@@ -389,7 +389,7 @@
| n2 => Const (@{const_name HOL.divide},
[num_T, num_T] ---> num_T)
$ mk_num n1 $ mk_num n2)
- | _ => raise TERM ("NitpickModel.reconstruct_term.term_\
+ | _ => raise TERM ("Nitpick_Model.reconstruct_term.term_\
\for_atom (Abs_Frac)", ts)
end
else if not for_auto andalso is_abs_fun thy constr_x then
@@ -421,7 +421,7 @@
and term_for_rep seen T T' Unit [[]] = term_for_atom seen T T' 0
| term_for_rep seen T T' (R as Atom (k, j0)) [[j]] =
if j >= j0 andalso j < j0 + k then term_for_atom seen T T' (j - j0)
- else raise REP ("NitpickModel.reconstruct_term.term_for_rep", [R])
+ else raise REP ("Nitpick_Model.reconstruct_term.term_for_rep", [R])
| term_for_rep seen (Type ("*", [T1, T2])) _ (Struct [R1, R2]) [js] =
let
val arity1 = arity_of_rep R1
@@ -454,7 +454,7 @@
| term_for_rep seen T T' (Opt R) jss =
if null jss then Const (unknown, T) else term_for_rep seen T T' R jss
| term_for_rep seen T _ R jss =
- raise ARG ("NitpickModel.reconstruct_term.term_for_rep",
+ raise ARG ("Nitpick_Model.reconstruct_term.term_for_rep",
Refute.string_of_typ T ^ " " ^ string_for_rep R ^ " " ^
string_of_int (length jss))
in
@@ -576,7 +576,7 @@
(assign_operator_for_const (s, T),
user_friendly_const ext_ctxt (base_name, step_name) formats (s, T),
T)
- | _ => raise NUT ("NitpickModel.reconstruct_hol_model.\
+ | _ => raise NUT ("Nitpick_Model.reconstruct_hol_model.\
\pretty_for_assign", [name])
val t2 = if rep_of name = Any then
Const (@{const_name undefined}, T')
@@ -601,7 +601,7 @@
fun integer_datatype T =
[{typ = T, card = card_of_type card_assigns T, co = false,
precise = false, constrs = []}]
- handle TYPE ("NitpickHOL.card_of_type", _, _) => []
+ handle TYPE ("Nitpick_HOL.card_of_type", _, _) => []
val (codatatypes, datatypes) =
List.partition #co datatypes
||> append (integer_datatype nat_T @ integer_datatype int_T)
@@ -637,7 +637,7 @@
free_names of
[name] => name
| [] => FreeName (s, T, Any)
- | _ => raise TERM ("NitpickModel.reconstruct_hol_model",
+ | _ => raise TERM ("Nitpick_Model.reconstruct_hol_model",
[Const x])) all_frees
val chunks = block_of_names true "Free variable" free_names @
block_of_names show_skolems "Skolem constant" skolem_names @
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Wed Oct 28 11:43:06 2009 +0000
@@ -7,17 +7,17 @@
signature NITPICK_MONO =
sig
- type extended_context = NitpickHOL.extended_context
+ type extended_context = Nitpick_HOL.extended_context
val formulas_monotonic :
extended_context -> typ -> term list -> term list -> term -> bool
end;
-structure NitpickMono : NITPICK_MONO =
+structure Nitpick_Mono : NITPICK_MONO =
struct
-open NitpickUtil
-open NitpickHOL
+open Nitpick_Util
+open Nitpick_HOL
type var = int
@@ -363,11 +363,11 @@
accum =
(accum |> fold (uncurry (do_ctype_comp cmp xs)) [(C11, C21), (C12, C22)]
handle Library.UnequalLengths =>
- raise CTYPE ("NitpickMono.do_ctype_comp", [C1, C2]))
+ raise CTYPE ("Nitpick_Mono.do_ctype_comp", [C1, C2]))
| do_ctype_comp cmp xs (CType _) (CType _) accum =
accum (* no need to compare them thanks to the cache *)
| do_ctype_comp _ _ C1 C2 _ =
- raise CTYPE ("NitpickMono.do_ctype_comp", [C1, C2])
+ raise CTYPE ("Nitpick_Mono.do_ctype_comp", [C1, C2])
(* comp_op -> ctype -> ctype -> constraint_set -> constraint_set *)
fun add_ctype_comp _ _ _ UnsolvableCSet = UnsolvableCSet
@@ -413,7 +413,7 @@
| do_notin_ctype_fv sn sexp (CType (_, Cs)) accum =
accum |> fold (do_notin_ctype_fv sn sexp) Cs
| do_notin_ctype_fv _ _ C _ =
- raise CTYPE ("NitpickMono.do_notin_ctype_fv", [C])
+ raise CTYPE ("Nitpick_Mono.do_notin_ctype_fv", [C])
(* sign -> ctype -> constraint_set -> constraint_set *)
fun add_notin_ctype_fv _ _ UnsolvableCSet = UnsolvableCSet
@@ -584,13 +584,13 @@
case ctype_for (nth_range_type 2 T) of
C as CPair (a_C, b_C) =>
(CFun (a_C, S Neg, CFun (b_C, S Neg, C)), accum)
- | C => raise CTYPE ("NitpickMono.consider_term.do_pair_constr", [C])
+ | C => raise CTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [C])
(* int -> typ -> accumulator -> ctype * accumulator *)
fun do_nth_pair_sel n T =
case ctype_for (domain_type T) of
C as CPair (a_C, b_C) =>
pair (CFun (C, S Neg, if n = 0 then a_C else b_C))
- | C => raise CTYPE ("NitpickMono.consider_term.do_nth_pair_sel", [C])
+ | C => raise CTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [C])
val unsolvable = (CType ("unsolvable", []), unsolvable_accum)
(* typ -> term -> accumulator -> ctype * accumulator *)
fun do_bounded_quantifier abs_T bound_t body_t accum =
@@ -770,7 +770,7 @@
| _ => case C1 of
CFun (C11, _, C12) =>
(C12, accum ||> add_is_sub_ctype C2 C11)
- | _ => raise CTYPE ("NitpickMono.consider_term.do_term \
+ | _ => raise CTYPE ("Nitpick_Mono.consider_term.do_term \
\(op $)", [C1])
end)
|> tap (fn (C, _) =>
@@ -906,7 +906,7 @@
| Const (@{const_name "op ="}, _) $ t1 $ t2 => do_equals t1 t2 accum
| @{const "op &"} $ t1 $ t2 => accum |> do_formula t1 |> do_formula t2
| @{const "op -->"} $ t1 $ t2 => do_implies t1 t2 accum
- | _ => raise TERM ("NitpickMono.consider_definitional_axiom.\
+ | _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\
\do_formula", [t])
in do_formula t end
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML Wed Oct 28 11:43:06 2009 +0000
@@ -7,11 +7,11 @@
signature NITPICK_NUT =
sig
- type special_fun = NitpickHOL.special_fun
- type extended_context = NitpickHOL.extended_context
- type scope = NitpickScope.scope
- type name_pool = NitpickPeephole.name_pool
- type rep = NitpickRep.rep
+ type special_fun = Nitpick_HOL.special_fun
+ type extended_context = Nitpick_HOL.extended_context
+ type scope = Nitpick_Scope.scope
+ type name_pool = Nitpick_Peephole.name_pool
+ type rep = Nitpick_Rep.rep
datatype cst =
Unity |
@@ -121,14 +121,14 @@
val rename_vars_in_nut : name_pool -> nut NameTable.table -> nut -> nut
end;
-structure NitpickNut : NITPICK_NUT =
+structure Nitpick_Nut : NITPICK_NUT =
struct
-open NitpickUtil
-open NitpickHOL
-open NitpickScope
-open NitpickPeephole
-open NitpickRep
+open Nitpick_Util
+open Nitpick_HOL
+open Nitpick_Scope
+open Nitpick_Peephole
+open Nitpick_Rep
datatype cst =
Unity |
@@ -357,16 +357,16 @@
| nickname_of (ConstName (s, _, _)) = s
| nickname_of (BoundRel (_, _, _, nick)) = nick
| nickname_of (FreeRel (_, _, _, nick)) = nick
- | nickname_of u = raise NUT ("NitpickNut.nickname_of", [u])
+ | nickname_of u = raise NUT ("Nitpick_Nut.nickname_of", [u])
(* nut -> bool *)
fun is_skolem_name u =
space_explode name_sep (nickname_of u)
|> exists (String.isPrefix skolem_prefix)
- handle NUT ("NitpickNut.nickname_of", _) => false
+ handle NUT ("Nitpick_Nut.nickname_of", _) => false
fun is_eval_name u =
String.isPrefix eval_prefix (nickname_of u)
- handle NUT ("NitpickNut.nickname_of", _) => false
+ handle NUT ("Nitpick_Nut.nickname_of", _) => false
(* cst -> nut -> bool *)
fun is_Cst cst (Cst (cst', _, _)) = (cst = cst')
| is_Cst _ _ = false
@@ -407,7 +407,7 @@
(case fast_string_ord (s1, s2) of
EQUAL => TermOrd.typ_ord (T1, T2)
| ord => ord)
- | name_ord (u1, u2) = raise NUT ("NitpickNut.name_ord", [u1, u2])
+ | name_ord (u1, u2) = raise NUT ("Nitpick_Nut.name_ord", [u1, u2])
(* nut -> nut -> int *)
fun num_occs_in_nut needle_u stack_u =
@@ -430,7 +430,7 @@
fun modify_name_rep (BoundName (j, T, _, nick)) R = BoundName (j, T, R, nick)
| modify_name_rep (FreeName (s, T, _)) R = FreeName (s, T, R)
| modify_name_rep (ConstName (s, T, _)) R = ConstName (s, T, R)
- | modify_name_rep u _ = raise NUT ("NitpickNut.modify_name_rep", [u])
+ | modify_name_rep u _ = raise NUT ("Nitpick_Nut.modify_name_rep", [u])
structure NameTable = Table(type key = nut val ord = name_ord)
@@ -438,12 +438,12 @@
fun the_name table name =
case NameTable.lookup table name of
SOME u => u
- | NONE => raise NUT ("NitpickNut.the_name", [name])
+ | NONE => raise NUT ("Nitpick_Nut.the_name", [name])
(* nut NameTable.table -> nut -> Kodkod.n_ary_index *)
fun the_rel table name =
case the_name table name of
FreeRel (x, _, _, _) => x
- | u => raise NUT ("NitpickNut.the_rel", [u])
+ | u => raise NUT ("Nitpick_Nut.the_rel", [u])
(* typ * term -> typ * term *)
fun mk_fst (_, Const (@{const_name Pair}, T) $ t1 $ _) = (domain_type T, t1)
@@ -669,13 +669,13 @@
(case arity_of_built_in_const fast_descrs x of
SOME n =>
(case n - length ts of
- 0 => raise TERM ("NitpickNut.nut_from_term.aux", [t])
+ 0 => raise TERM ("Nitpick_Nut.nut_from_term.aux", [t])
| k => if k > 0 then sub (eta_expand Ts t k)
else do_apply t0 ts)
| NONE => if null ts then ConstName (s, T, Any)
else do_apply t0 ts)
| (Free (s, T), []) => FreeName (s, T, Any)
- | (Var _, []) => raise TERM ("NitpickNut.nut_from_term", [t])
+ | (Var _, []) => raise TERM ("Nitpick_Nut.nut_from_term", [t])
| (Bound j, []) =>
BoundName (length Ts - j - 1, nth Ts j, Any, nth ss j)
| (Abs (s, T, t1), []) =>
@@ -1106,7 +1106,7 @@
else
unopt_rep R
in s_op2 Lambda T R u1' u2' end
- | R => raise NUT ("NitpickNut.aux.choose_reps_in_nut", [u]))
+ | R => raise NUT ("Nitpick_Nut.aux.choose_reps_in_nut", [u]))
| Op2 (oper, T, _, u1, u2) =>
if oper mem [All, Exist] then
let
@@ -1274,9 +1274,9 @@
| shape_tuple (T as Type ("fun", [_, T2])) (R as Vect (k, R')) us =
Tuple (T, R, map (shape_tuple T2 R') (batch_list (length us div k) us))
| shape_tuple T R [u] =
- if type_of u = T then u else raise NUT ("NitpickNut.shape_tuple", [u])
+ if type_of u = T then u else raise NUT ("Nitpick_Nut.shape_tuple", [u])
| shape_tuple T Unit [] = Cst (Unity, T, Unit)
- | shape_tuple _ _ us = raise NUT ("NitpickNut.shape_tuple", us)
+ | shape_tuple _ _ us = raise NUT ("Nitpick_Nut.shape_tuple", us)
(* bool -> nut -> nut list * name_pool * nut NameTable.table
-> nut list * name_pool * nut NameTable.table *)
--- a/src/HOL/Tools/Nitpick/nitpick_peephole.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_peephole.ML Wed Oct 28 11:43:06 2009 +0000
@@ -86,11 +86,11 @@
val kodkod_constrs : bool -> int -> int -> int -> kodkod_constrs
end;
-structure NitpickPeephole : NITPICK_PEEPHOLE =
+structure Nitpick_Peephole : NITPICK_PEEPHOLE =
struct
open Kodkod
-open NitpickUtil
+open Nitpick_Util
type name_pool = {
rels: n_ary_index list,
@@ -188,20 +188,20 @@
if is_none_product r1 orelse is_none_product r2 then SOME false else NONE
(* int -> rel_expr *)
-fun empty_n_ary_rel 0 = raise ARG ("NitpickPeephole.empty_n_ary_rel", "0")
+fun empty_n_ary_rel 0 = raise ARG ("Nitpick_Peephole.empty_n_ary_rel", "0")
| empty_n_ary_rel n = funpow (n - 1) (curry Product None) None
(* decl -> rel_expr *)
fun decl_one_set (DeclOne (_, r)) = r
| decl_one_set _ =
- raise ARG ("NitpickPeephole.decl_one_set", "not \"DeclOne\"")
+ raise ARG ("Nitpick_Peephole.decl_one_set", "not \"DeclOne\"")
(* int_expr -> bool *)
fun is_Num (Num _) = true
| is_Num _ = false
(* int_expr -> int *)
fun dest_Num (Num k) = k
- | dest_Num _ = raise ARG ("NitpickPeephole.dest_Num", "not \"Num\"")
+ | dest_Num _ = raise ARG ("Nitpick_Peephole.dest_Num", "not \"Num\"")
(* int -> int -> int_expr list *)
fun num_seq j0 n = map Num (index_seq j0 n)
--- a/src/HOL/Tools/Nitpick/nitpick_rep.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_rep.ML Wed Oct 28 11:43:06 2009 +0000
@@ -7,8 +7,8 @@
signature NITPICK_REP =
sig
- type polarity = NitpickUtil.polarity
- type scope = NitpickScope.scope
+ type polarity = Nitpick_Util.polarity
+ type scope = Nitpick_Scope.scope
datatype rep =
Any |
@@ -58,12 +58,12 @@
val all_combinations_for_reps : rep list -> int list list
end;
-structure NitpickRep : NITPICK_REP =
+structure Nitpick_Rep : NITPICK_REP =
struct
-open NitpickUtil
-open NitpickHOL
-open NitpickScope
+open Nitpick_Util
+open Nitpick_HOL
+open Nitpick_Scope
datatype rep =
Any |
@@ -111,7 +111,7 @@
| is_opt_rep _ = false
(* rep -> int *)
-fun card_of_rep Any = raise REP ("NitpickRep.card_of_rep", [Any])
+fun card_of_rep Any = raise REP ("Nitpick_Rep.card_of_rep", [Any])
| card_of_rep (Formula _) = 2
| card_of_rep Unit = 1
| card_of_rep (Atom (k, _)) = k
@@ -120,7 +120,7 @@
| card_of_rep (Func (R1, R2)) =
reasonable_power (card_of_rep R2) (card_of_rep R1)
| card_of_rep (Opt R) = card_of_rep R
-fun arity_of_rep Any = raise REP ("NitpickRep.arity_of_rep", [Any])
+fun arity_of_rep Any = raise REP ("Nitpick_Rep.arity_of_rep", [Any])
| arity_of_rep (Formula _) = 0
| arity_of_rep Unit = 0
| arity_of_rep (Atom _) = 1
@@ -129,7 +129,7 @@
| arity_of_rep (Func (R1, R2)) = arity_of_rep R1 + arity_of_rep R2
| arity_of_rep (Opt R) = arity_of_rep R
fun min_univ_card_of_rep Any =
- raise REP ("NitpickRep.min_univ_card_of_rep", [Any])
+ raise REP ("Nitpick_Rep.min_univ_card_of_rep", [Any])
| min_univ_card_of_rep (Formula _) = 0
| min_univ_card_of_rep Unit = 0
| min_univ_card_of_rep (Atom (k, j0)) = k + j0 + 1
@@ -151,7 +151,7 @@
(* rep -> rep * rep *)
fun dest_Func (Func z) = z
- | dest_Func R = raise REP ("NitpickRep.dest_Func", [R])
+ | dest_Func R = raise REP ("Nitpick_Rep.dest_Func", [R])
(* int Typtab.table -> typ -> (unit -> int) -> rep -> rep *)
fun smart_range_rep _ _ _ Unit = Unit
| smart_range_rep _ _ _ (Vect (_, R)) = R
@@ -162,7 +162,7 @@
Atom (1, offset_of_type ofs T2)
| smart_range_rep ofs (Type ("fun", [_, T2])) ran_card (Atom _) =
Atom (ran_card (), offset_of_type ofs T2)
- | smart_range_rep _ _ _ R = raise REP ("NitpickRep.smart_range_rep", [R])
+ | smart_range_rep _ _ _ R = raise REP ("Nitpick_Rep.smart_range_rep", [R])
(* rep -> rep list *)
fun binder_reps (Func (R1, R2)) = R1 :: binder_reps R2
@@ -177,7 +177,7 @@
| flip_rep_polarity R = R
(* int Typtab.table -> rep -> rep *)
-fun one_rep _ _ Any = raise REP ("NitpickRep.one_rep", [Any])
+fun one_rep _ _ Any = raise REP ("Nitpick_Rep.one_rep", [Any])
| one_rep _ _ (Atom x) = Atom x
| one_rep _ _ (Struct Rs) = Struct Rs
| one_rep _ _ (Vect z) = Vect z
@@ -203,7 +203,7 @@
else if polar2 = Neut then
polar1
else
- raise ARG ("NitpickRep.min_polarity",
+ raise ARG ("Nitpick_Rep.min_polarity",
commas (map (quote o string_for_polarity) [polar1, polar2]))
(* It's important that Func is before Vect, because if the range is Opt we
@@ -236,7 +236,7 @@
if k1 < k2 then Vect (k1, R1)
else if k1 > k2 then Vect (k2, R2)
else Vect (k1, min_rep R1 R2)
- | min_rep R1 R2 = raise REP ("NitpickRep.min_rep", [R1, R2])
+ | min_rep R1 R2 = raise REP ("Nitpick_Rep.min_rep", [R1, R2])
(* rep list -> rep list -> rep list *)
and min_reps [] _ = []
| min_reps _ [] = []
@@ -253,7 +253,7 @@
| Vect (k, _) => k
| Func (R1, _) => card_of_rep R1
| Opt R => card_of_domain_from_rep ran_card R
- | _ => raise REP ("NitpickRep.card_of_domain_from_rep", [R])
+ | _ => raise REP ("Nitpick_Rep.card_of_domain_from_rep", [R])
(* int Typtab.table -> typ -> rep -> rep *)
fun rep_to_binary_rel_rep ofs T R =
@@ -306,10 +306,10 @@
(optable_rep ofs T1 (best_one_rep_for_type scope T1),
optable_rep ofs T2 (best_one_rep_for_type scope T2))
| best_non_opt_symmetric_reps_for_fun_type _ T =
- raise TYPE ("NitpickRep.best_non_opt_symmetric_reps_for_fun_type", [T], [])
+ raise TYPE ("Nitpick_Rep.best_non_opt_symmetric_reps_for_fun_type", [T], [])
(* rep -> (int * int) list *)
-fun atom_schema_of_rep Any = raise REP ("NitpickRep.atom_schema_of_rep", [Any])
+fun atom_schema_of_rep Any = raise REP ("Nitpick_Rep.atom_schema_of_rep", [Any])
| atom_schema_of_rep (Formula _) = []
| atom_schema_of_rep Unit = []
| atom_schema_of_rep (Atom x) = [x]
@@ -332,7 +332,7 @@
| type_schema_of_rep (Type ("fun", [T1, T2])) (Func (R1, R2)) =
type_schema_of_rep T1 R1 @ type_schema_of_rep T2 R2
| type_schema_of_rep T (Opt R) = type_schema_of_rep T R
- | type_schema_of_rep T R = raise REP ("NitpickRep.type_schema_of_rep", [R])
+ | type_schema_of_rep T R = raise REP ("Nitpick_Rep.type_schema_of_rep", [R])
(* typ list -> rep list -> typ list *)
and type_schema_of_reps Ts Rs = flat (map2 type_schema_of_rep Ts Rs)
--- a/src/HOL/Tools/Nitpick/nitpick_scope.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_scope.ML Wed Oct 28 11:43:06 2009 +0000
@@ -7,7 +7,7 @@
signature NITPICK_SCOPE =
sig
- type extended_context = NitpickHOL.extended_context
+ type extended_context = Nitpick_HOL.extended_context
type constr_spec = {
const: styp,
@@ -47,11 +47,11 @@
-> int list -> typ list -> typ list -> scope list
end;
-structure NitpickScope : NITPICK_SCOPE =
+structure Nitpick_Scope : NITPICK_SCOPE =
struct
-open NitpickUtil
-open NitpickHOL
+open Nitpick_Util
+open Nitpick_HOL
type constr_spec = {
const: styp,
@@ -85,7 +85,7 @@
List.find (equal T o #typ) dtypes
(* dtype_spec list -> styp -> constr_spec *)
-fun constr_spec [] x = raise TERM ("NitpickScope.constr_spec", [Const x])
+fun constr_spec [] x = raise TERM ("Nitpick_Scope.constr_spec", [Const x])
| constr_spec ({constrs, ...} :: dtypes : dtype_spec list) (x as (s, T)) =
case List.find (equal (s, body_type T) o (apsnd body_type o #const))
constrs of
@@ -115,7 +115,7 @@
(* scope -> typ -> int * int *)
fun spec_of_type ({card_assigns, ofs, ...} : scope) T =
(card_of_type card_assigns T
- handle TYPE ("NitpickHOL.card_of_type", _, _) => ~1, offset_of_type ofs T)
+ handle TYPE ("Nitpick_HOL.card_of_type", _, _) => ~1, offset_of_type ofs T)
(* (string -> string) -> scope
-> string list * string list * string list * string list * string list *)
@@ -205,17 +205,17 @@
fun lookup_ints_assign eq asgns key =
case triple_lookup eq asgns key of
SOME ks => ks
- | NONE => raise ARG ("NitpickScope.lookup_ints_assign", "")
+ | NONE => raise ARG ("Nitpick_Scope.lookup_ints_assign", "")
(* theory -> (typ option * int list) list -> typ -> int list *)
fun lookup_type_ints_assign thy asgns T =
map (curry Int.max 1) (lookup_ints_assign (type_match thy) asgns T)
- handle ARG ("NitpickScope.lookup_ints_assign", _) =>
- raise TYPE ("NitpickScope.lookup_type_ints_assign", [T], [])
+ handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
+ raise TYPE ("Nitpick_Scope.lookup_type_ints_assign", [T], [])
(* theory -> (styp option * int list) list -> styp -> int list *)
fun lookup_const_ints_assign thy asgns x =
lookup_ints_assign (const_match thy) asgns x
- handle ARG ("NitpickScope.lookup_ints_assign", _) =>
- raise TERM ("NitpickScope.lookup_const_ints_assign", [Const x])
+ handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
+ raise TERM ("Nitpick_Scope.lookup_const_ints_assign", [Const x])
(* theory -> (styp option * int list) list -> styp -> row option *)
fun row_for_constr thy maxes_asgns constr =
@@ -315,7 +315,7 @@
max < k
orelse (forall (not_equal 0) precise_cards andalso doms_card () < k)
end
- handle TYPE ("NitpickHOL.card_of_type", _, _) => false
+ handle TYPE ("Nitpick_HOL.card_of_type", _, _) => false
(* theory -> scope_desc -> bool *)
fun is_surely_inconsistent_scope_description thy (desc as (card_asgns, _)) =
--- a/src/HOL/Tools/Nitpick/nitpick_tests.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_tests.ML Wed Oct 28 11:43:06 2009 +0000
@@ -10,15 +10,14 @@
val run_all_tests : unit -> unit
end
-structure NitpickTests =
+structure Nitpick_Tests =
struct
-open NitpickUtil
-open NitpickPeephole
-open NitpickRep
-open NitpickNut
-open NitpickKodkod
-open Nitpick
+open Nitpick_Util
+open Nitpick_Peephole
+open Nitpick_Rep
+open Nitpick_Nut
+open Nitpick_Kodkod
val settings =
[("solver", "\"zChaff\""),
--- a/src/HOL/Tools/Nitpick/nitpick_util.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Nitpick/nitpick_util.ML Wed Oct 28 11:43:06 2009 +0000
@@ -72,7 +72,7 @@
val maybe_quote : string -> string
end
-structure NitpickUtil : NITPICK_UTIL =
+structure Nitpick_Util : NITPICK_UTIL =
struct
type styp = string * typ
@@ -107,7 +107,7 @@
| reasonable_power 1 _ = 1
| reasonable_power a b =
if b < 0 orelse b > max_exponent then
- raise LIMIT ("NitpickUtil.reasonable_power",
+ raise LIMIT ("Nitpick_Util.reasonable_power",
"too large exponent (" ^ signed_string_of_int b ^ ")")
else
let
@@ -123,7 +123,7 @@
if reasonable_power m r = n then
r
else
- raise ARG ("NitpickUtil.exact_log",
+ raise ARG ("Nitpick_Util.exact_log",
commas (map signed_string_of_int [m, n]))
end
@@ -133,7 +133,7 @@
if reasonable_power r m = n then
r
else
- raise ARG ("NitpickUtil.exact_root",
+ raise ARG ("Nitpick_Util.exact_root",
commas (map signed_string_of_int [m, n]))
end
@@ -156,7 +156,7 @@
fun aux _ [] _ = []
| aux i (j :: js) (x :: xs) =
if i = j then x :: aux (i + 1) js xs else aux (i + 1) (j :: js) xs
- | aux _ _ _ = raise ARG ("NitpickUtil.filter_indices",
+ | aux _ _ _ = raise ARG ("Nitpick_Util.filter_indices",
"indices unordered or out of range")
in aux 0 js xs end
fun filter_out_indices js xs =
@@ -165,7 +165,7 @@
fun aux _ [] xs = xs
| aux i (j :: js) (x :: xs) =
if i = j then aux (i + 1) js xs else x :: aux (i + 1) (j :: js) xs
- | aux _ _ _ = raise ARG ("NitpickUtil.filter_out_indices",
+ | aux _ _ _ = raise ARG ("Nitpick_Util.filter_out_indices",
"indices unordered or out of range")
in aux 0 js xs end
@@ -303,5 +303,5 @@
end;
-structure BasicNitpickUtil : BASIC_NITPICK_UTIL = NitpickUtil;
-open BasicNitpickUtil;
+structure Basic_Nitpick_Util : BASIC_NITPICK_UTIL = Nitpick_Util;
+open Basic_Nitpick_Util;
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_aux.ML Wed Oct 28 11:42:31 2009 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,236 +0,0 @@
-(* Author: Lukas Bulwahn, TU Muenchen
-
-Auxilary functions for predicate compiler
-*)
-
-structure Predicate_Compile_Aux =
-struct
-
-(* general syntactic functions *)
-
-(*Like dest_conj, but flattens conjunctions however nested*)
-fun conjuncts_aux (Const ("op &", _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs)
- | conjuncts_aux t conjs = t::conjs;
-
-fun conjuncts t = conjuncts_aux t [];
-
-(* syntactic functions *)
-
-fun is_equationlike_term (Const ("==", _) $ _ $ _) = true
- | is_equationlike_term (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _)) = true
- | is_equationlike_term _ = false
-
-val is_equationlike = is_equationlike_term o prop_of
-
-fun is_pred_equation_term (Const ("==", _) $ u $ v) =
- (fastype_of u = @{typ bool}) andalso (fastype_of v = @{typ bool})
- | is_pred_equation_term _ = false
-
-val is_pred_equation = is_pred_equation_term o prop_of
-
-fun is_intro_term constname t =
- case fst (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl t))) of
- Const (c, _) => c = constname
- | _ => false
-
-fun is_intro constname t = is_intro_term constname (prop_of t)
-
-fun is_pred thy constname =
- let
- val T = (Sign.the_const_type thy constname)
- in body_type T = @{typ "bool"} end;
-
-
-fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT)
- | is_predT _ = false
-
-
-(*** check if a term contains only constructor functions ***)
-fun is_constrt thy =
- let
- val cnstrs = flat (maps
- (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
- (Symtab.dest (Datatype.get_all thy)));
- fun check t = (case strip_comb t of
- (Free _, []) => true
- | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
- (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts
- | _ => false)
- | _ => false)
- in check end;
-
-fun strip_ex (Const ("Ex", _) $ Abs (x, T, t)) =
- let
- val (xTs, t') = strip_ex t
- in
- ((x, T) :: xTs, t')
- end
- | strip_ex t = ([], t)
-
-fun focus_ex t nctxt =
- let
- val ((xs, Ts), t') = apfst split_list (strip_ex t)
- val (xs', nctxt') = Name.variants xs nctxt;
- val ps' = xs' ~~ Ts;
- val vs = map Free ps';
- val t'' = Term.subst_bounds (rev vs, t');
- in ((ps', t''), nctxt') end;
-
-
-(* introduction rule combinators *)
-
-(* combinators to apply a function to all literals of an introduction rules *)
-
-fun map_atoms f intro =
- let
- val (literals, head) = Logic.strip_horn intro
- fun appl t = (case t of
- (@{term "Not"} $ t') => HOLogic.mk_not (f t')
- | _ => f t)
- in
- Logic.list_implies
- (map (HOLogic.mk_Trueprop o appl o HOLogic.dest_Trueprop) literals, head)
- end
-
-fun fold_atoms f intro s =
- let
- val (literals, head) = Logic.strip_horn intro
- fun appl t s = (case t of
- (@{term "Not"} $ t') => f t' s
- | _ => f t s)
- in fold appl (map HOLogic.dest_Trueprop literals) s end
-
-fun fold_map_atoms f intro s =
- let
- val (literals, head) = Logic.strip_horn intro
- fun appl t s = (case t of
- (@{term "Not"} $ t') => apfst HOLogic.mk_not (f t' s)
- | _ => f t s)
- val (literals', s') = fold_map appl (map HOLogic.dest_Trueprop literals) s
- in
- (Logic.list_implies (map HOLogic.mk_Trueprop literals', head), s')
- end;
-
-fun maps_premises f intro =
- let
- val (premises, head) = Logic.strip_horn intro
- in
- Logic.list_implies (maps f premises, head)
- end
-
-(* lifting term operations to theorems *)
-
-fun map_term thy f th =
- Skip_Proof.make_thm thy (f (prop_of th))
-
-(*
-fun equals_conv lhs_cv rhs_cv ct =
- case Thm.term_of ct of
- Const ("==", _) $ _ $ _ => Conv.arg_conv cv ct
- | _ => error "equals_conv"
-*)
-
-(* Different options for compiler *)
-
-datatype options = Options of {
- expected_modes : (string * int list list) option,
- show_steps : bool,
- show_mode_inference : bool,
- show_proof_trace : bool,
- show_intermediate_results : bool,
- show_compilation : bool,
- skip_proof : bool,
-
- inductify : bool,
- rpred : bool,
- depth_limited : bool
-};
-
-fun expected_modes (Options opt) = #expected_modes opt
-fun show_steps (Options opt) = #show_steps opt
-fun show_mode_inference (Options opt) = #show_mode_inference opt
-fun show_intermediate_results (Options opt) = #show_intermediate_results opt
-fun show_proof_trace (Options opt) = #show_proof_trace opt
-fun show_compilation (Options opt) = #show_compilation opt
-fun skip_proof (Options opt) = #skip_proof opt
-
-fun is_inductify (Options opt) = #inductify opt
-fun is_rpred (Options opt) = #rpred opt
-fun is_depth_limited (Options opt) = #depth_limited opt
-
-val default_options = Options {
- expected_modes = NONE,
- show_steps = false,
- show_intermediate_results = false,
- show_proof_trace = false,
- show_mode_inference = false,
- show_compilation = false,
- skip_proof = false,
-
- inductify = false,
- rpred = false,
- depth_limited = false
-}
-
-
-fun print_step options s =
- if show_steps options then tracing s else ()
-
-(* tuple processing *)
-
-fun expand_tuples thy intro =
- let
- fun rewrite_args [] (pats, intro_t, ctxt) = (pats, intro_t, ctxt)
- | rewrite_args (arg::args) (pats, intro_t, ctxt) =
- (case HOLogic.strip_tupleT (fastype_of arg) of
- (Ts as _ :: _ :: _) =>
- let
- fun rewrite_arg' (Const ("Pair", _) $ _ $ t2, Type ("*", [_, T2]))
- (args, (pats, intro_t, ctxt)) = rewrite_arg' (t2, T2) (args, (pats, intro_t, ctxt))
- | rewrite_arg' (t, Type ("*", [T1, T2])) (args, (pats, intro_t, ctxt)) =
- let
- val ([x, y], ctxt') = Variable.variant_fixes ["x", "y"] ctxt
- val pat = (t, HOLogic.mk_prod (Free (x, T1), Free (y, T2)))
- val intro_t' = Pattern.rewrite_term thy [pat] [] intro_t
- val args' = map (Pattern.rewrite_term thy [pat] []) args
- in
- rewrite_arg' (Free (y, T2), T2) (args', (pat::pats, intro_t', ctxt'))
- end
- | rewrite_arg' _ (args, (pats, intro_t, ctxt)) = (args, (pats, intro_t, ctxt))
- val (args', (pats, intro_t', ctxt')) = rewrite_arg' (arg, fastype_of arg)
- (args, (pats, intro_t, ctxt))
- in
- rewrite_args args' (pats, intro_t', ctxt')
- end
- | _ => rewrite_args args (pats, intro_t, ctxt))
- fun rewrite_prem atom =
- let
- val (_, args) = strip_comb atom
- in rewrite_args args end
- val ctxt = ProofContext.init thy
- val (((T_insts, t_insts), [intro']), ctxt1) = Variable.import false [intro] ctxt
- val intro_t = prop_of intro'
- val concl = Logic.strip_imp_concl intro_t
- val (p, args) = strip_comb (HOLogic.dest_Trueprop concl)
- val (pats', intro_t', ctxt2) = rewrite_args args ([], intro_t, ctxt1)
- val (pats', intro_t', ctxt3) =
- fold_atoms rewrite_prem intro_t' (pats', intro_t', ctxt2)
- fun rewrite_pat (ct1, ct2) =
- (ct1, cterm_of thy (Pattern.rewrite_term thy pats' [] (term_of ct2)))
- val t_insts' = map rewrite_pat t_insts
- val intro'' = Thm.instantiate (T_insts, t_insts') intro
- val [intro'''] = Variable.export ctxt3 ctxt [intro'']
- val intro'''' = Simplifier.full_simplify
- (HOL_basic_ss addsimps [@{thm fst_conv}, @{thm snd_conv}, @{thm Pair_eq}])
- intro'''
- (* splitting conjunctions introduced by Pair_eq*)
- fun split_conj prem =
- map HOLogic.mk_Trueprop (conjuncts (HOLogic.dest_Trueprop prem))
- val intro''''' = map_term thy (maps_premises split_conj) intro''''
- in
- intro'''''
- end
-
-
-
-end;
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_data.ML Wed Oct 28 11:42:31 2009 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,210 +0,0 @@
-(* Author: Lukas Bulwahn, TU Muenchen
-
-Book-keeping datastructure for the predicate compiler
-
-*)
-signature PREDICATE_COMPILE_DATA =
-sig
- type specification_table;
- val make_const_spec_table : theory -> specification_table
- val get_specification : specification_table -> string -> thm list
- val obtain_specification_graph : theory -> specification_table -> string -> thm list Graph.T
- val normalize_equation : theory -> thm -> thm
-end;
-
-structure Predicate_Compile_Data : PREDICATE_COMPILE_DATA =
-struct
-
-open Predicate_Compile_Aux;
-
-structure Data = TheoryDataFun
-(
- type T =
- {const_spec_table : thm list Symtab.table};
- val empty =
- {const_spec_table = Symtab.empty};
- val copy = I;
- val extend = I;
- fun merge _
- ({const_spec_table = const_spec_table1},
- {const_spec_table = const_spec_table2}) =
- {const_spec_table = Symtab.merge (K true) (const_spec_table1, const_spec_table2)}
-);
-
-fun mk_data c = {const_spec_table = c}
-fun map_data f {const_spec_table = c} = mk_data (f c)
-
-type specification_table = thm list Symtab.table
-
-fun defining_const_of_introrule_term t =
- let
- val _ $ u = Logic.strip_imp_concl t
- val (pred, all_args) = strip_comb u
- in case pred of
- Const (c, T) => c
- | _ => raise TERM ("defining_const_of_introrule_term failed: Not a constant", [t])
- end
-
-val defining_const_of_introrule = defining_const_of_introrule_term o prop_of
-
-(*TODO*)
-fun is_introlike_term t = true
-
-val is_introlike = is_introlike_term o prop_of
-
-fun check_equation_format_term (t as (Const ("==", _) $ u $ v)) =
- (case strip_comb u of
- (Const (c, T), args) =>
- if (length (binder_types T) = length args) then
- true
- else
- raise TERM ("check_equation_format_term failed: Number of arguments mismatch", [t])
- | _ => raise TERM ("check_equation_format_term failed: Not a constant", [t]))
- | check_equation_format_term t =
- raise TERM ("check_equation_format_term failed: Not an equation", [t])
-
-val check_equation_format = check_equation_format_term o prop_of
-
-fun defining_const_of_equation_term (t as (Const ("==", _) $ u $ v)) =
- (case fst (strip_comb u) of
- Const (c, _) => c
- | _ => raise TERM ("defining_const_of_equation_term failed: Not a constant", [t]))
- | defining_const_of_equation_term t =
- raise TERM ("defining_const_of_equation_term failed: Not an equation", [t])
-
-val defining_const_of_equation = defining_const_of_equation_term o prop_of
-
-(* Normalizing equations *)
-
-fun mk_meta_equation th =
- case prop_of th of
- Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _) => th RS @{thm eq_reflection}
- | _ => th
-
-val meta_fun_cong = @{lemma "f == g ==> f x == g x" by simp}
-
-fun full_fun_cong_expand th =
- let
- val (f, args) = strip_comb (fst (Logic.dest_equals (prop_of th)))
- val i = length (binder_types (fastype_of f)) - length args
- in funpow i (fn th => th RS meta_fun_cong) th end;
-
-fun declare_names s xs ctxt =
- let
- val res = Name.names ctxt s xs
- in (res, fold Name.declare (map fst res) ctxt) end
-
-fun split_all_pairs thy th =
- let
- val ctxt = ProofContext.init thy
- val ((_, [th']), ctxt') = Variable.import true [th] ctxt
- val t = prop_of th'
- val frees = Term.add_frees t []
- val freenames = Term.add_free_names t []
- val nctxt = Name.make_context freenames
- fun mk_tuple_rewrites (x, T) nctxt =
- let
- val Ts = HOLogic.flatten_tupleT T
- val (xTs, nctxt') = declare_names x Ts nctxt
- val paths = HOLogic.flat_tupleT_paths T
- in ((Free (x, T), HOLogic.mk_ptuple paths T (map Free xTs)), nctxt') end
- val (rewr, _) = fold_map mk_tuple_rewrites frees nctxt
- val t' = Pattern.rewrite_term thy rewr [] t
- val tac = Skip_Proof.cheat_tac thy
- val th'' = Goal.prove ctxt (Term.add_free_names t' []) [] t' (fn {...} => tac)
- val th''' = LocalDefs.unfold ctxt [@{thm split_conv}] th''
- in
- th'''
- end;
-
-fun normalize_equation thy th =
- mk_meta_equation th
- |> Predicate_Compile_Set.unfold_set_notation
- |> full_fun_cong_expand
- |> split_all_pairs thy
- |> tap check_equation_format
-
-fun inline_equations thy th = Conv.fconv_rule (Simplifier.rewrite
-((Simplifier.theory_context thy Simplifier.empty_ss) addsimps (Predicate_Compile_Inline_Defs.get (ProofContext.init thy)))) th
-
-fun store_thm_in_table ignore_consts thy th=
- let
- val th = th
- |> inline_equations thy
- |> Predicate_Compile_Set.unfold_set_notation
- |> AxClass.unoverload thy
- val (const, th) =
- if is_equationlike th then
- let
- val eq = normalize_equation thy th
- in
- (defining_const_of_equation eq, eq)
- end
- else if (is_introlike th) then (defining_const_of_introrule th, th)
- else error "store_thm: unexpected definition format"
- in
- if not (member (op =) ignore_consts const) then
- Symtab.cons_list (const, th)
- else I
- end
-
-fun make_const_spec_table thy =
- let
- fun store ignore_const f = fold (store_thm_in_table ignore_const thy) (map (Thm.transfer thy) (f (ProofContext.init thy)))
- val table = Symtab.empty
- |> store [] Predicate_Compile_Alternative_Defs.get
- val ignore_consts = Symtab.keys table
- in
- table
- |> store ignore_consts Predicate_Compile_Preproc_Const_Defs.get
- |> store ignore_consts Nitpick_Simps.get
- |> store ignore_consts Nitpick_Intros.get
- end
-
-fun get_specification table constname =
- case Symtab.lookup table constname of
- SOME thms => thms
- | NONE => error ("get_specification: lookup of constant " ^ quote constname ^ " failed")
-
-val logic_operator_names =
- [@{const_name "=="}, @{const_name "op ="}, @{const_name "op -->"}, @{const_name "All"}, @{const_name "Ex"},
- @{const_name "op &"}]
-
-val special_cases = member (op =) [
- @{const_name "False"},
- @{const_name "Suc"}, @{const_name Nat.zero_nat_inst.zero_nat},
- @{const_name Nat.one_nat_inst.one_nat},
-@{const_name "HOL.ord_class.less"}, @{const_name "HOL.ord_class.less_eq"}, @{const_name "HOL.zero_class.zero"},
-@{const_name "HOL.one_class.one"}, @{const_name HOL.plus_class.plus},
-@{const_name Nat.ord_nat_inst.less_eq_nat},
-@{const_name number_nat_inst.number_of_nat},
- @{const_name Int.Bit0},
- @{const_name Int.Bit1},
- @{const_name Int.Pls},
-@{const_name "Int.zero_int_inst.zero_int"},
-@{const_name "List.filter"}]
-
-fun case_consts thy s = is_some (Datatype.info_of_case thy s)
-
-fun obtain_specification_graph thy table constname =
- let
- fun is_nondefining_constname c = member (op =) logic_operator_names c
- val is_defining_constname = member (op =) (Symtab.keys table)
- fun has_code_pred_intros c = is_some (try (Predicate_Compile_Core.intros_of thy) c)
- fun defiants_of specs =
- fold (Term.add_const_names o prop_of) specs []
- |> filter is_defining_constname
- |> filter_out is_nondefining_constname
- |> filter_out has_code_pred_intros
- |> filter_out (case_consts thy)
- |> filter_out special_cases
- fun extend constname =
- let
- val specs = get_specification table constname
- in (specs, defiants_of specs) end;
- in
- Graph.extend extend constname Graph.empty
- end;
-
-
-end;
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_pred.ML Wed Oct 28 11:42:31 2009 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,189 +0,0 @@
-(* Author: Lukas Bulwahn, TU Muenchen
-
-Preprocessing definitions of predicates to introduction rules
-*)
-
-signature PREDICATE_COMPILE_PRED =
-sig
- (* preprocesses an equation to a set of intro rules; defines new constants *)
- (*
- val preprocess_pred_equation : thm -> theory -> thm list * theory
- *)
- val preprocess : string -> theory -> (thm list list * theory)
- (* output is the term list of clauses of an unknown predicate *)
- val preprocess_term : term -> theory -> (term list * theory)
-
- (*val rewrite : thm -> thm*)
-
-end;
-(* : PREDICATE_COMPILE_PREPROC_PRED *)
-structure Predicate_Compile_Pred =
-struct
-
-open Predicate_Compile_Aux
-
-fun is_compound ((Const ("Not", _)) $ t) =
- error "is_compound: Negation should not occur; preprocessing is defect"
- | is_compound ((Const ("Ex", _)) $ _) = true
- | is_compound ((Const (@{const_name "op |"}, _)) $ _ $ _) = true
- | is_compound ((Const ("op &", _)) $ _ $ _) =
- error "is_compound: Conjunction should not occur; preprocessing is defect"
- | is_compound _ = false
-
-fun flatten constname atom (defs, thy) =
- if is_compound atom then
- let
- val constname = Name.variant (map (Long_Name.base_name o fst) defs)
- ((Long_Name.base_name constname) ^ "_aux")
- val full_constname = Sign.full_bname thy constname
- val (params, args) = List.partition (is_predT o fastype_of)
- (map Free (Term.add_frees atom []))
- val constT = map fastype_of (params @ args) ---> HOLogic.boolT
- val lhs = list_comb (Const (full_constname, constT), params @ args)
- val def = Logic.mk_equals (lhs, atom)
- val ([definition], thy') = thy
- |> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)]
- |> PureThy.add_defs false [((Binding.name (constname ^ "_def"), def), [])]
- in
- (lhs, ((full_constname, [definition]) :: defs, thy'))
- end
- else
- (atom, (defs, thy))
-
-fun flatten_intros constname intros thy =
- let
- val ctxt = ProofContext.init thy
- val ((_, intros), ctxt') = Variable.import true intros ctxt
- val (intros', (local_defs, thy')) = (fold_map o fold_map_atoms)
- (flatten constname) (map prop_of intros) ([], thy)
- val tac = fn _ => Skip_Proof.cheat_tac thy'
- val intros'' = map (fn t => Goal.prove ctxt' [] [] t tac) intros'
- |> Variable.export ctxt' ctxt
- in
- (intros'', (local_defs, thy'))
- end
-
-(* TODO: same function occurs in inductive package *)
-fun select_disj 1 1 = []
- | select_disj _ 1 = [rtac @{thm disjI1}]
- | select_disj n i = (rtac @{thm disjI2})::(select_disj (n - 1) (i - 1));
-
-fun introrulify thy ths =
- let
- val ctxt = ProofContext.init thy
- val ((_, ths'), ctxt') = Variable.import true ths ctxt
- fun introrulify' th =
- let
- val (lhs, rhs) = Logic.dest_equals (prop_of th)
- val frees = Term.add_free_names rhs []
- val disjuncts = HOLogic.dest_disj rhs
- val nctxt = Name.make_context frees
- fun mk_introrule t =
- let
- val ((ps, t'), nctxt') = focus_ex t nctxt
- val prems = map HOLogic.mk_Trueprop (HOLogic.dest_conj t')
- in
- (ps, Logic.list_implies (prems, HOLogic.mk_Trueprop lhs))
- end
- val x = ((cterm_of thy) o the_single o snd o strip_comb o HOLogic.dest_Trueprop o fst o
- Logic.dest_implies o prop_of) @{thm exI}
- fun prove_introrule (index, (ps, introrule)) =
- let
- val tac = Simplifier.simp_tac (HOL_basic_ss addsimps [th]) 1
- THEN EVERY1 (select_disj (length disjuncts) (index + 1))
- THEN (EVERY (map (fn y =>
- rtac (Drule.cterm_instantiate [(x, cterm_of thy (Free y))] @{thm exI}) 1) ps))
- THEN REPEAT_DETERM (rtac @{thm conjI} 1 THEN atac 1)
- THEN TRY (atac 1)
- in
- Goal.prove ctxt' (map fst ps) [] introrule (fn {...} => tac)
- end
- in
- map_index prove_introrule (map mk_introrule disjuncts)
- end
- in maps introrulify' ths' |> Variable.export ctxt' ctxt end
-
-val rewrite =
- Simplifier.simplify (HOL_basic_ss addsimps [@{thm Ball_def}, @{thm Bex_def}])
- #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm all_not_ex}])
- #> Conv.fconv_rule nnf_conv
- #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm ex_disj_distrib}])
-
-val rewrite_intros =
- Simplifier.simplify (HOL_basic_ss addsimps @{thms HOL.simp_thms(9)})
-
-fun preprocess (constname, specs) thy =
- let
- val ctxt = ProofContext.init thy
- val intros =
- if forall is_pred_equation specs then
- introrulify thy (map rewrite specs)
- else if forall (is_intro constname) specs then
- map rewrite_intros specs
- else
- error ("unexpected specification for constant " ^ quote constname ^ ":\n"
- ^ commas (map (quote o Display.string_of_thm_global thy) specs))
- val (intros', (local_defs, thy')) = flatten_intros constname intros thy
- val (intross, thy'') = fold_map preprocess local_defs thy'
- in
- ((constname, intros') :: flat intross,thy'')
- end;
-
-fun preprocess_term t thy = error "preprocess_pred_term: to implement"
-
-fun is_Abs (Abs _) = true
- | is_Abs _ = false
-
-fun flat_higher_order_arguments (intross, thy) =
- let
- fun process constname atom (new_defs, thy) =
- let
- val (pred, args) = strip_comb atom
- val abs_args = filter is_Abs args
- fun replace_abs_arg (abs_arg as Abs _ ) (new_defs, thy) =
- let
- val _ = tracing ("Introduce new constant for " ^
- Syntax.string_of_term_global thy abs_arg)
- val vars = map Var (Term.add_vars abs_arg [])
- val abs_arg' = Logic.unvarify abs_arg
- val frees = map Free (Term.add_frees abs_arg' [])
- val constname = Name.variant (map (Long_Name.base_name o fst) new_defs)
- ((Long_Name.base_name constname) ^ "_hoaux")
- val full_constname = Sign.full_bname thy constname
- val constT = map fastype_of frees ---> (fastype_of abs_arg')
- val const = Const (full_constname, constT)
- val lhs = list_comb (const, frees)
- val def = Logic.mk_equals (lhs, abs_arg')
- val _ = tracing (Syntax.string_of_term_global thy def)
- val ([definition], thy') = thy
- |> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)]
- |> PureThy.add_defs false [((Binding.name (constname ^ "_def"), def), [])]
- in
- (list_comb (Logic.varify const, vars), ((full_constname, [definition])::new_defs, thy'))
- end
- | replace_abs_arg arg (new_defs, thy) = (arg, (new_defs, thy))
- val (args', (new_defs', thy')) = fold_map replace_abs_arg args (new_defs, thy)
- in
- (list_comb (pred, args'), (new_defs', thy'))
- end
- fun flat_intro intro (new_defs, thy) =
- let
- val constname = fst (dest_Const (fst (strip_comb
- (HOLogic.dest_Trueprop (Logic.strip_imp_concl (prop_of intro))))))
- val (intro_ts, (new_defs, thy)) = fold_map_atoms (process constname) (prop_of intro) (new_defs, thy)
- val th = Skip_Proof.make_thm thy intro_ts
- in
- (th, (new_defs, thy))
- end
- fun fold_map_spec f [] s = ([], s)
- | fold_map_spec f ((c, ths) :: specs) s =
- let
- val (ths', s') = f ths s
- val (specs', s'') = fold_map_spec f specs s'
- in ((c, ths') :: specs', s'') end
- val (intross', (new_defs, thy')) = fold_map_spec (fold_map flat_intro) intross ([], thy)
- in
- (intross', (new_defs, thy'))
- end
-
-end;
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_quickcheck.ML Wed Oct 28 11:42:31 2009 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,96 +0,0 @@
-(* Author: Lukas Bulwahn, TU Muenchen
-
-A quickcheck generator based on the predicate compiler
-*)
-
-signature PREDICATE_COMPILE_QUICKCHECK =
-sig
- val quickcheck : Proof.context -> term -> int -> term list option
- val test_ref :
- ((unit -> int -> int * int -> term list Predicate.pred * (int * int)) option) Unsynchronized.ref
-end;
-
-structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK =
-struct
-
-val test_ref =
- Unsynchronized.ref (NONE : (unit -> int -> int * int -> term list Predicate.pred * (int * int)) option)
-
-val target = "Quickcheck"
-
-fun dest_compfuns (Predicate_Compile_Core.CompilationFuns funs) = funs
-val mk_predT = #mk_predT (dest_compfuns Predicate_Compile_Core.pred_compfuns)
-val mk_rpredT = #mk_predT (dest_compfuns Predicate_Compile_Core.rpred_compfuns)
-val mk_return = #mk_single (dest_compfuns Predicate_Compile_Core.rpred_compfuns)
-val mk_bind = #mk_bind (dest_compfuns Predicate_Compile_Core.rpred_compfuns)
-val lift_pred = #lift_pred (dest_compfuns Predicate_Compile_Core.rpred_compfuns)
-
-fun mk_split_lambda [] t = Abs ("u", HOLogic.unitT, t)
- | mk_split_lambda [x] t = lambda x t
- | mk_split_lambda xs t =
- let
- fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
- | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
- in
- mk_split_lambda' xs t
- end;
-
-fun strip_imp_prems (Const("op -->", _) $ A $ B) = A :: strip_imp_prems B
- | strip_imp_prems _ = [];
-
-fun strip_imp_concl (Const("op -->", _) $ A $ B) = strip_imp_concl B
- | strip_imp_concl A = A : term;
-
-fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
-
-fun quickcheck ctxt t =
- let
- val _ = tracing ("Starting quickcheck with " ^ (Syntax.string_of_term ctxt t))
- val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
- val thy = (ProofContext.theory_of ctxt')
- val (vs, t') = strip_abs t
- val vs' = Variable.variant_frees ctxt' [] vs
- val t'' = subst_bounds (map Free (rev vs'), t')
- val (prems, concl) = strip_horn t''
- val constname = "pred_compile_quickcheck"
- val full_constname = Sign.full_bname thy constname
- val constT = map snd vs' ---> @{typ bool}
- val thy' = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy
- val t = Logic.list_implies
- (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),
- HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs')))
- val tac = fn _ => Skip_Proof.cheat_tac thy'
- val intro = Goal.prove (ProofContext.init thy') (map fst vs') [] t tac
- val _ = tracing (Display.string_of_thm ctxt' intro)
- val thy'' = thy'
- |> Context.theory_map (Predicate_Compile_Preproc_Const_Defs.add_thm intro)
- |> Predicate_Compile.preprocess Predicate_Compile_Aux.default_options full_constname
- (* |> Predicate_Compile_Core.add_equations Predicate_Compile_Aux.default_options [full_constname]*)
- (* |> Predicate_Compile_Core.add_depth_limited_equations Predicate_Compile_Aux.default_options [full_constname]*)
- (* |> Predicate_Compile_Core.add_quickcheck_equations Predicate_Compile_Aux.default_options [full_constname] *)
- val depth_limited_modes = Predicate_Compile_Core.depth_limited_modes_of thy'' full_constname
- val modes = Predicate_Compile_Core.generator_modes_of thy'' full_constname
- val prog =
- if member (op =) modes ([], []) then
- let
- val name = Predicate_Compile_Core.generator_name_of thy'' full_constname ([], [])
- val T = [@{typ bool}, @{typ code_numeral}] ---> (mk_rpredT (HOLogic.mk_tupleT (map snd vs')))
- in Const (name, T) $ @{term True} $ Bound 0 end
- else if member (op =) depth_limited_modes ([], []) then
- let
- val name = Predicate_Compile_Core.depth_limited_function_name_of thy'' full_constname ([], [])
- val T = @{typ code_numeral} --> (mk_predT (HOLogic.mk_tupleT (map snd vs')))
- in lift_pred (Const (name, T) $ Bound 0) end
- else error "Predicate Compile Quickcheck failed"
- val qc_term = Abs ("size", @{typ code_numeral}, mk_bind (prog,
- mk_split_lambda (map Free vs') (mk_return (HOLogic.mk_list @{typ term}
- (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))))
- val _ = tracing (Syntax.string_of_term ctxt' qc_term)
- val compile = Code_ML.eval (SOME target) ("Predicate_Compile_Quickcheck.test_ref", test_ref)
- (fn proc => fn g => fn s => g s #>> (Predicate.map o map) proc)
- thy'' qc_term []
- in
- ((compile #> Random_Engine.run) #> (Option.map fst o Predicate.yield))
- end
-
-end;
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_set.ML Wed Oct 28 11:42:31 2009 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,52 +0,0 @@
-(* Author: Lukas Bulwahn, TU Muenchen
-
-Preprocessing sets to predicates
-*)
-
-signature PREDICATE_COMPILE_SET =
-sig
-(*
- val preprocess_intro : thm -> theory -> thm * theory
- val preprocess_clause : term -> theory -> term * theory
-*)
- val unfold_set_notation : thm -> thm;
-end;
-
-structure Predicate_Compile_Set : PREDICATE_COMPILE_SET =
-struct
-(*FIXME: unfolding Ball in pretty adhoc here *)
-val unfold_set_lemmas = [@{thm Collect_def}, @{thm mem_def},
-@{thm Ball_def}, @{thm Bex_def}]
-
-val unfold_set_notation = Simplifier.rewrite_rule unfold_set_lemmas
-
-(*
-fun find_set_theorems ctxt cname =
- let
- val _ = cname
-*)
-
-(*
-fun preprocess_term t ctxt =
- case t of
- Const ("op :", _) $ x $ A =>
- case strip_comb A of
- (Const (cname, T), params) =>
- let
- val _ = find_set_theorems
- in
- (t, ctxt)
- end
- | _ => (t, ctxt)
- | _ => (t, ctxt)
-*)
-(*
-fun preprocess_intro th thy =
- let
- val cnames = find_heads_of_prems
- val set_cname = filter (has_set_definition
- val _ = define_preds
- val _ = prep_pred_def
- in
-*)
-end;
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Wed Oct 28 11:43:06 2009 +0000
@@ -1,6 +1,9 @@
-(* Author: Lukas Bulwahn, TU Muenchen
+(* Title: HOL/Tools/Predicate_Compile/predicate_compile.ML
+ Author: Lukas Bulwahn, TU Muenchen
+FIXME.
*)
+
signature PREDICATE_COMPILE =
sig
val setup : theory -> theory
@@ -99,7 +102,7 @@
let
val _ = print_step options "Fetching definitions from theory..."
val table = Predicate_Compile_Data.make_const_spec_table thy
- val gr = Predicate_Compile_Data.obtain_specification_graph thy table const
+ val gr = Predicate_Compile_Data.obtain_specification_graph options thy table const
val gr = Graph.subgraph (member (op =) (Graph.all_succs gr [const])) gr
in fold_rev (preprocess_strong_conn_constnames options gr)
(Graph.strong_conn gr) thy
@@ -114,6 +117,7 @@
show_steps = chk "show_steps",
show_intermediate_results = chk "show_intermediate_results",
show_proof_trace = chk "show_proof_trace",
+ show_modes = chk "show_modes",
show_mode_inference = chk "show_mode_inference",
show_compilation = chk "show_compilation",
skip_proof = chk "skip_proof",
@@ -146,7 +150,7 @@
val setup = Predicate_Compile_Fun.setup_oracle #> Predicate_Compile_Core.setup
-val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace",
+val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace", "show_modes",
"show_mode_inference", "show_compilation", "skip_proof", "inductify", "rpred", "depth_limited"]
local structure P = OuterParse
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Wed Oct 28 11:43:06 2009 +0000
@@ -0,0 +1,240 @@
+(* Title: HOL/Tools/Predicate_Compile/predicate_compile_aux.ML
+ Author: Lukas Bulwahn, TU Muenchen
+
+Auxilary functions for predicate compiler.
+*)
+
+(* FIXME proper signature *)
+
+structure Predicate_Compile_Aux =
+struct
+
+(* general syntactic functions *)
+
+(*Like dest_conj, but flattens conjunctions however nested*)
+fun conjuncts_aux (Const ("op &", _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs)
+ | conjuncts_aux t conjs = t::conjs;
+
+fun conjuncts t = conjuncts_aux t [];
+
+(* syntactic functions *)
+
+fun is_equationlike_term (Const ("==", _) $ _ $ _) = true
+ | is_equationlike_term (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _)) = true
+ | is_equationlike_term _ = false
+
+val is_equationlike = is_equationlike_term o prop_of
+
+fun is_pred_equation_term (Const ("==", _) $ u $ v) =
+ (fastype_of u = @{typ bool}) andalso (fastype_of v = @{typ bool})
+ | is_pred_equation_term _ = false
+
+val is_pred_equation = is_pred_equation_term o prop_of
+
+fun is_intro_term constname t =
+ case fst (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl t))) of
+ Const (c, _) => c = constname
+ | _ => false
+
+fun is_intro constname t = is_intro_term constname (prop_of t)
+
+fun is_pred thy constname =
+ let
+ val T = (Sign.the_const_type thy constname)
+ in body_type T = @{typ "bool"} end;
+
+
+fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT)
+ | is_predT _ = false
+
+
+(*** check if a term contains only constructor functions ***)
+fun is_constrt thy =
+ let
+ val cnstrs = flat (maps
+ (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
+ (Symtab.dest (Datatype.get_all thy)));
+ fun check t = (case strip_comb t of
+ (Free _, []) => true
+ | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
+ (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts
+ | _ => false)
+ | _ => false)
+ in check end;
+
+fun strip_ex (Const ("Ex", _) $ Abs (x, T, t)) =
+ let
+ val (xTs, t') = strip_ex t
+ in
+ ((x, T) :: xTs, t')
+ end
+ | strip_ex t = ([], t)
+
+fun focus_ex t nctxt =
+ let
+ val ((xs, Ts), t') = apfst split_list (strip_ex t)
+ val (xs', nctxt') = Name.variants xs nctxt;
+ val ps' = xs' ~~ Ts;
+ val vs = map Free ps';
+ val t'' = Term.subst_bounds (rev vs, t');
+ in ((ps', t''), nctxt') end;
+
+
+(* introduction rule combinators *)
+
+(* combinators to apply a function to all literals of an introduction rules *)
+
+fun map_atoms f intro =
+ let
+ val (literals, head) = Logic.strip_horn intro
+ fun appl t = (case t of
+ (@{term "Not"} $ t') => HOLogic.mk_not (f t')
+ | _ => f t)
+ in
+ Logic.list_implies
+ (map (HOLogic.mk_Trueprop o appl o HOLogic.dest_Trueprop) literals, head)
+ end
+
+fun fold_atoms f intro s =
+ let
+ val (literals, head) = Logic.strip_horn intro
+ fun appl t s = (case t of
+ (@{term "Not"} $ t') => f t' s
+ | _ => f t s)
+ in fold appl (map HOLogic.dest_Trueprop literals) s end
+
+fun fold_map_atoms f intro s =
+ let
+ val (literals, head) = Logic.strip_horn intro
+ fun appl t s = (case t of
+ (@{term "Not"} $ t') => apfst HOLogic.mk_not (f t' s)
+ | _ => f t s)
+ val (literals', s') = fold_map appl (map HOLogic.dest_Trueprop literals) s
+ in
+ (Logic.list_implies (map HOLogic.mk_Trueprop literals', head), s')
+ end;
+
+fun maps_premises f intro =
+ let
+ val (premises, head) = Logic.strip_horn intro
+ in
+ Logic.list_implies (maps f premises, head)
+ end
+
+(* lifting term operations to theorems *)
+
+fun map_term thy f th =
+ Skip_Proof.make_thm thy (f (prop_of th))
+
+(*
+fun equals_conv lhs_cv rhs_cv ct =
+ case Thm.term_of ct of
+ Const ("==", _) $ _ $ _ => Conv.arg_conv cv ct
+ | _ => error "equals_conv"
+*)
+
+(* Different options for compiler *)
+
+datatype options = Options of {
+ expected_modes : (string * int list list) option,
+ show_steps : bool,
+ show_proof_trace : bool,
+ show_intermediate_results : bool,
+ show_mode_inference : bool,
+ show_modes : bool,
+ show_compilation : bool,
+ skip_proof : bool,
+
+ inductify : bool,
+ rpred : bool,
+ depth_limited : bool
+};
+
+fun expected_modes (Options opt) = #expected_modes opt
+fun show_steps (Options opt) = #show_steps opt
+fun show_intermediate_results (Options opt) = #show_intermediate_results opt
+fun show_proof_trace (Options opt) = #show_proof_trace opt
+fun show_modes (Options opt) = #show_modes opt
+fun show_mode_inference (Options opt) = #show_mode_inference opt
+fun show_compilation (Options opt) = #show_compilation opt
+fun skip_proof (Options opt) = #skip_proof opt
+
+fun is_inductify (Options opt) = #inductify opt
+fun is_rpred (Options opt) = #rpred opt
+fun is_depth_limited (Options opt) = #depth_limited opt
+
+val default_options = Options {
+ expected_modes = NONE,
+ show_steps = false,
+ show_intermediate_results = false,
+ show_proof_trace = false,
+ show_modes = false,
+ show_mode_inference = false,
+ show_compilation = false,
+ skip_proof = false,
+
+ inductify = false,
+ rpred = false,
+ depth_limited = false
+}
+
+
+fun print_step options s =
+ if show_steps options then tracing s else ()
+
+(* tuple processing *)
+
+fun expand_tuples thy intro =
+ let
+ fun rewrite_args [] (pats, intro_t, ctxt) = (pats, intro_t, ctxt)
+ | rewrite_args (arg::args) (pats, intro_t, ctxt) =
+ (case HOLogic.strip_tupleT (fastype_of arg) of
+ (Ts as _ :: _ :: _) =>
+ let
+ fun rewrite_arg' (Const ("Pair", _) $ _ $ t2, Type ("*", [_, T2]))
+ (args, (pats, intro_t, ctxt)) = rewrite_arg' (t2, T2) (args, (pats, intro_t, ctxt))
+ | rewrite_arg' (t, Type ("*", [T1, T2])) (args, (pats, intro_t, ctxt)) =
+ let
+ val ([x, y], ctxt') = Variable.variant_fixes ["x", "y"] ctxt
+ val pat = (t, HOLogic.mk_prod (Free (x, T1), Free (y, T2)))
+ val intro_t' = Pattern.rewrite_term thy [pat] [] intro_t
+ val args' = map (Pattern.rewrite_term thy [pat] []) args
+ in
+ rewrite_arg' (Free (y, T2), T2) (args', (pat::pats, intro_t', ctxt'))
+ end
+ | rewrite_arg' _ (args, (pats, intro_t, ctxt)) = (args, (pats, intro_t, ctxt))
+ val (args', (pats, intro_t', ctxt')) = rewrite_arg' (arg, fastype_of arg)
+ (args, (pats, intro_t, ctxt))
+ in
+ rewrite_args args' (pats, intro_t', ctxt')
+ end
+ | _ => rewrite_args args (pats, intro_t, ctxt))
+ fun rewrite_prem atom =
+ let
+ val (_, args) = strip_comb atom
+ in rewrite_args args end
+ val ctxt = ProofContext.init thy
+ val (((T_insts, t_insts), [intro']), ctxt1) = Variable.import false [intro] ctxt
+ val intro_t = prop_of intro'
+ val concl = Logic.strip_imp_concl intro_t
+ val (p, args) = strip_comb (HOLogic.dest_Trueprop concl)
+ val (pats', intro_t', ctxt2) = rewrite_args args ([], intro_t, ctxt1)
+ val (pats', intro_t', ctxt3) =
+ fold_atoms rewrite_prem intro_t' (pats', intro_t', ctxt2)
+ fun rewrite_pat (ct1, ct2) =
+ (ct1, cterm_of thy (Pattern.rewrite_term thy pats' [] (term_of ct2)))
+ val t_insts' = map rewrite_pat t_insts
+ val intro'' = Thm.instantiate (T_insts, t_insts') intro
+ val [intro'''] = Variable.export ctxt3 ctxt [intro'']
+ val intro'''' = Simplifier.full_simplify
+ (HOL_basic_ss addsimps [@{thm fst_conv}, @{thm snd_conv}, @{thm Pair_eq}])
+ intro'''
+ (* splitting conjunctions introduced by Pair_eq*)
+ fun split_conj prem =
+ map HOLogic.mk_Trueprop (conjuncts (HOLogic.dest_Trueprop prem))
+ val intro''''' = map_term thy (maps_premises split_conj) intro''''
+ in
+ intro'''''
+ end
+
+end;
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Wed Oct 28 11:43:06 2009 +0000
@@ -1,7 +1,7 @@
-(* Author: Lukas Bulwahn, TU Muenchen
+(* Title: HOL/Tools/Predicate_Compile/predicate_compile_core.ML
+ Author: Lukas Bulwahn, TU Muenchen
-(Prototype of) A compiler from predicates specified by intro/elim rules
-to equations.
+A compiler from predicates specified by intro/elim rules to equations.
*)
signature PREDICATE_COMPILE_CORE =
@@ -49,15 +49,13 @@
mk_sup : term * term -> term,
mk_if : term -> term,
mk_not : term -> term,
- mk_map : typ -> typ -> term -> term -> term,
- lift_pred : term -> term
+ mk_map : typ -> typ -> term -> term -> term
};
val pred_compfuns : compilation_funs
- val rpred_compfuns : compilation_funs
- (* val add_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
+ val randompred_compfuns : compilation_funs
+ val add_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
val add_quickcheck_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
val add_depth_limited_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
- *)
end;
structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
@@ -333,10 +331,12 @@
(* diagnostic display functions *)
-fun print_modes modes =
- Output.tracing ("Inferred modes:\n" ^
- cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
- string_of_mode ms)) modes));
+fun print_modes options modes =
+ if show_modes options then
+ Output.tracing ("Inferred modes:\n" ^
+ cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
+ string_of_mode ms)) modes))
+ else ()
fun print_pred_mode_table string_of_entry thy pred_mode_table =
let
@@ -720,12 +720,6 @@
PredData.map (Graph.map_node pred (map_pred_data set))
end
-(** data structures for generic compilation for different monads **)
-
-(* maybe rename functions more generic:
- mk_predT -> mk_monadT; dest_predT -> dest_monadT
- mk_single -> mk_return (?)
-*)
datatype compilation_funs = CompilationFuns of {
mk_predT : typ -> typ,
dest_predT : typ -> typ,
@@ -735,8 +729,7 @@
mk_sup : term * term -> term,
mk_if : term -> term,
mk_not : term -> term,
- mk_map : typ -> typ -> term -> term -> term,
- lift_pred : term -> term
+ mk_map : typ -> typ -> term -> term -> term
};
fun mk_predT (CompilationFuns funs) = #mk_predT funs
@@ -748,7 +741,6 @@
fun mk_if (CompilationFuns funs) = #mk_if funs
fun mk_not (CompilationFuns funs) = #mk_not funs
fun mk_map (CompilationFuns funs) = #mk_map funs
-fun lift_pred (CompilationFuns funs) = #lift_pred funs
fun funT_of compfuns (iss, is) T =
let
@@ -766,9 +758,9 @@
structure PredicateCompFuns =
struct
-fun mk_predT T = Type (@{type_name "Predicate.pred"}, [T])
+fun mk_predT T = Type (@{type_name Predicate.pred}, [T])
-fun dest_predT (Type (@{type_name "Predicate.pred"}, [T])) = T
+fun dest_predT (Type (@{type_name Predicate.pred}, [T])) = T
| dest_predT T = raise TYPE ("dest_predT", [T], []);
fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T);
@@ -807,75 +799,65 @@
fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map},
(T1 --> T2) --> mk_predT T1 --> mk_predT T2) $ tf $ tp;
-val lift_pred = I
-
val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot,
mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
- mk_map = mk_map, lift_pred = lift_pred};
+ mk_map = mk_map};
end;
-structure RPredCompFuns =
+structure RandomPredCompFuns =
struct
-fun mk_rpredT T =
- @{typ "Random.seed"} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ "Random.seed"})
+fun mk_randompredT T =
+ @{typ Random.seed} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ Random.seed})
-fun dest_rpredT (Type ("fun", [_,
- Type (@{type_name "*"}, [Type (@{type_name "Predicate.pred"}, [T]), _])])) = T
- | dest_rpredT T = raise TYPE ("dest_rpredT", [T], []);
+fun dest_randompredT (Type ("fun", [@{typ Random.seed}, Type (@{type_name "*"},
+ [Type (@{type_name "Predicate.pred"}, [T]), @{typ Random.seed}])])) = T
+ | dest_randompredT T = raise TYPE ("dest_randompredT", [T], []);
-fun mk_bot T = Const(@{const_name RPred.bot}, mk_rpredT T)
+fun mk_bot T = Const(@{const_name Quickcheck.empty}, mk_randompredT T)
fun mk_single t =
let
val T = fastype_of t
in
- Const (@{const_name RPred.return}, T --> mk_rpredT T) $ t
+ Const (@{const_name Quickcheck.single}, T --> mk_randompredT T) $ t
end;
fun mk_bind (x, f) =
let
val T as (Type ("fun", [_, U])) = fastype_of f
in
- Const (@{const_name RPred.bind}, fastype_of x --> T --> U) $ x $ f
+ Const (@{const_name Quickcheck.bind}, fastype_of x --> T --> U) $ x $ f
end
-val mk_sup = HOLogic.mk_binop @{const_name RPred.supp}
-
-fun mk_if cond = Const (@{const_name RPred.if_rpred},
- HOLogic.boolT --> mk_rpredT HOLogic.unitT) $ cond;
+val mk_sup = HOLogic.mk_binop @{const_name Quickcheck.union}
-fun mk_not t = let val T = mk_rpredT HOLogic.unitT
- in Const (@{const_name RPred.not_rpred}, T --> T) $ t end
-
-fun mk_map T1 T2 tf tp = Const (@{const_name RPred.map},
- (T1 --> T2) --> mk_rpredT T1 --> mk_rpredT T2) $ tf $ tp
+fun mk_if cond = Const (@{const_name Quickcheck.if_randompred},
+ HOLogic.boolT --> mk_randompredT HOLogic.unitT) $ cond;
-fun lift_pred t =
- let
- val T = PredicateCompFuns.dest_predT (fastype_of t)
- val lift_predT = PredicateCompFuns.mk_predT T --> mk_rpredT T
- in
- Const (@{const_name "RPred.lift_pred"}, lift_predT) $ t
- end;
+fun mk_not t = let val T = mk_randompredT HOLogic.unitT
+ in Const (@{const_name Quickcheck.not_randompred}, T --> T) $ t end
-val compfuns = CompilationFuns {mk_predT = mk_rpredT, dest_predT = dest_rpredT, mk_bot = mk_bot,
+fun mk_map T1 T2 tf tp = Const (@{const_name Quickcheck.map},
+ (T1 --> T2) --> mk_randompredT T1 --> mk_randompredT T2) $ tf $ tp
+
+val compfuns = CompilationFuns {mk_predT = mk_randompredT, dest_predT = dest_randompredT, mk_bot = mk_bot,
mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
- mk_map = mk_map, lift_pred = lift_pred};
+ mk_map = mk_map};
end;
(* for external use with interactive mode *)
val pred_compfuns = PredicateCompFuns.compfuns
-val rpred_compfuns = RPredCompFuns.compfuns;
+val randompred_compfuns = RandomPredCompFuns.compfuns;
fun lift_random random =
let
val T = dest_randomT (fastype_of random)
in
- Const (@{const_name lift_random}, (@{typ Random.seed} -->
+ Const (@{const_name Quickcheck.Random}, (@{typ Random.seed} -->
HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) -->
- RPredCompFuns.mk_rpredT T) $ random
+ RandomPredCompFuns.mk_randompredT T) $ random
end;
fun depth_limited_funT_of compfuns (iss, is) T =
@@ -1158,51 +1140,51 @@
fun mk_Eval_of additional_arguments ((x, T), NONE) names = (x, names)
| mk_Eval_of additional_arguments ((x, T), SOME mode) names =
- let
+ let
val Ts = binder_types T
(*val argnames = Name.variant_list names
(map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
val args = map Free (argnames ~~ Ts)
val (inargs, outargs) = split_smode mode args*)
- fun mk_split_lambda [] t = lambda (Free (Name.variant names "x", HOLogic.unitT)) t
- | mk_split_lambda [x] t = lambda x t
- | mk_split_lambda xs t =
- let
- fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
- | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
- in
- mk_split_lambda' xs t
- end;
- fun mk_arg (i, T) =
- let
- val vname = Name.variant names ("x" ^ string_of_int i)
- val default = Free (vname, T)
- in
- case AList.lookup (op =) mode i of
- NONE => (([], [default]), [default])
- | SOME NONE => (([default], []), [default])
- | SOME (SOME pis) =>
- case HOLogic.strip_tupleT T of
- [] => error "pair mode but unit tuple" (*(([default], []), [default])*)
- | [_] => error "pair mode but not a tuple" (*(([default], []), [default])*)
- | Ts =>
- let
- val vnames = Name.variant_list names
- (map (fn j => "x" ^ string_of_int i ^ "p" ^ string_of_int j)
- (1 upto length Ts))
- val args = map Free (vnames ~~ Ts)
- fun split_args (i, arg) (ins, outs) =
- if member (op =) pis i then
- (arg::ins, outs)
- else
- (ins, arg::outs)
- val (inargs, outargs) = fold_rev split_args ((1 upto length Ts) ~~ args) ([], [])
- fun tuple args = if null args then [] else [HOLogic.mk_tuple args]
- in ((tuple inargs, tuple outargs), args) end
- end
- val (inoutargs, args) = split_list (map mk_arg (1 upto (length Ts) ~~ Ts))
+ fun mk_split_lambda [] t = lambda (Free (Name.variant names "x", HOLogic.unitT)) t
+ | mk_split_lambda [x] t = lambda x t
+ | mk_split_lambda xs t =
+ let
+ fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
+ | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
+ in
+ mk_split_lambda' xs t
+ end;
+ fun mk_arg (i, T) =
+ let
+ val vname = Name.variant names ("x" ^ string_of_int i)
+ val default = Free (vname, T)
+ in
+ case AList.lookup (op =) mode i of
+ NONE => (([], [default]), [default])
+ | SOME NONE => (([default], []), [default])
+ | SOME (SOME pis) =>
+ case HOLogic.strip_tupleT T of
+ [] => error "pair mode but unit tuple" (*(([default], []), [default])*)
+ | [_] => error "pair mode but not a tuple" (*(([default], []), [default])*)
+ | Ts =>
+ let
+ val vnames = Name.variant_list names
+ (map (fn j => "x" ^ string_of_int i ^ "p" ^ string_of_int j)
+ (1 upto length Ts))
+ val args = map Free (vnames ~~ Ts)
+ fun split_args (i, arg) (ins, outs) =
+ if member (op =) pis i then
+ (arg::ins, outs)
+ else
+ (ins, arg::outs)
+ val (inargs, outargs) = fold_rev split_args ((1 upto length Ts) ~~ args) ([], [])
+ fun tuple args = if null args then [] else [HOLogic.mk_tuple args]
+ in ((tuple inargs, tuple outargs), args) end
+ end
+ val (inoutargs, args) = split_list (map mk_arg (1 upto (length Ts) ~~ Ts))
val (inargs, outargs) = pairself flat (split_list inoutargs)
- val r = PredicateCompFuns.mk_Eval
+ val r = PredicateCompFuns.mk_Eval
(list_comb (x, inargs @ additional_arguments), HOLogic.mk_tuple outargs)
val t = fold_rev mk_split_lambda args r
in
@@ -1371,22 +1353,22 @@
fun compile_pred compilation_modifiers compfuns thy all_vs param_vs s T mode moded_cls =
let
- val (Ts1, Ts2) = chop (length (fst mode)) (binder_types T)
+ val (Ts1, Ts2) = chop (length (fst mode)) (binder_types T)
val (Us1, Us2) = split_smodeT (snd mode) Ts2
val Ts1' =
map2 (fn NONE => I | SOME is => #funT_of compilation_modifiers compfuns ([], is)) (fst mode) Ts1
fun mk_input_term (i, NONE) =
- [Free (Name.variant (all_vs @ param_vs) ("x" ^ string_of_int i), nth Ts2 (i - 1))]
- | mk_input_term (i, SOME pis) = case HOLogic.strip_tupleT (nth Ts2 (i - 1)) of
- [] => error "strange unit input"
- | [T] => [Free (Name.variant (all_vs @ param_vs) ("x" ^ string_of_int i), nth Ts2 (i - 1))]
- | Ts => let
- val vnames = Name.variant_list (all_vs @ param_vs)
- (map (fn j => "x" ^ string_of_int i ^ "p" ^ string_of_int j)
- pis)
- in if null pis then []
- else [HOLogic.mk_tuple (map Free (vnames ~~ map (fn j => nth Ts (j - 1)) pis))] end
- val in_ts = maps mk_input_term (snd mode)
+ [Free (Name.variant (all_vs @ param_vs) ("x" ^ string_of_int i), nth Ts2 (i - 1))]
+ | mk_input_term (i, SOME pis) = case HOLogic.strip_tupleT (nth Ts2 (i - 1)) of
+ [] => error "strange unit input"
+ | [T] => [Free (Name.variant (all_vs @ param_vs) ("x" ^ string_of_int i), nth Ts2 (i - 1))]
+ | Ts => let
+ val vnames = Name.variant_list (all_vs @ param_vs)
+ (map (fn j => "x" ^ string_of_int i ^ "p" ^ string_of_int j)
+ pis)
+ in if null pis then []
+ else [HOLogic.mk_tuple (map Free (vnames ~~ map (fn j => nth Ts (j - 1)) pis))] end
+ val in_ts = maps mk_input_term (snd mode)
val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1'
val additional_arguments = #additional_arguments compilation_modifiers (all_vs @ param_vs)
val cl_ts =
@@ -1407,7 +1389,7 @@
(* special setup for simpset *)
val HOL_basic_ss' = HOL_basic_ss addsimps (@{thms "HOL.simp_thms"} @ [@{thm Pair_eq}])
setSolver (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
- setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI}))
+ setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI}))
(* Definition of executable functions and their intro and elim rules *)
@@ -1423,29 +1405,29 @@
val funtrm = Const (mode_id, funT)
val (Ts1, Ts2) = chop (length iss) Ts;
val Ts1' = map2 (fn NONE => I | SOME is => funT_of (PredicateCompFuns.compfuns) ([], is)) iss Ts1
- val param_names = Name.variant_list []
+ val param_names = Name.variant_list []
(map (fn i => "x" ^ string_of_int i) (1 upto (length Ts1)));
val params = map Free (param_names ~~ Ts1')
- fun mk_args (i, T) argnames =
+ fun mk_args (i, T) argnames =
let
- val vname = Name.variant (param_names @ argnames) ("x" ^ string_of_int (length Ts1' + i))
- val default = (Free (vname, T), vname :: argnames)
- in
- case AList.lookup (op =) is i of
- NONE => default
- | SOME NONE => default
- | SOME (SOME pis) =>
- case HOLogic.strip_tupleT T of
- [] => default
- | [_] => default
- | Ts =>
- let
- val vnames = Name.variant_list (param_names @ argnames)
- (map (fn j => "x" ^ string_of_int (length Ts1' + i) ^ "p" ^ string_of_int j)
- (1 upto (length Ts)))
- in (HOLogic.mk_tuple (map Free (vnames ~~ Ts)), vnames @ argnames) end
- end
- val (args, argnames) = fold_map mk_args (1 upto (length Ts2) ~~ Ts2) []
+ val vname = Name.variant (param_names @ argnames) ("x" ^ string_of_int (length Ts1' + i))
+ val default = (Free (vname, T), vname :: argnames)
+ in
+ case AList.lookup (op =) is i of
+ NONE => default
+ | SOME NONE => default
+ | SOME (SOME pis) =>
+ case HOLogic.strip_tupleT T of
+ [] => default
+ | [_] => default
+ | Ts =>
+ let
+ val vnames = Name.variant_list (param_names @ argnames)
+ (map (fn j => "x" ^ string_of_int (length Ts1' + i) ^ "p" ^ string_of_int j)
+ (1 upto (length Ts)))
+ in (HOLogic.mk_tuple (map Free (vnames ~~ Ts)), vnames @ argnames) end
+ end
+ val (args, argnames) = fold_map mk_args (1 upto (length Ts2) ~~ Ts2) []
val (inargs, outargs) = split_smode is args
val param_names' = Name.variant_list (param_names @ argnames)
(map (fn i => "p" ^ string_of_int i) (1 upto (length iss)))
@@ -1461,7 +1443,7 @@
HOLogic.mk_tuple outargs))
val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
val simprules = [defthm, @{thm eval_pred},
- @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}]
+ @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}]
val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
val introthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ param_names' @ ["y"]) [] introtrm (fn {...} => unfolddef_tac)
val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
@@ -1484,30 +1466,30 @@
end;
fun split_tupleT is T =
- let
- fun split_tuple' _ _ [] = ([], [])
- | split_tuple' is i (T::Ts) =
- (if i mem is then apfst else apsnd) (cons T)
- (split_tuple' is (i+1) Ts)
- in
- split_tuple' is 1 (HOLogic.strip_tupleT T)
+ let
+ fun split_tuple' _ _ [] = ([], [])
+ | split_tuple' is i (T::Ts) =
+ (if i mem is then apfst else apsnd) (cons T)
+ (split_tuple' is (i+1) Ts)
+ in
+ split_tuple' is 1 (HOLogic.strip_tupleT T)
end
-
+
fun mk_arg xin xout pis T =
let
- val n = length (HOLogic.strip_tupleT T)
- val ni = length pis
- fun mk_proj i j t =
- (if i = j then I else HOLogic.mk_fst)
- (funpow (i - 1) HOLogic.mk_snd t)
- fun mk_arg' i (si, so) = if i mem pis then
- (mk_proj si ni xin, (si+1, so))
- else
- (mk_proj so (n - ni) xout, (si, so+1))
- val (args, _) = fold_map mk_arg' (1 upto n) (1, 1)
- in
- HOLogic.mk_tuple args
- end
+ val n = length (HOLogic.strip_tupleT T)
+ val ni = length pis
+ fun mk_proj i j t =
+ (if i = j then I else HOLogic.mk_fst)
+ (funpow (i - 1) HOLogic.mk_snd t)
+ fun mk_arg' i (si, so) = if i mem pis then
+ (mk_proj si ni xin, (si+1, so))
+ else
+ (mk_proj so (n - ni) xout, (si, so+1))
+ val (args, _) = fold_map mk_arg' (1 upto n) (1, 1)
+ in
+ HOLogic.mk_tuple args
+ end
fun create_definitions preds (name, modes) thy =
let
@@ -1523,37 +1505,37 @@
val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (HOLogic.mk_tupleT Us2))
val names = Name.variant_list []
(map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
- (* old *)
- (*
- val xs = map Free (names ~~ (Ts1' @ Ts2))
+ (* old *)
+ (*
+ val xs = map Free (names ~~ (Ts1' @ Ts2))
val (xparams, xargs) = chop (length iss) xs
val (xins, xouts) = split_smode is xargs
- *)
- (* new *)
- val param_names = Name.variant_list []
- (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts1')))
- val xparams = map Free (param_names ~~ Ts1')
+ *)
+ (* new *)
+ val param_names = Name.variant_list []
+ (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts1')))
+ val xparams = map Free (param_names ~~ Ts1')
fun mk_vars (i, T) names =
- let
- val vname = Name.variant names ("x" ^ string_of_int (length Ts1' + i))
- in
- case AList.lookup (op =) is i of
- NONE => ((([], [Free (vname, T)]), Free (vname, T)), vname :: names)
- | SOME NONE => ((([Free (vname, T)], []), Free (vname, T)), vname :: names)
- | SOME (SOME pis) =>
- let
- val (Tins, Touts) = split_tupleT pis T
- val name_in = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "in")
- val name_out = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "out")
- val xin = Free (name_in, HOLogic.mk_tupleT Tins)
- val xout = Free (name_out, HOLogic.mk_tupleT Touts)
- val xarg = mk_arg xin xout pis T
- in (((if null Tins then [] else [xin], if null Touts then [] else [xout]), xarg), name_in :: name_out :: names) end
- end
- val (xinoutargs, names) = fold_map mk_vars ((1 upto (length Ts2)) ~~ Ts2) param_names
+ let
+ val vname = Name.variant names ("x" ^ string_of_int (length Ts1' + i))
+ in
+ case AList.lookup (op =) is i of
+ NONE => ((([], [Free (vname, T)]), Free (vname, T)), vname :: names)
+ | SOME NONE => ((([Free (vname, T)], []), Free (vname, T)), vname :: names)
+ | SOME (SOME pis) =>
+ let
+ val (Tins, Touts) = split_tupleT pis T
+ val name_in = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "in")
+ val name_out = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "out")
+ val xin = Free (name_in, HOLogic.mk_tupleT Tins)
+ val xout = Free (name_out, HOLogic.mk_tupleT Touts)
+ val xarg = mk_arg xin xout pis T
+ in (((if null Tins then [] else [xin], if null Touts then [] else [xout]), xarg), name_in :: name_out :: names) end
+ end
+ val (xinoutargs, names) = fold_map mk_vars ((1 upto (length Ts2)) ~~ Ts2) param_names
val (xinout, xargs) = split_list xinoutargs
- val (xins, xouts) = pairself flat (split_list xinout)
- val (xparams', names') = fold_map (mk_Eval_of []) ((xparams ~~ Ts1) ~~ iss) names
+ val (xins, xouts) = pairself flat (split_list xinout)
+ val (xparams', names') = fold_map (mk_Eval_of []) ((xparams ~~ Ts1) ~~ iss) names
fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
| mk_split_lambda [x] t = lambda x t
| mk_split_lambda xs t =
@@ -1573,7 +1555,7 @@
val (intro, elim) =
create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
in thy'
- |> add_predfun name mode (mode_cname, definition, intro, elim)
+ |> add_predfun name mode (mode_cname, definition, intro, elim)
|> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
|> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim) |> snd
|> Theory.checkpoint
@@ -1604,7 +1586,7 @@
val paramTs' = map2 (fn SOME is => generator_funT_of ([], is) | NONE => I) iss paramTs
in
(paramTs' @ inargTs @ [@{typ code_numeral}]) --->
- (mk_predT RPredCompFuns.compfuns (HOLogic.mk_tupleT outargTs))
+ (mk_predT RandomPredCompFuns.compfuns (HOLogic.mk_tupleT outargTs))
end
fun rpred_create_definitions preds (name, modes) thy =
@@ -1653,7 +1635,7 @@
Const (name, T) => simp_tac (HOL_basic_ss addsimps
([@{thm eval_pred}, (predfun_definition_of thy name mode),
@{thm "split_eta"}, @{thm "split_beta"}, @{thm "fst_conv"},
- @{thm "snd_conv"}, @{thm pair_collapse}, @{thm "Product_Type.split_conv"}])) 1
+ @{thm "snd_conv"}, @{thm pair_collapse}, @{thm "Product_Type.split_conv"}])) 1
| Free _ => TRY (rtac @{thm refl} 1)
| Abs _ => error "prove_param: No valid parameter term"
in
@@ -1687,12 +1669,12 @@
THEN (REPEAT_DETERM (atac 1))
end
| _ => rtac @{thm bindI} 1
- THEN asm_full_simp_tac
- (HOL_basic_ss' addsimps [@{thm "split_eta"}, @{thm "split_beta"}, @{thm "fst_conv"},
- @{thm "snd_conv"}, @{thm pair_collapse}]) 1
- THEN (atac 1)
- THEN print_tac "after prove parameter call"
-
+ THEN asm_full_simp_tac
+ (HOL_basic_ss' addsimps [@{thm "split_eta"}, @{thm "split_beta"}, @{thm "fst_conv"},
+ @{thm "snd_conv"}, @{thm pair_collapse}]) 1
+ THEN (atac 1)
+ THEN print_tac "after prove parameter call"
+
fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st;
@@ -1743,9 +1725,9 @@
val (in_ts, clause_out_ts) = split_smode is ts;
fun prove_prems out_ts [] =
(prove_match thy out_ts)
- THEN print_tac "before simplifying assumptions"
+ THEN print_tac "before simplifying assumptions"
THEN asm_full_simp_tac HOL_basic_ss' 1
- THEN print_tac "before single intro rule"
+ THEN print_tac "before single intro rule"
THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
| prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
let
@@ -1811,7 +1793,7 @@
val pred_case_rule = the_elim_of thy pred
in
REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
- THEN print_tac' options "before applying elim rule"
+ THEN print_tac' options "before applying elim rule"
THEN etac (predfun_elim_of thy pred mode) 1
THEN etac pred_case_rule 1
THEN (EVERY (map
@@ -1929,7 +1911,7 @@
THEN (print_tac "state before assumption matching")
THEN (REPEAT (atac 1 ORELSE
(CHANGED (asm_full_simp_tac (HOL_basic_ss' addsimps
- [@{thm split_eta}, @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}]) 1)
+ [@{thm split_eta}, @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}]) 1)
THEN print_tac "state after simp_tac:"))))
| prove_prems2 out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
let
@@ -2005,7 +1987,7 @@
(if not (skip_proof options) then
(fn _ =>
rtac @{thm pred_iffI} 1
- THEN print_tac' options "after pred_iffI"
+ THEN print_tac' options "after pred_iffI"
THEN prove_one_direction options thy clauses preds modes pred mode moded_clauses
THEN print_tac' options "proved one direction"
THEN prove_other_direction options thy modes pred mode moded_clauses
@@ -2169,7 +2151,7 @@
val moded_clauses = #infer_modes steps options thy extra_modes all_modes param_vs clauses
val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
val _ = check_expected_modes options modes
- val _ = print_modes modes
+ val _ = print_modes options modes
(*val _ = print_moded_clauses thy moded_clauses*)
val _ = print_step options "Defining executable functions..."
val thy' = fold (#create_definitions steps preds) modes thy
@@ -2291,7 +2273,7 @@
val add_quickcheck_equations = gen_add_equations
{infer_modes = infer_modes_with_generator,
create_definitions = rpred_create_definitions,
- compile_preds = compile_preds rpred_comp_modifiers RPredCompFuns.compfuns,
+ compile_preds = compile_preds rpred_comp_modifiers RandomPredCompFuns.compfuns,
prove = prove_by_skip,
are_not_defined = fn thy => forall (null o rpred_modes_of thy),
qname = "rpred_equation"}
@@ -2420,7 +2402,7 @@
fun eval thy (options as (depth_limit, random)) t_compr =
let
- val compfuns = if random then RPredCompFuns.compfuns else PredicateCompFuns.compfuns
+ val compfuns = if random then RandomPredCompFuns.compfuns else PredicateCompFuns.compfuns
val t = analyze_compr thy compfuns options t_compr;
val T = dest_predT compfuns (fastype_of t);
val t' = mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t;
@@ -2467,12 +2449,10 @@
val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
-val _ = List.app OuterKeyword.keyword ["depth_limit", "random"]
-
val options =
let
- val depth_limit = Scan.optional (P.$$$ "depth_limit" |-- P.$$$ "=" |-- P.nat >> SOME) NONE
- val random = Scan.optional (P.$$$ "random" >> K true) false
+ val depth_limit = Scan.optional (Args.$$$ "depth_limit" |-- P.$$$ "=" |-- P.nat >> SOME) NONE
+ val random = Scan.optional (Args.$$$ "random" >> K true) false
in
Scan.optional (P.$$$ "[" |-- depth_limit -- random --| P.$$$ "]") (NONE, false)
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Wed Oct 28 11:43:06 2009 +0000
@@ -0,0 +1,219 @@
+(* Title: HOL/Tools/Predicate_Compile/predicate_compile_data.ML
+ Author: Lukas Bulwahn, TU Muenchen
+
+Book-keeping datastructure for the predicate compiler.
+*)
+
+signature PREDICATE_COMPILE_DATA =
+sig
+ type specification_table;
+ val make_const_spec_table : theory -> specification_table
+ val get_specification : specification_table -> string -> thm list
+ val obtain_specification_graph : Predicate_Compile_Aux.options -> theory ->
+ specification_table -> string -> thm list Graph.T
+ val normalize_equation : theory -> thm -> thm
+end;
+
+structure Predicate_Compile_Data : PREDICATE_COMPILE_DATA =
+struct
+
+open Predicate_Compile_Aux;
+
+structure Data = TheoryDataFun
+(
+ type T =
+ {const_spec_table : thm list Symtab.table};
+ val empty =
+ {const_spec_table = Symtab.empty};
+ val copy = I;
+ val extend = I;
+ fun merge _
+ ({const_spec_table = const_spec_table1},
+ {const_spec_table = const_spec_table2}) =
+ {const_spec_table = Symtab.merge (K true) (const_spec_table1, const_spec_table2)}
+);
+
+fun mk_data c = {const_spec_table = c}
+fun map_data f {const_spec_table = c} = mk_data (f c)
+
+type specification_table = thm list Symtab.table
+
+fun defining_const_of_introrule_term t =
+ let
+ val _ $ u = Logic.strip_imp_concl t
+ val (pred, all_args) = strip_comb u
+ in case pred of
+ Const (c, T) => c
+ | _ => raise TERM ("defining_const_of_introrule_term failed: Not a constant", [t])
+ end
+
+val defining_const_of_introrule = defining_const_of_introrule_term o prop_of
+
+(*TODO*)
+fun is_introlike_term t = true
+
+val is_introlike = is_introlike_term o prop_of
+
+fun check_equation_format_term (t as (Const ("==", _) $ u $ v)) =
+ (case strip_comb u of
+ (Const (c, T), args) =>
+ if (length (binder_types T) = length args) then
+ true
+ else
+ raise TERM ("check_equation_format_term failed: Number of arguments mismatch", [t])
+ | _ => raise TERM ("check_equation_format_term failed: Not a constant", [t]))
+ | check_equation_format_term t =
+ raise TERM ("check_equation_format_term failed: Not an equation", [t])
+
+val check_equation_format = check_equation_format_term o prop_of
+
+fun defining_const_of_equation_term (t as (Const ("==", _) $ u $ v)) =
+ (case fst (strip_comb u) of
+ Const (c, _) => c
+ | _ => raise TERM ("defining_const_of_equation_term failed: Not a constant", [t]))
+ | defining_const_of_equation_term t =
+ raise TERM ("defining_const_of_equation_term failed: Not an equation", [t])
+
+val defining_const_of_equation = defining_const_of_equation_term o prop_of
+
+(* Normalizing equations *)
+
+fun mk_meta_equation th =
+ case prop_of th of
+ Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _) => th RS @{thm eq_reflection}
+ | _ => th
+
+val meta_fun_cong = @{lemma "f == g ==> f x == g x" by simp}
+
+fun full_fun_cong_expand th =
+ let
+ val (f, args) = strip_comb (fst (Logic.dest_equals (prop_of th)))
+ val i = length (binder_types (fastype_of f)) - length args
+ in funpow i (fn th => th RS meta_fun_cong) th end;
+
+fun declare_names s xs ctxt =
+ let
+ val res = Name.names ctxt s xs
+ in (res, fold Name.declare (map fst res) ctxt) end
+
+fun split_all_pairs thy th =
+ let
+ val ctxt = ProofContext.init thy
+ val ((_, [th']), ctxt') = Variable.import true [th] ctxt
+ val t = prop_of th'
+ val frees = Term.add_frees t []
+ val freenames = Term.add_free_names t []
+ val nctxt = Name.make_context freenames
+ fun mk_tuple_rewrites (x, T) nctxt =
+ let
+ val Ts = HOLogic.flatten_tupleT T
+ val (xTs, nctxt') = declare_names x Ts nctxt
+ val paths = HOLogic.flat_tupleT_paths T
+ in ((Free (x, T), HOLogic.mk_ptuple paths T (map Free xTs)), nctxt') end
+ val (rewr, _) = fold_map mk_tuple_rewrites frees nctxt
+ val t' = Pattern.rewrite_term thy rewr [] t
+ val tac = Skip_Proof.cheat_tac thy
+ val th'' = Goal.prove ctxt (Term.add_free_names t' []) [] t' (fn {...} => tac)
+ val th''' = LocalDefs.unfold ctxt [@{thm split_conv}] th''
+ in
+ th'''
+ end;
+
+fun normalize_equation thy th =
+ mk_meta_equation th
+ |> Predicate_Compile_Set.unfold_set_notation
+ |> full_fun_cong_expand
+ |> split_all_pairs thy
+ |> tap check_equation_format
+
+fun inline_equations thy th = Conv.fconv_rule (Simplifier.rewrite
+((Simplifier.theory_context thy Simplifier.empty_ss) addsimps (Predicate_Compile_Inline_Defs.get (ProofContext.init thy)))) th
+
+fun store_thm_in_table ignore_consts thy th=
+ let
+ val th = th
+ |> inline_equations thy
+ |> Predicate_Compile_Set.unfold_set_notation
+ |> AxClass.unoverload thy
+ val (const, th) =
+ if is_equationlike th then
+ let
+ val eq = normalize_equation thy th
+ in
+ (defining_const_of_equation eq, eq)
+ end
+ else if (is_introlike th) then (defining_const_of_introrule th, th)
+ else error "store_thm: unexpected definition format"
+ in
+ if not (member (op =) ignore_consts const) then
+ Symtab.cons_list (const, th)
+ else I
+ end
+
+fun make_const_spec_table thy =
+ let
+ fun store ignore_const f = fold (store_thm_in_table ignore_const thy) (map (Thm.transfer thy) (f (ProofContext.init thy)))
+ val table = Symtab.empty
+ |> store [] Predicate_Compile_Alternative_Defs.get
+ val ignore_consts = Symtab.keys table
+ in
+ table
+ |> store ignore_consts Predicate_Compile_Preproc_Const_Defs.get
+ |> store ignore_consts Nitpick_Simps.get
+ |> store ignore_consts Nitpick_Intros.get
+ end
+
+fun get_specification table constname =
+ case Symtab.lookup table constname of
+ SOME thms => thms
+ | NONE => error ("get_specification: lookup of constant " ^ quote constname ^ " failed")
+
+val logic_operator_names =
+ [@{const_name "=="}, @{const_name "op ="}, @{const_name "op -->"}, @{const_name "All"}, @{const_name "Ex"},
+ @{const_name "op &"}]
+
+val special_cases = member (op =) [
+ @{const_name "False"},
+ @{const_name "Suc"}, @{const_name Nat.zero_nat_inst.zero_nat},
+ @{const_name Nat.one_nat_inst.one_nat},
+@{const_name "HOL.ord_class.less"}, @{const_name "HOL.ord_class.less_eq"}, @{const_name "HOL.zero_class.zero"},
+@{const_name "HOL.one_class.one"}, @{const_name HOL.plus_class.plus},
+@{const_name Nat.ord_nat_inst.less_eq_nat},
+@{const_name number_nat_inst.number_of_nat},
+ @{const_name Int.Bit0},
+ @{const_name Int.Bit1},
+ @{const_name Int.Pls},
+@{const_name "Int.zero_int_inst.zero_int"},
+@{const_name "List.filter"}]
+
+fun case_consts thy s = is_some (Datatype.info_of_case thy s)
+
+fun print_specification options thy constname specs =
+ if show_intermediate_results options then
+ tracing ("Specification of " ^ constname ^ ":\n" ^
+ cat_lines (map (Display.string_of_thm_global thy) specs))
+ else ()
+
+fun obtain_specification_graph options thy table constname =
+ let
+ fun is_nondefining_constname c = member (op =) logic_operator_names c
+ val is_defining_constname = member (op =) (Symtab.keys table)
+ fun has_code_pred_intros c = is_some (try (Predicate_Compile_Core.intros_of thy) c)
+ fun defiants_of specs =
+ fold (Term.add_const_names o prop_of) specs []
+ |> filter is_defining_constname
+ |> filter_out is_nondefining_constname
+ |> filter_out has_code_pred_intros
+ |> filter_out (case_consts thy)
+ |> filter_out special_cases
+ fun extend constname =
+ let
+ val specs = get_specification table constname
+ val _ = print_specification options thy constname specs
+ in (specs, defiants_of specs) end;
+ in
+ Graph.extend extend constname Graph.empty
+ end;
+
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Wed Oct 28 11:43:06 2009 +0000
@@ -0,0 +1,438 @@
+(* Title: HOL/Tools/Predicate_Compile/predicate_compile_fun.ML
+ Author: Lukas Bulwahn, TU Muenchen
+
+Preprocessing functions to predicates.
+*)
+
+signature PREDICATE_COMPILE_FUN =
+sig
+val define_predicates : (string * thm list) list -> theory -> (string * thm list) list * theory
+ val rewrite_intro : theory -> thm -> thm list
+ val setup_oracle : theory -> theory
+ val pred_of_function : theory -> string -> string option
+end;
+
+structure Predicate_Compile_Fun : PREDICATE_COMPILE_FUN =
+struct
+
+
+(* Oracle for preprocessing *)
+
+val (oracle : (string * (cterm -> thm)) option Unsynchronized.ref) = Unsynchronized.ref NONE;
+
+fun the_oracle () =
+ case !oracle of
+ NONE => error "Oracle is not setup"
+ | SOME (_, oracle) => oracle
+
+val setup_oracle = Thm.add_oracle (Binding.name "pred_compile_preprocessing", I) #->
+ (fn ora => fn thy => let val _ = (oracle := SOME ora) in thy end)
+
+
+fun is_funtype (Type ("fun", [_, _])) = true
+ | is_funtype _ = false;
+
+fun is_Type (Type _) = true
+ | is_Type _ = false
+
+(* returns true if t is an application of an datatype constructor *)
+(* which then consequently would be splitted *)
+(* else false *)
+(*
+fun is_constructor thy t =
+ if (is_Type (fastype_of t)) then
+ (case DatatypePackage.get_datatype thy ((fst o dest_Type o fastype_of) t) of
+ NONE => false
+ | SOME info => (let
+ val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
+ val (c, _) = strip_comb t
+ in (case c of
+ Const (name, _) => name mem_string constr_consts
+ | _ => false) end))
+ else false
+*)
+
+(* must be exported in code.ML *)
+fun is_constr thy = is_some o Code.get_datatype_of_constr thy;
+
+(* Table from constant name (string) to term of inductive predicate *)
+structure Pred_Compile_Preproc = TheoryDataFun
+(
+ type T = string Symtab.table;
+ val empty = Symtab.empty;
+ val copy = I;
+ val extend = I;
+ fun merge _ = Symtab.merge (op =);
+)
+
+fun pred_of_function thy name = Symtab.lookup (Pred_Compile_Preproc.get thy) name
+
+fun defined thy = Symtab.defined (Pred_Compile_Preproc.get thy)
+
+
+fun transform_ho_typ (T as Type ("fun", _)) =
+ let
+ val (Ts, T') = strip_type T
+ in if T' = @{typ "bool"} then T else (Ts @ [T']) ---> HOLogic.boolT end
+| transform_ho_typ t = t
+
+fun transform_ho_arg arg =
+ case (fastype_of arg) of
+ (T as Type ("fun", _)) =>
+ (case arg of
+ Free (name, _) => Free (name, transform_ho_typ T)
+ | _ => error "I am surprised")
+| _ => arg
+
+fun pred_type T =
+ let
+ val (Ts, T') = strip_type T
+ val Ts' = map transform_ho_typ Ts
+ in
+ (Ts' @ [T']) ---> HOLogic.boolT
+ end;
+
+(* FIXME: create new predicate name -- does not avoid nameclashing *)
+fun pred_of f =
+ let
+ val (name, T) = dest_Const f
+ in
+ if (body_type T = @{typ bool}) then
+ (Free (Long_Name.base_name name ^ "P", T))
+ else
+ (Free (Long_Name.base_name name ^ "P", pred_type T))
+ end
+
+fun mk_param thy lookup_pred (t as Free (v, _)) = lookup_pred t
+ | mk_param thy lookup_pred t =
+ let
+ val _ = tracing ("called param with " ^ (Syntax.string_of_term_global thy t))
+ in if Predicate_Compile_Aux.is_predT (fastype_of t) then
+ t
+ else
+ let
+ val (vs, body) = strip_abs t
+ val names = Term.add_free_names body []
+ val vs_names = Name.variant_list names (map fst vs)
+ val vs' = map2 (curry Free) vs_names (map snd vs)
+ val body' = subst_bounds (rev vs', body)
+ val (f, args) = strip_comb body'
+ val resname = Name.variant (vs_names @ names) "res"
+ val resvar = Free (resname, body_type (fastype_of body'))
+ (*val P = case try lookup_pred f of SOME P => P | NONE => error "mk_param"
+ val pred_body = list_comb (P, args @ [resvar])
+ *)
+ val pred_body = HOLogic.mk_eq (body', resvar)
+ val param = fold_rev lambda (vs' @ [resvar]) pred_body
+ in param end
+ end
+(* creates the list of premises for every intro rule *)
+(* theory -> term -> (string list, term list list) *)
+
+fun dest_code_eqn eqn = let
+ val (lhs, rhs) = Logic.dest_equals (Logic.unvarify (Thm.prop_of eqn))
+ val (func, args) = strip_comb lhs
+in ((func, args), rhs) end;
+
+fun string_of_typ T = Syntax.string_of_typ_global @{theory} T
+
+fun string_of_term t =
+ case t of
+ Const (c, T) => "Const (" ^ c ^ ", " ^ string_of_typ T ^ ")"
+ | Free (c, T) => "Free (" ^ c ^ ", " ^ string_of_typ T ^ ")"
+ | Var ((c, i), T) => "Var ((" ^ c ^ ", " ^ string_of_int i ^ "), " ^ string_of_typ T ^ ")"
+ | Bound i => "Bound " ^ string_of_int i
+ | Abs (x, T, t) => "Abs (" ^ x ^ ", " ^ string_of_typ T ^ ", " ^ string_of_term t ^ ")"
+ | t1 $ t2 => "(" ^ string_of_term t1 ^ ") $ (" ^ string_of_term t2 ^ ")"
+
+fun ind_package_get_nparams thy name =
+ case try (Inductive.the_inductive (ProofContext.init thy)) name of
+ SOME (_, result) => length (Inductive.params_of (#raw_induct result))
+ | NONE => error ("No such predicate: " ^ quote name)
+
+(* TODO: does not work with higher order functions yet *)
+fun mk_rewr_eq (func, pred) =
+ let
+ val (argTs, resT) = (strip_type (fastype_of func))
+ val nctxt =
+ Name.make_context (Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) (func $ pred) [])
+ val (argnames, nctxt') = Name.variants (replicate (length argTs) "a") nctxt
+ val ([resname], nctxt'') = Name.variants ["r"] nctxt'
+ val args = map Free (argnames ~~ argTs)
+ val res = Free (resname, resT)
+ in Logic.mk_equals
+ (HOLogic.mk_eq (res, list_comb (func, args)), list_comb (pred, args @ [res]))
+ end;
+
+fun has_split_rule_cname @{const_name "nat_case"} = true
+ | has_split_rule_cname @{const_name "list_case"} = true
+ | has_split_rule_cname _ = false
+
+fun has_split_rule_term thy (Const (@{const_name "nat_case"}, _)) = true
+ | has_split_rule_term thy (Const (@{const_name "list_case"}, _)) = true
+ | has_split_rule_term thy _ = false
+
+fun has_split_rule_term' thy (Const (@{const_name "If"}, _)) = true
+ | has_split_rule_term' thy (Const (@{const_name "Let"}, _)) = true
+ | has_split_rule_term' thy c = has_split_rule_term thy c
+
+fun prepare_split_thm ctxt split_thm =
+ (split_thm RS @{thm iffD2})
+ |> LocalDefs.unfold ctxt [@{thm atomize_conjL[symmetric]},
+ @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}]
+
+fun find_split_thm thy (Const (name, typ)) =
+ let
+ fun split_name str =
+ case first_field "." str
+ of (SOME (field, rest)) => field :: split_name rest
+ | NONE => [str]
+ val splitted_name = split_name name
+ in
+ if length splitted_name > 0 andalso
+ String.isSuffix "_case" (List.last splitted_name)
+ then
+ (List.take (splitted_name, length splitted_name - 1)) @ ["split"]
+ |> space_implode "."
+ |> PureThy.get_thm thy
+ |> SOME
+ handle ERROR msg => NONE
+ else NONE
+ end
+ | find_split_thm _ _ = NONE
+
+fun find_split_thm' thy (Const (@{const_name "If"}, _)) = SOME @{thm split_if}
+ | find_split_thm' thy (Const (@{const_name "Let"}, _)) = SOME @{thm refl} (* TODO *)
+ | find_split_thm' thy c = find_split_thm thy c
+
+fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t)
+
+fun folds_map f xs y =
+ let
+ fun folds_map' acc [] y = [(rev acc, y)]
+ | folds_map' acc (x :: xs) y =
+ maps (fn (x, y) => folds_map' (x :: acc) xs y) (f x y)
+ in
+ folds_map' [] xs y
+ end;
+
+fun mk_prems thy (lookup_pred, get_nparams) t (names, prems) =
+ let
+ fun mk_prems' (t as Const (name, T)) (names, prems) =
+ if is_constr thy name orelse (is_none (try lookup_pred t)) then
+ [(t, (names, prems))]
+ else [(lookup_pred t, (names, prems))]
+ | mk_prems' (t as Free (f, T)) (names, prems) =
+ [(lookup_pred t, (names, prems))]
+ | mk_prems' (t as Abs _) (names, prems) =
+ if Predicate_Compile_Aux.is_predT (fastype_of t) then
+ [(t, (names, prems))] else error "mk_prems': Abs "
+ (* mk_param *)
+ | mk_prems' t (names, prems) =
+ if Predicate_Compile_Aux.is_constrt thy t then
+ [(t, (names, prems))]
+ else
+ if has_split_rule_term' thy (fst (strip_comb t)) then
+ let
+ val (f, args) = strip_comb t
+ val split_thm = prepare_split_thm (ProofContext.init thy) (the (find_split_thm' thy f))
+ (* TODO: contextify things - this line is to unvarify the split_thm *)
+ (*val ((_, [isplit_thm]), _) = Variable.import true [split_thm] (ProofContext.init thy)*)
+ val (assms, concl) = Logic.strip_horn (Thm.prop_of split_thm)
+ val (P, [split_t]) = strip_comb (HOLogic.dest_Trueprop concl)
+ val subst = Pattern.match thy (split_t, t) (Vartab.empty, Vartab.empty)
+ val (_, split_args) = strip_comb split_t
+ val match = split_args ~~ args
+ fun mk_prems_of_assm assm =
+ let
+ val (vTs, assm') = strip_all (Envir.beta_norm (Envir.subst_term subst assm))
+ val var_names = Name.variant_list names (map fst vTs)
+ val vars = map Free (var_names ~~ (map snd vTs))
+ val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm'))
+ val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res)
+ in
+ mk_prems' inner_t (var_names @ names, prems' @ prems)
+ end
+ in
+ maps mk_prems_of_assm assms
+ end
+ else
+ let
+ val (f, args) = strip_comb t
+ (* TODO: special procedure for higher-order functions: split arguments in
+ simple types and function types *)
+ val resname = Name.variant names "res"
+ val resvar = Free (resname, body_type (fastype_of t))
+ val names' = resname :: names
+ fun mk_prems'' (t as Const (c, _)) =
+ if is_constr thy c orelse (is_none (try lookup_pred t)) then
+ folds_map mk_prems' args (names', prems) |>
+ map
+ (fn (argvs, (names'', prems')) =>
+ let
+ val prem = HOLogic.mk_Trueprop (HOLogic.mk_eq (resvar, list_comb (f, argvs)))
+ in (names'', prem :: prems') end)
+ else
+ let
+ val pred = lookup_pred t
+ val nparams = get_nparams pred
+ val (params, args) = chop nparams args
+ val params' = map (mk_param thy lookup_pred) params
+ in
+ folds_map mk_prems' args (names', prems)
+ |> map (fn (argvs, (names'', prems')) =>
+ let
+ val prem = HOLogic.mk_Trueprop (list_comb (pred, params' @ argvs @ [resvar]))
+ in (names'', prem :: prems') end)
+ end
+ | mk_prems'' (t as Free (_, _)) =
+ let
+ (* higher order argument call *)
+ val pred = lookup_pred t
+ in
+ folds_map mk_prems' args (resname :: names, prems)
+ |> map (fn (argvs, (names', prems')) =>
+ let
+ val prem = HOLogic.mk_Trueprop (list_comb (pred, argvs @ [resvar]))
+ in (names', prem :: prems') end)
+ end
+ | mk_prems'' t =
+ error ("Invalid term: " ^ Syntax.string_of_term_global thy t)
+ in
+ map (pair resvar) (mk_prems'' f)
+ end
+ in
+ mk_prems' t (names, prems)
+ end;
+
+(* assumption: mutual recursive predicates all have the same parameters. *)
+fun define_predicates specs thy =
+ if forall (fn (const, _) => member (op =) (Symtab.keys (Pred_Compile_Preproc.get thy)) const) specs then
+ ([], thy)
+ else
+ let
+ val consts = map fst specs
+ val eqns = maps snd specs
+ (*val eqns = maps (Predicate_Compile_Preproc_Data.get_specification thy) consts*)
+ (* create prednames *)
+ val ((funs, argss), rhss) = map_split dest_code_eqn eqns |>> split_list
+ val argss' = map (map transform_ho_arg) argss
+ val pnames = map dest_Free (distinct (op =) (maps (filter (is_funtype o fastype_of)) argss'))
+ val preds = map pred_of funs
+ val prednames = map (fst o dest_Free) preds
+ val funnames = map (fst o dest_Const) funs
+ val fun_pred_names = (funnames ~~ prednames)
+ (* mapping from term (Free or Const) to term *)
+ fun lookup_pred (Const (name, T)) =
+ (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of
+ SOME c => Const (c, pred_type T)
+ | NONE =>
+ (case AList.lookup op = fun_pred_names name of
+ SOME f => Free (f, pred_type T)
+ | NONE => Const (name, T)))
+ | lookup_pred (Free (name, T)) =
+ if member op = (map fst pnames) name then
+ Free (name, transform_ho_typ T)
+ else
+ Free (name, T)
+ | lookup_pred t =
+ error ("lookup function is not defined for " ^ Syntax.string_of_term_global thy t)
+
+ (* mapping from term (predicate term, not function term!) to int *)
+ fun get_nparams (Const (name, _)) =
+ the_default 0 (try (ind_package_get_nparams thy) name)
+ | get_nparams (Free (name, _)) =
+ (if member op = prednames name then
+ length pnames
+ else 0)
+ | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
+
+ (* create intro rules *)
+
+ fun mk_intros ((func, pred), (args, rhs)) =
+ if (body_type (fastype_of func) = @{typ bool}) then
+ (*TODO: preprocess predicate definition of rhs *)
+ [Logic.list_implies ([HOLogic.mk_Trueprop rhs], HOLogic.mk_Trueprop (list_comb (pred, args)))]
+ else
+ let
+ val names = Term.add_free_names rhs []
+ in mk_prems thy (lookup_pred, get_nparams) rhs (names, [])
+ |> map (fn (resultt, (names', prems)) =>
+ Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt]))))
+ end
+ fun mk_rewr_thm (func, pred) = @{thm refl}
+ in
+ case try (maps mk_intros) ((funs ~~ preds) ~~ (argss' ~~ rhss)) of
+ NONE => ([], thy)
+ | SOME intr_ts =>
+ if is_some (try (map (cterm_of thy)) intr_ts) then
+ let
+ val (ind_result, thy') =
+ Inductive.add_inductive_global (serial ())
+ {quiet_mode = false, verbose = false, kind = Thm.internalK,
+ alt_name = Binding.empty, coind = false, no_elim = false,
+ no_ind = false, skip_mono = false, fork_mono = false}
+ (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (distinct (op =) (map dest_Free preds)))
+ pnames
+ (map (fn x => (Attrib.empty_binding, x)) intr_ts)
+ [] thy
+ val prednames = map (fst o dest_Const) (#preds ind_result)
+ (* val rewr_thms = map mk_rewr_eq ((distinct (op =) funs) ~~ (#preds ind_result)) *)
+ (* add constants to my table *)
+ val specs = map (fn predname => (predname, filter (Predicate_Compile_Aux.is_intro predname) (#intrs ind_result))) prednames
+ val thy'' = Pred_Compile_Preproc.map (fold Symtab.update_new (consts ~~ prednames)) thy'
+ in
+ (specs, thy'')
+ end
+ else
+ let
+ val _ = tracing "Introduction rules of function_predicate are not welltyped"
+ in ([], thy) end
+ end
+
+(* preprocessing intro rules - uses oracle *)
+
+(* theory -> thm -> thm *)
+fun rewrite_intro thy intro =
+ let
+ fun lookup_pred (Const (name, T)) =
+ (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of
+ SOME c => Const (c, pred_type T)
+ | NONE => error ("Function " ^ name ^ " is not inductified"))
+ | lookup_pred (Free (name, T)) = Free (name, T)
+ | lookup_pred _ = error "lookup function is not defined!"
+
+ fun get_nparams (Const (name, _)) =
+ the_default 0 (try (ind_package_get_nparams thy) name)
+ | get_nparams (Free _) = 0
+ | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
+
+ val intro_t = (Logic.unvarify o prop_of) intro
+ val (prems, concl) = Logic.strip_horn intro_t
+ val frees = map fst (Term.add_frees intro_t [])
+ fun rewrite prem names =
+ let
+ val t = (HOLogic.dest_Trueprop prem)
+ val (lit, mk_lit) = case try HOLogic.dest_not t of
+ SOME t => (t, HOLogic.mk_not)
+ | NONE => (t, I)
+ val (P, args) = (strip_comb lit)
+ in
+ folds_map (
+ fn t => if (is_funtype (fastype_of t)) then (fn x => [(t, x)])
+ else mk_prems thy (lookup_pred, get_nparams) t) args (names, [])
+ |> map (fn (resargs, (names', prems')) =>
+ let
+ val prem' = HOLogic.mk_Trueprop (mk_lit (list_comb (P, resargs)))
+ in (prem'::prems', names') end)
+ end
+ val intro_ts' = folds_map rewrite prems frees
+ |> maps (fn (prems', frees') =>
+ rewrite concl frees'
+ |> map (fn (concl'::conclprems, _) =>
+ Logic.list_implies ((flat prems') @ conclprems, concl')))
+ in
+ map (Drule.standard o the_oracle () o cterm_of thy) intro_ts'
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Wed Oct 28 11:43:06 2009 +0000
@@ -0,0 +1,191 @@
+(* Title: HOL/Tools/Predicate_Compile/predicate_compile_pred.ML
+ Author: Lukas Bulwahn, TU Muenchen
+
+Preprocessing definitions of predicates to introduction rules.
+*)
+
+signature PREDICATE_COMPILE_PRED =
+sig
+ (* preprocesses an equation to a set of intro rules; defines new constants *)
+ (*
+ val preprocess_pred_equation : thm -> theory -> thm list * theory
+ *)
+ val preprocess : string -> theory -> (thm list list * theory)
+ (* output is the term list of clauses of an unknown predicate *)
+ val preprocess_term : term -> theory -> (term list * theory)
+
+ (*val rewrite : thm -> thm*)
+
+end;
+
+(* : PREDICATE_COMPILE_PREPROC_PRED *) (* FIXME *)
+structure Predicate_Compile_Pred =
+struct
+
+open Predicate_Compile_Aux
+
+fun is_compound ((Const ("Not", _)) $ t) =
+ error "is_compound: Negation should not occur; preprocessing is defect"
+ | is_compound ((Const ("Ex", _)) $ _) = true
+ | is_compound ((Const (@{const_name "op |"}, _)) $ _ $ _) = true
+ | is_compound ((Const ("op &", _)) $ _ $ _) =
+ error "is_compound: Conjunction should not occur; preprocessing is defect"
+ | is_compound _ = false
+
+fun flatten constname atom (defs, thy) =
+ if is_compound atom then
+ let
+ val constname = Name.variant (map (Long_Name.base_name o fst) defs)
+ ((Long_Name.base_name constname) ^ "_aux")
+ val full_constname = Sign.full_bname thy constname
+ val (params, args) = List.partition (is_predT o fastype_of)
+ (map Free (Term.add_frees atom []))
+ val constT = map fastype_of (params @ args) ---> HOLogic.boolT
+ val lhs = list_comb (Const (full_constname, constT), params @ args)
+ val def = Logic.mk_equals (lhs, atom)
+ val ([definition], thy') = thy
+ |> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)]
+ |> PureThy.add_defs false [((Binding.name (constname ^ "_def"), def), [])]
+ in
+ (lhs, ((full_constname, [definition]) :: defs, thy'))
+ end
+ else
+ (atom, (defs, thy))
+
+fun flatten_intros constname intros thy =
+ let
+ val ctxt = ProofContext.init thy
+ val ((_, intros), ctxt') = Variable.import true intros ctxt
+ val (intros', (local_defs, thy')) = (fold_map o fold_map_atoms)
+ (flatten constname) (map prop_of intros) ([], thy)
+ val tac = fn _ => Skip_Proof.cheat_tac thy'
+ val intros'' = map (fn t => Goal.prove ctxt' [] [] t tac) intros'
+ |> Variable.export ctxt' ctxt
+ in
+ (intros'', (local_defs, thy'))
+ end
+
+(* TODO: same function occurs in inductive package *)
+fun select_disj 1 1 = []
+ | select_disj _ 1 = [rtac @{thm disjI1}]
+ | select_disj n i = (rtac @{thm disjI2})::(select_disj (n - 1) (i - 1));
+
+fun introrulify thy ths =
+ let
+ val ctxt = ProofContext.init thy
+ val ((_, ths'), ctxt') = Variable.import true ths ctxt
+ fun introrulify' th =
+ let
+ val (lhs, rhs) = Logic.dest_equals (prop_of th)
+ val frees = Term.add_free_names rhs []
+ val disjuncts = HOLogic.dest_disj rhs
+ val nctxt = Name.make_context frees
+ fun mk_introrule t =
+ let
+ val ((ps, t'), nctxt') = focus_ex t nctxt
+ val prems = map HOLogic.mk_Trueprop (HOLogic.dest_conj t')
+ in
+ (ps, Logic.list_implies (prems, HOLogic.mk_Trueprop lhs))
+ end
+ val x = ((cterm_of thy) o the_single o snd o strip_comb o HOLogic.dest_Trueprop o fst o
+ Logic.dest_implies o prop_of) @{thm exI}
+ fun prove_introrule (index, (ps, introrule)) =
+ let
+ val tac = Simplifier.simp_tac (HOL_basic_ss addsimps [th]) 1
+ THEN EVERY1 (select_disj (length disjuncts) (index + 1))
+ THEN (EVERY (map (fn y =>
+ rtac (Drule.cterm_instantiate [(x, cterm_of thy (Free y))] @{thm exI}) 1) ps))
+ THEN REPEAT_DETERM (rtac @{thm conjI} 1 THEN atac 1)
+ THEN TRY (atac 1)
+ in
+ Goal.prove ctxt' (map fst ps) [] introrule (fn {...} => tac)
+ end
+ in
+ map_index prove_introrule (map mk_introrule disjuncts)
+ end
+ in maps introrulify' ths' |> Variable.export ctxt' ctxt end
+
+val rewrite =
+ Simplifier.simplify (HOL_basic_ss addsimps [@{thm Ball_def}, @{thm Bex_def}])
+ #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm all_not_ex}])
+ #> Conv.fconv_rule nnf_conv
+ #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm ex_disj_distrib}])
+
+val rewrite_intros =
+ Simplifier.simplify (HOL_basic_ss addsimps @{thms HOL.simp_thms(9)})
+
+fun preprocess (constname, specs) thy =
+ let
+ val ctxt = ProofContext.init thy
+ val intros =
+ if forall is_pred_equation specs then
+ introrulify thy (map rewrite specs)
+ else if forall (is_intro constname) specs then
+ map rewrite_intros specs
+ else
+ error ("unexpected specification for constant " ^ quote constname ^ ":\n"
+ ^ commas (map (quote o Display.string_of_thm_global thy) specs))
+ val (intros', (local_defs, thy')) = flatten_intros constname intros thy
+ val (intross, thy'') = fold_map preprocess local_defs thy'
+ in
+ ((constname, intros') :: flat intross,thy'')
+ end;
+
+fun preprocess_term t thy = error "preprocess_pred_term: to implement"
+
+fun is_Abs (Abs _) = true
+ | is_Abs _ = false
+
+fun flat_higher_order_arguments (intross, thy) =
+ let
+ fun process constname atom (new_defs, thy) =
+ let
+ val (pred, args) = strip_comb atom
+ val abs_args = filter is_Abs args
+ fun replace_abs_arg (abs_arg as Abs _ ) (new_defs, thy) =
+ let
+ val _ = tracing ("Introduce new constant for " ^
+ Syntax.string_of_term_global thy abs_arg)
+ val vars = map Var (Term.add_vars abs_arg [])
+ val abs_arg' = Logic.unvarify abs_arg
+ val frees = map Free (Term.add_frees abs_arg' [])
+ val constname = Name.variant (map (Long_Name.base_name o fst) new_defs)
+ ((Long_Name.base_name constname) ^ "_hoaux")
+ val full_constname = Sign.full_bname thy constname
+ val constT = map fastype_of frees ---> (fastype_of abs_arg')
+ val const = Const (full_constname, constT)
+ val lhs = list_comb (const, frees)
+ val def = Logic.mk_equals (lhs, abs_arg')
+ val _ = tracing (Syntax.string_of_term_global thy def)
+ val ([definition], thy') = thy
+ |> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)]
+ |> PureThy.add_defs false [((Binding.name (constname ^ "_def"), def), [])]
+ in
+ (list_comb (Logic.varify const, vars), ((full_constname, [definition])::new_defs, thy'))
+ end
+ | replace_abs_arg arg (new_defs, thy) = (arg, (new_defs, thy))
+ val (args', (new_defs', thy')) = fold_map replace_abs_arg args (new_defs, thy)
+ in
+ (list_comb (pred, args'), (new_defs', thy'))
+ end
+ fun flat_intro intro (new_defs, thy) =
+ let
+ val constname = fst (dest_Const (fst (strip_comb
+ (HOLogic.dest_Trueprop (Logic.strip_imp_concl (prop_of intro))))))
+ val (intro_ts, (new_defs, thy)) = fold_map_atoms (process constname) (prop_of intro) (new_defs, thy)
+ val th = Skip_Proof.make_thm thy intro_ts
+ in
+ (th, (new_defs, thy))
+ end
+ fun fold_map_spec f [] s = ([], s)
+ | fold_map_spec f ((c, ths) :: specs) s =
+ let
+ val (ths', s') = f ths s
+ val (specs', s'') = fold_map_spec f specs s'
+ in ((c, ths') :: specs', s'') end
+ val (intross', (new_defs, thy')) = fold_map_spec (fold_map flat_intro) intross ([], thy)
+ in
+ (intross', (new_defs, thy'))
+ end
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Wed Oct 28 11:43:06 2009 +0000
@@ -0,0 +1,113 @@
+(* Title: HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML
+ Author: Lukas Bulwahn, TU Muenchen
+
+A quickcheck generator based on the predicate compiler.
+*)
+
+signature PREDICATE_COMPILE_QUICKCHECK =
+sig
+ val quickcheck : Proof.context -> term -> int -> term list option
+ val test_ref :
+ ((unit -> int -> int * int -> term list Predicate.pred * (int * int)) option) Unsynchronized.ref
+end;
+
+structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK =
+struct
+
+open Predicate_Compile_Aux;
+
+val test_ref =
+ Unsynchronized.ref (NONE : (unit -> int -> int * int -> term list Predicate.pred * (int * int)) option)
+
+val target = "Quickcheck"
+
+val options = Options {
+ expected_modes = NONE,
+ show_steps = true,
+ show_intermediate_results = true,
+ show_proof_trace = false,
+ show_modes = true,
+ show_mode_inference = false,
+ show_compilation = true,
+ skip_proof = false,
+
+ inductify = false,
+ rpred = false,
+ depth_limited = false
+}
+
+fun dest_compfuns (Predicate_Compile_Core.CompilationFuns funs) = funs
+val mk_predT = #mk_predT (dest_compfuns Predicate_Compile_Core.pred_compfuns)
+val mk_randompredT = #mk_predT (dest_compfuns Predicate_Compile_Core.randompred_compfuns)
+val mk_return = #mk_single (dest_compfuns Predicate_Compile_Core.randompred_compfuns)
+val mk_bind = #mk_bind (dest_compfuns Predicate_Compile_Core.randompred_compfuns)
+
+fun mk_split_lambda [] t = Abs ("u", HOLogic.unitT, t)
+ | mk_split_lambda [x] t = lambda x t
+ | mk_split_lambda xs t =
+ let
+ fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
+ | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
+ in
+ mk_split_lambda' xs t
+ end;
+
+fun strip_imp_prems (Const("op -->", _) $ A $ B) = A :: strip_imp_prems B
+ | strip_imp_prems _ = [];
+
+fun strip_imp_concl (Const("op -->", _) $ A $ B) = strip_imp_concl B
+ | strip_imp_concl A = A : term;
+
+fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
+
+fun quickcheck ctxt t =
+ let
+ val _ = tracing ("Starting quickcheck with " ^ (Syntax.string_of_term ctxt t))
+ val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
+ val thy = (ProofContext.theory_of ctxt')
+ val (vs, t') = strip_abs t
+ val vs' = Variable.variant_frees ctxt' [] vs
+ val t'' = subst_bounds (map Free (rev vs'), t')
+ val (prems, concl) = strip_horn t''
+ val constname = "pred_compile_quickcheck"
+ val full_constname = Sign.full_bname thy constname
+ val constT = map snd vs' ---> @{typ bool}
+ val thy' = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy
+ val t = Logic.list_implies
+ (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),
+ HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs')))
+ val tac = fn _ => Skip_Proof.cheat_tac thy'
+ val intro = Goal.prove (ProofContext.init thy') (map fst vs') [] t tac
+ val _ = tracing (Display.string_of_thm ctxt' intro)
+ val thy'' = thy'
+ |> Context.theory_map (Predicate_Compile_Preproc_Const_Defs.add_thm intro)
+ |> Predicate_Compile.preprocess options full_constname
+ |> Predicate_Compile_Core.add_equations options [full_constname]
+ (* |> Predicate_Compile_Core.add_depth_limited_equations Predicate_Compile_Aux.default_options [full_constname]*)
+ |> Predicate_Compile_Core.add_quickcheck_equations options [full_constname]
+ val depth_limited_modes = Predicate_Compile_Core.depth_limited_modes_of thy'' full_constname
+ val modes = Predicate_Compile_Core.generator_modes_of thy'' full_constname
+ val prog =
+ if member (op =) modes ([], []) then
+ let
+ val name = Predicate_Compile_Core.generator_name_of thy'' full_constname ([], [])
+ val T = [@{typ code_numeral}] ---> (mk_randompredT (HOLogic.mk_tupleT (map snd vs')))
+ in Const (name, T) $ Bound 0 end
+ (*else if member (op =) depth_limited_modes ([], []) then
+ let
+ val name = Predicate_Compile_Core.depth_limited_function_name_of thy'' full_constname ([], [])
+ val T = @{typ code_numeral} --> (mk_predT (HOLogic.mk_tupleT (map snd vs')))
+ in lift_pred (Const (name, T) $ Bound 0) end*)
+ else error "Predicate Compile Quickcheck failed"
+ val qc_term = Abs ("size", @{typ code_numeral}, mk_bind (prog,
+ mk_split_lambda (map Free vs') (mk_return (HOLogic.mk_list @{typ term}
+ (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))))
+ val _ = tracing (Syntax.string_of_term ctxt' qc_term)
+ val compile = Code_ML.eval (SOME target) ("Predicate_Compile_Quickcheck.test_ref", test_ref)
+ (fn proc => fn g => fn s => g s #>> (Predicate.map o map) proc)
+ thy'' qc_term []
+ in
+ ((compile #> Random_Engine.run) #> (Option.map fst o Predicate.yield))
+ end
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_set.ML Wed Oct 28 11:43:06 2009 +0000
@@ -0,0 +1,53 @@
+(* Title: HOL/Tools/Predicate_Compile/predicate_compile_set.ML
+ Author: Lukas Bulwahn, TU Muenchen
+
+Preprocessing sets to predicates.
+*)
+
+signature PREDICATE_COMPILE_SET =
+sig
+(*
+ val preprocess_intro : thm -> theory -> thm * theory
+ val preprocess_clause : term -> theory -> term * theory
+*)
+ val unfold_set_notation : thm -> thm;
+end;
+
+structure Predicate_Compile_Set : PREDICATE_COMPILE_SET =
+struct
+(*FIXME: unfolding Ball in pretty adhoc here *)
+val unfold_set_lemmas = [@{thm Collect_def}, @{thm mem_def},
+@{thm Ball_def}, @{thm Bex_def}]
+
+val unfold_set_notation = Simplifier.rewrite_rule unfold_set_lemmas
+
+(*
+fun find_set_theorems ctxt cname =
+ let
+ val _ = cname
+*)
+
+(*
+fun preprocess_term t ctxt =
+ case t of
+ Const ("op :", _) $ x $ A =>
+ case strip_comb A of
+ (Const (cname, T), params) =>
+ let
+ val _ = find_set_theorems
+ in
+ (t, ctxt)
+ end
+ | _ => (t, ctxt)
+ | _ => (t, ctxt)
+*)
+(*
+fun preprocess_intro th thy =
+ let
+ val cnames = find_heads_of_prems
+ val set_cname = filter (has_set_definition
+ val _ = define_preds
+ val _ = prep_pred_def
+ in
+*)
+end;
--- a/src/HOL/Tools/TFL/post.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/TFL/post.ML Wed Oct 28 11:43:06 2009 +0000
@@ -189,12 +189,11 @@
in
fun derive_init_eqs sgn rules eqs =
let
- val eqths = map (Thm.trivial o (Thm.cterm_of sgn) o HOLogic.mk_Trueprop)
- eqs
- fun countlist l =
- (rev o snd o (Library.foldl (fn ((i,L), e) => (i + 1,(e,i) :: L)))) ((0,[]), l)
+ val eqths = map (Thm.trivial o (Thm.cterm_of sgn) o HOLogic.mk_Trueprop) eqs
+ fun countlist l =
+ rev (snd (fold (fn e => fn (i, L) => (i + 1, (e, i) :: L)) l (0, [])))
in
- maps (fn (e,i) => solve_eq (e, (get_related_thms i rules), i)) (countlist eqths)
+ maps (fn (e, i) => solve_eq (e, (get_related_thms i rules), i)) (countlist eqths)
end;
end;
--- a/src/HOL/Tools/choice_specification.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/choice_specification.ML Wed Oct 28 11:43:06 2009 +0000
@@ -165,7 +165,7 @@
val cnames = map (fst o dest_Const) proc_consts
fun post_process (arg as (thy,thm)) =
let
- fun inst_all thy (thm,v) =
+ fun inst_all thy v thm =
let
val cv = cterm_of thy v
val cT = ctyp_of_term cv
@@ -174,7 +174,7 @@
thm RS spec'
end
fun remove_alls frees thm =
- Library.foldl (inst_all (Thm.theory_of_thm thm)) (thm,frees)
+ fold (inst_all (Thm.theory_of_thm thm)) frees thm
fun process_single ((name,atts),rew_imp,frees) args =
let
fun undo_imps thm =
--- a/src/HOL/Tools/hologic.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/hologic.ML Wed Oct 28 11:43:06 2009 +0000
@@ -638,8 +638,8 @@
val Ts = map snd vs;
val t = Const (c, Ts ---> T);
val tt = mk_prod (t, Abs ("u", unitT, reflect_term t));
- fun app (t, (fT, (v, T))) = valapp T fT $ t $ Free (v, termifyT T);
- in Library.foldl app (tt, mk_fTs Ts T ~~ vs) end;
+ fun app (fT, (v, T)) t = valapp T fT $ t $ Free (v, termifyT T);
+ in fold app (mk_fTs Ts T ~~ vs) tt end;
(* open state monads *)
--- a/src/HOL/Tools/inductive_codegen.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/inductive_codegen.ML Wed Oct 28 11:43:06 2009 +0000
@@ -41,7 +41,7 @@
fun warn thm = warning ("InductiveCodegen: Not a proper clause:\n" ^
Display.string_of_thm_without_context thm);
-fun add_node (g, x) = Graph.new_node (x, ()) g handle Graph.DUP _ => g;
+fun add_node x g = Graph.new_node (x, ()) g handle Graph.DUP _ => g;
fun add optmod optnparms = Thm.declaration_attribute (fn thm => Context.mapping (fn thy =>
let
@@ -72,7 +72,7 @@
Symtab.update (s, (AList.update Thm.eq_thm_prop
(thm, (thyname_of s, nparms)) rules)),
graph = List.foldr (uncurry (Graph.add_edge o pair s))
- (Library.foldl add_node (graph, s :: cs)) cs,
+ (fold add_node (s :: cs) graph) cs,
eqns = eqns} thy
end
| _ => (warn thm; thy))
@@ -266,19 +266,22 @@
flat (separate [str ",", Pretty.brk 1] (map single xs)) @
[str ")"]);
-fun mk_v ((names, vs), s) = (case AList.lookup (op =) vs s of
- NONE => ((names, (s, [s])::vs), s)
- | SOME xs => let val s' = Name.variant names s in
- ((s'::names, AList.update (op =) (s, s'::xs) vs), s') end);
+fun mk_v s (names, vs) =
+ (case AList.lookup (op =) vs s of
+ NONE => (s, (names, (s, [s])::vs))
+ | SOME xs =>
+ let val s' = Name.variant names s
+ in (s', (s'::names, AList.update (op =) (s, s'::xs) vs)) end);
-fun distinct_v (nvs, Var ((s, 0), T)) =
- apsnd (Var o rpair T o rpair 0) (mk_v (nvs, s))
- | distinct_v (nvs, t $ u) =
+fun distinct_v (Var ((s, 0), T)) nvs =
+ let val (s', nvs') = mk_v s nvs
+ in (Var ((s', 0), T), nvs') end
+ | distinct_v (t $ u) nvs =
let
- val (nvs', t') = distinct_v (nvs, t);
- val (nvs'', u') = distinct_v (nvs', u);
- in (nvs'', t' $ u') end
- | distinct_v x = x;
+ val (t', nvs') = distinct_v t nvs;
+ val (u', nvs'') = distinct_v u nvs';
+ in (t' $ u', nvs'') end
+ | distinct_v t nvs = (t, nvs);
fun is_exhaustive (Var _) = true
| is_exhaustive (Const ("Pair", _) $ t $ u) =
@@ -346,30 +349,29 @@
(fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
(arg_vs ~~ iss);
- fun check_constrt ((names, eqs), t) =
- if is_constrt thy t then ((names, eqs), t) else
+ fun check_constrt t (names, eqs) =
+ if is_constrt thy t then (t, (names, eqs))
+ else
let val s = Name.variant names "x";
- in ((s::names, (s, t)::eqs), Var ((s, 0), fastype_of t)) end;
+ in (Var ((s, 0), fastype_of t), (s::names, (s, t)::eqs)) end;
fun compile_eq (s, t) gr =
apfst (Pretty.block o cons (str (s ^ " = ")) o single)
(invoke_codegen thy defs dep module false t gr);
val (in_ts, out_ts) = get_args is 1 ts;
- val ((all_vs', eqs), in_ts') =
- Library.foldl_map check_constrt ((all_vs, []), in_ts);
+ val (in_ts', (all_vs', eqs)) = fold_map check_constrt in_ts (all_vs, []);
fun compile_prems out_ts' vs names [] gr =
let
- val (out_ps, gr2) = fold_map
- (invoke_codegen thy defs dep module false) out_ts gr;
+ val (out_ps, gr2) =
+ fold_map (invoke_codegen thy defs dep module false) out_ts gr;
val (eq_ps, gr3) = fold_map compile_eq eqs gr2;
- val ((names', eqs'), out_ts'') =
- Library.foldl_map check_constrt ((names, []), out_ts');
- val (nvs, out_ts''') = Library.foldl_map distinct_v
- ((names', map (fn x => (x, [x])) vs), out_ts'');
- val (out_ps', gr4) = fold_map
- (invoke_codegen thy defs dep module false) (out_ts''') gr3;
+ val (out_ts'', (names', eqs')) = fold_map check_constrt out_ts' (names, []);
+ val (out_ts''', nvs) =
+ fold_map distinct_v out_ts'' (names', map (fn x => (x, [x])) vs);
+ val (out_ps', gr4) =
+ fold_map (invoke_codegen thy defs dep module false) out_ts''' gr3;
val (eq_ps', gr5) = fold_map compile_eq eqs' gr4;
in
(compile_match (snd nvs) (eq_ps @ eq_ps') out_ps'
@@ -379,15 +381,11 @@
| compile_prems out_ts vs names ps gr =
let
val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
- val SOME (p, mode as SOME (Mode (_, js, _))) =
- select_mode_prem thy modes' vs' ps;
+ val SOME (p, mode as SOME (Mode (_, js, _))) = select_mode_prem thy modes' vs' ps;
val ps' = filter_out (equal p) ps;
- val ((names', eqs), out_ts') =
- Library.foldl_map check_constrt ((names, []), out_ts);
- val (nvs, out_ts'') = Library.foldl_map distinct_v
- ((names', map (fn x => (x, [x])) vs), out_ts');
- val (out_ps, gr0) = fold_map
- (invoke_codegen thy defs dep module false) out_ts'' gr;
+ val (out_ts', (names', eqs)) = fold_map check_constrt out_ts (names, []);
+ val (out_ts'', nvs) = fold_map distinct_v out_ts' (names', map (fn x => (x, [x])) vs);
+ val (out_ps, gr0) = fold_map (invoke_codegen thy defs dep module false) out_ts'' gr;
val (eq_ps, gr1) = fold_map compile_eq eqs gr0;
in
(case p of
@@ -480,19 +478,22 @@
fun prep_intrs intrs = map (rename_term o #prop o rep_thm o Drule.standard) intrs;
fun constrain cs [] = []
- | constrain cs ((s, xs) :: ys) = (s, case AList.lookup (op =) cs (s : string) of
- NONE => xs
- | SOME xs' => inter (op =) xs' xs) :: constrain cs ys;
+ | constrain cs ((s, xs) :: ys) =
+ (s,
+ case AList.lookup (op =) cs (s : string) of
+ NONE => xs
+ | SOME xs' => inter (op =) xs' xs) :: constrain cs ys;
fun mk_extra_defs thy defs gr dep names module ts =
- Library.foldl (fn (gr, name) =>
+ fold (fn name => fn gr =>
if name mem names then gr
- else (case get_clauses thy name of
+ else
+ (case get_clauses thy name of
NONE => gr
| SOME (names, thyname, nparms, intrs) =>
mk_ind_def thy defs gr dep names (if_library thyname module)
[] (prep_intrs intrs) nparms))
- (gr, fold Term.add_const_names ts [])
+ (fold Term.add_const_names ts []) gr
and mk_ind_def thy defs gr dep names module modecs intrs nparms =
add_edge_acyclic (hd names, dep) gr handle
@@ -562,17 +563,16 @@
val (ts1, ts2) = chop k ts;
val u = list_comb (Const (s, T), ts1);
- fun mk_mode (((ts, mode), i), Const ("dummy_pattern", _)) =
- ((ts, mode), i+1)
- | mk_mode (((ts, mode), i), t) = ((ts @ [t], mode @ [i]), i+1);
+ fun mk_mode (Const ("dummy_pattern", _)) ((ts, mode), i) = ((ts, mode), i + 1)
+ | mk_mode t ((ts, mode), i) = ((ts @ [t], mode @ [i]), i + 1);
val module' = if_library thyname module;
val gr1 = mk_extra_defs thy defs
(mk_ind_def thy defs gr dep names module'
[] (prep_intrs intrs) k) dep names module' [u];
val (modes, _) = lookup_modes gr1 dep;
- val (ts', is) = if is_query then
- fst (Library.foldl mk_mode ((([], []), 1), ts2))
+ val (ts', is) =
+ if is_query then fst (fold mk_mode ts2 (([], []), 1))
else (ts2, 1 upto length (binder_types T) - k);
val mode = find_mode gr1 dep s u modes is;
val (in_ps, gr2) = fold_map (invoke_codegen thy defs dep module true) ts' gr1;
@@ -697,8 +697,9 @@
val setup =
add_codegen "inductive" inductive_codegen #>
- Attrib.setup @{binding code_ind} (Scan.lift (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) --
- Scan.option (Args.$$$ "params" |-- Args.colon |-- OuterParse.nat) >> uncurry add))
+ Attrib.setup @{binding code_ind}
+ (Scan.lift (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) --
+ Scan.option (Args.$$$ "params" |-- Args.colon |-- OuterParse.nat) >> uncurry add))
"introduction rules for executable predicates";
end;
--- a/src/HOL/Tools/inductive_realizer.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/inductive_realizer.ML Wed Oct 28 11:43:06 2009 +0000
@@ -271,7 +271,7 @@
fun rename tab = map (fn x => the_default x (AList.lookup op = tab x));
-fun add_ind_realizer rsets intrs induct raw_induct elims (thy, vs) =
+fun add_ind_realizer rsets intrs induct raw_induct elims vs thy =
let
val qualifier = Long_Name.qualifier (name_of_thm induct);
val inducts = PureThy.get_thms thy (Long_Name.qualify qualifier "inducts");
@@ -367,7 +367,7 @@
SOME (fst (fst (dest_Var (head_of P)))) else NONE)
(HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct)));
- fun add_ind_realizer (thy, Ps) =
+ fun add_ind_realizer Ps thy =
let
val vs' = rename (map (pairself (fst o fst o dest_Var))
(params ~~ List.take (snd (strip_comb (HOLogic.dest_Trueprop
@@ -426,13 +426,12 @@
let
val (prem :: prems) = prems_of elim;
fun reorder1 (p, (_, intr)) =
- Library.foldl (fn (t, ((s, _), T)) => Logic.all (Free (s, T)) t)
- (strip_all p, subtract (op =) params' (Term.add_vars (prop_of intr) []));
+ fold (fn ((s, _), T) => Logic.all (Free (s, T)))
+ (subtract (op =) params' (Term.add_vars (prop_of intr) []))
+ (strip_all p);
fun reorder2 ((ivs, intr), i) =
let val fs = subtract (op =) params' (Term.add_vars (prop_of intr) [])
- in Library.foldl (fn (t, x) => lambda (Var x) t)
- (list_comb (Bound (i + length ivs), ivs), fs)
- end;
+ in fold (lambda o Var) fs (list_comb (Bound (i + length ivs), ivs)) end;
val p = Logic.list_implies
(map reorder1 (prems ~~ intrs) @ [prem], concl_of elim);
val T' = Extraction.etype_of thy (vs @ Ps) [] p;
@@ -465,7 +464,7 @@
(** add realizers to theory **)
- val thy4 = Library.foldl add_ind_realizer (thy3, subsets Ps);
+ val thy4 = fold add_ind_realizer (subsets Ps) thy3;
val thy5 = Extraction.add_realizers_i
(map (mk_realizer thy4 vs) (map (fn (((rule, rrule), rlz), c) =>
(name_of_thm rule, rule, rrule, rlz,
@@ -474,10 +473,11 @@
val elimps = map_filter (fn ((s, intrs), p) =>
if s mem rsets then SOME (p, intrs) else NONE)
(rss' ~~ (elims ~~ #elims ind_info));
- val thy6 = Library.foldl (fn (thy, p as (((((elim, _), _), _), _), _)) => thy |>
- add_elim_realizer [] p |> add_elim_realizer [fst (fst (dest_Var
- (HOLogic.dest_Trueprop (concl_of elim))))] p) (thy5,
- elimps ~~ case_thms ~~ case_names ~~ dummies)
+ val thy6 =
+ fold (fn p as (((((elim, _), _), _), _), _) =>
+ add_elim_realizer [] p #>
+ add_elim_realizer [fst (fst (dest_Var (HOLogic.dest_Trueprop (concl_of elim))))] p)
+ (elimps ~~ case_thms ~~ case_names ~~ dummies) thy5;
in Sign.restore_naming thy thy6 end;
@@ -488,7 +488,7 @@
val vss = sort (int_ord o pairself length)
(subsets (map fst (relevant_vars (concl_of (hd intrs)))))
in
- Library.foldl (add_ind_realizer rsets intrs induct raw_induct elims) (thy, vss)
+ fold (add_ind_realizer rsets intrs induct raw_induct elims) vss thy
end
fun rlz_attrib arg = Thm.declaration_attribute (fn thm => Context.mapping
--- a/src/HOL/Tools/meson.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/meson.ML Wed Oct 28 11:43:06 2009 +0000
@@ -323,7 +323,7 @@
Strips universal quantifiers and breaks up conjunctions.
Eliminates existential quantifiers using skoths: Skolemization theorems.*)
fun cnf skoths ctxt (th,ths) =
- let val ctxtr = Unsynchronized.ref ctxt
+ let val ctxtr = Unsynchronized.ref ctxt (* FIXME ??? *)
fun cnf_aux (th,ths) =
if not (can HOLogic.dest_Trueprop (prop_of th)) then ths (*meta-level: ignore*)
else if not (has_conns ["All","Ex","op &"] (prop_of th))
@@ -462,14 +462,13 @@
(** Detecting repeated assumptions in a subgoal **)
(*The stringtree detects repeated assumptions.*)
-fun ins_term (net,t) = Net.insert_term (op aconv) (t,t) net;
+fun ins_term t net = Net.insert_term (op aconv) (t, t) net;
(*detects repetitions in a list of terms*)
fun has_reps [] = false
| has_reps [_] = false
| has_reps [t,u] = (t aconv u)
- | has_reps ts = (Library.foldl ins_term (Net.empty, ts); false)
- handle Net.INSERT => true;
+ | has_reps ts = (fold ins_term ts Net.empty; false) handle Net.INSERT => true;
(*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*)
fun TRYING_eq_assume_tac 0 st = Seq.single st
--- a/src/HOL/Tools/metis_tools.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/metis_tools.ML Wed Oct 28 11:43:06 2009 +0000
@@ -21,8 +21,6 @@
val trace = Unsynchronized.ref false;
fun trace_msg msg = if ! trace then tracing (msg ()) else ();
-structure Recon = ResReconstruct;
-
val (type_lits, type_lits_setup) = Attrib.config_bool "metis_type_lits" true;
datatype mode = FO | HO | FT (*first-order, higher-order, fully-typed*)
@@ -211,14 +209,14 @@
| strip_happ args x = (x, args);
fun fol_type_to_isa _ (Metis.Term.Var v) =
- (case Recon.strip_prefix ResClause.tvar_prefix v of
- SOME w => Recon.make_tvar w
- | NONE => Recon.make_tvar v)
+ (case ResReconstruct.strip_prefix ResClause.tvar_prefix v of
+ SOME w => ResReconstruct.make_tvar w
+ | NONE => ResReconstruct.make_tvar v)
| fol_type_to_isa ctxt (Metis.Term.Fn(x, tys)) =
- (case Recon.strip_prefix ResClause.tconst_prefix x of
- SOME tc => Term.Type (Recon.invert_type_const tc, map (fol_type_to_isa ctxt) tys)
+ (case ResReconstruct.strip_prefix ResClause.tconst_prefix x of
+ SOME tc => Term.Type (ResReconstruct.invert_type_const tc, map (fol_type_to_isa ctxt) tys)
| NONE =>
- case Recon.strip_prefix ResClause.tfree_prefix x of
+ case ResReconstruct.strip_prefix ResClause.tfree_prefix x of
SOME tf => mk_tfree ctxt tf
| NONE => error ("fol_type_to_isa: " ^ x));
@@ -227,10 +225,10 @@
let val thy = ProofContext.theory_of ctxt
val _ = trace_msg (fn () => "fol_term_to_hol: " ^ Metis.Term.toString fol_tm)
fun tm_to_tt (Metis.Term.Var v) =
- (case Recon.strip_prefix ResClause.tvar_prefix v of
- SOME w => Type (Recon.make_tvar w)
+ (case ResReconstruct.strip_prefix ResClause.tvar_prefix v of
+ SOME w => Type (ResReconstruct.make_tvar w)
| NONE =>
- case Recon.strip_prefix ResClause.schematic_var_prefix v of
+ case ResReconstruct.strip_prefix ResClause.schematic_var_prefix v of
SOME w => Term (mk_var (w, HOLogic.typeT))
| NONE => Term (mk_var (v, HOLogic.typeT)) )
(*Var from Metis with a name like _nnn; possibly a type variable*)
@@ -247,14 +245,14 @@
and applic_to_tt ("=",ts) =
Term (list_comb(Const ("op =", HOLogic.typeT), terms_of (map tm_to_tt ts)))
| applic_to_tt (a,ts) =
- case Recon.strip_prefix ResClause.const_prefix a of
+ case ResReconstruct.strip_prefix ResClause.const_prefix a of
SOME b =>
- let val c = Recon.invert_const b
- val ntypes = Recon.num_typargs thy c
+ let val c = ResReconstruct.invert_const b
+ val ntypes = ResReconstruct.num_typargs thy c
val nterms = length ts - ntypes
val tts = map tm_to_tt ts
val tys = types_of (List.take(tts,ntypes))
- val ntyargs = Recon.num_typargs thy c
+ val ntyargs = ResReconstruct.num_typargs thy c
in if length tys = ntyargs then
apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes))
else error ("Constant " ^ c ^ " expects " ^ Int.toString ntyargs ^
@@ -265,13 +263,14 @@
cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
end
| NONE => (*Not a constant. Is it a type constructor?*)
- case Recon.strip_prefix ResClause.tconst_prefix a of
- SOME b => Type (Term.Type(Recon.invert_type_const b, types_of (map tm_to_tt ts)))
+ case ResReconstruct.strip_prefix ResClause.tconst_prefix a of
+ SOME b =>
+ Type (Term.Type (ResReconstruct.invert_type_const b, types_of (map tm_to_tt ts)))
| NONE => (*Maybe a TFree. Should then check that ts=[].*)
- case Recon.strip_prefix ResClause.tfree_prefix a of
+ case ResReconstruct.strip_prefix ResClause.tfree_prefix a of
SOME b => Type (mk_tfree ctxt b)
| NONE => (*a fixed variable? They are Skolem functions.*)
- case Recon.strip_prefix ResClause.fixed_var_prefix a of
+ case ResReconstruct.strip_prefix ResClause.fixed_var_prefix a of
SOME b =>
let val opr = Term.Free(b, HOLogic.typeT)
in apply_list opr (length ts) (map tm_to_tt ts) end
@@ -282,16 +281,16 @@
fun fol_term_to_hol_FT ctxt fol_tm =
let val _ = trace_msg (fn () => "fol_term_to_hol_FT: " ^ Metis.Term.toString fol_tm)
fun cvt (Metis.Term.Fn ("ti", [Metis.Term.Var v, _])) =
- (case Recon.strip_prefix ResClause.schematic_var_prefix v of
+ (case ResReconstruct.strip_prefix ResClause.schematic_var_prefix v of
SOME w => mk_var(w, dummyT)
| NONE => mk_var(v, dummyT))
| cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn ("=",[]), _])) =
Const ("op =", HOLogic.typeT)
| cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (x,[]), ty])) =
- (case Recon.strip_prefix ResClause.const_prefix x of
- SOME c => Const (Recon.invert_const c, dummyT)
+ (case ResReconstruct.strip_prefix ResClause.const_prefix x of
+ SOME c => Const (ResReconstruct.invert_const c, dummyT)
| NONE => (*Not a constant. Is it a fixed variable??*)
- case Recon.strip_prefix ResClause.fixed_var_prefix x of
+ case ResReconstruct.strip_prefix ResClause.fixed_var_prefix x of
SOME v => Free (v, fol_type_to_isa ctxt ty)
| NONE => error ("fol_term_to_hol_FT bad constant: " ^ x))
| cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (".",[tm1,tm2]), _])) =
@@ -302,13 +301,15 @@
| cvt (Metis.Term.Fn ("=", [tm1,tm2])) =
list_comb(Const ("op =", HOLogic.typeT), map cvt [tm1,tm2])
| cvt (t as Metis.Term.Fn (x, [])) =
- (case Recon.strip_prefix ResClause.const_prefix x of
- SOME c => Const (Recon.invert_const c, dummyT)
+ (case ResReconstruct.strip_prefix ResClause.const_prefix x of
+ SOME c => Const (ResReconstruct.invert_const c, dummyT)
| NONE => (*Not a constant. Is it a fixed variable??*)
- case Recon.strip_prefix ResClause.fixed_var_prefix x of
+ case ResReconstruct.strip_prefix ResClause.fixed_var_prefix x of
SOME v => Free (v, dummyT)
- | NONE => (trace_msg (fn () => "fol_term_to_hol_FT bad const: " ^ x); fol_term_to_hol_RAW ctxt t))
- | cvt t = (trace_msg (fn () => "fol_term_to_hol_FT bad term: " ^ Metis.Term.toString t); fol_term_to_hol_RAW ctxt t)
+ | NONE => (trace_msg (fn () => "fol_term_to_hol_FT bad const: " ^ x);
+ fol_term_to_hol_RAW ctxt t))
+ | cvt t = (trace_msg (fn () => "fol_term_to_hol_FT bad term: " ^ Metis.Term.toString t);
+ fol_term_to_hol_RAW ctxt t)
in cvt fol_tm end;
fun fol_term_to_hol ctxt FO = fol_term_to_hol_RAW ctxt
@@ -382,9 +383,9 @@
" in " ^ Display.string_of_thm ctxt i_th);
NONE)
fun remove_typeinst (a, t) =
- case Recon.strip_prefix ResClause.schematic_var_prefix a of
+ case ResReconstruct.strip_prefix ResClause.schematic_var_prefix a of
SOME b => SOME (b, t)
- | NONE => case Recon.strip_prefix ResClause.tvar_prefix a of
+ | NONE => case ResReconstruct.strip_prefix ResClause.tvar_prefix a of
SOME _ => NONE (*type instantiations are forbidden!*)
| NONE => SOME (a,t) (*internal Metis var?*)
val _ = trace_msg (fn () => " isa th: " ^ Display.string_of_thm ctxt i_th)
@@ -451,7 +452,7 @@
in cterm_instantiate [(refl_x, c_t)] REFL_THM end;
fun get_ty_arg_size _ (Const ("op =", _)) = 0 (*equality has no type arguments*)
- | get_ty_arg_size thy (Const (c, _)) = (Recon.num_typargs thy c handle TYPE _ => 0)
+ | get_ty_arg_size thy (Const (c, _)) = (ResReconstruct.num_typargs thy c handle TYPE _ => 0)
| get_ty_arg_size _ _ = 0;
(* INFERENCE RULE: EQUALITY *)
@@ -522,7 +523,7 @@
val subst' = incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
val _ = trace_msg (fn () => "subst' " ^ Display.string_of_thm ctxt subst')
val eq_terms = map (pairself (cterm_of thy))
- (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
+ (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
in cterm_instantiate eq_terms subst' end;
val factor = Seq.hd o distinct_subgoals_tac;
@@ -588,7 +589,7 @@
(* ------------------------------------------------------------------------- *)
type logic_map =
- {axioms : (Metis.Thm.thm * Thm.thm) list,
+ {axioms : (Metis.Thm.thm * thm) list,
tfrees : ResClause.type_literal list};
fun const_in_metis c (pred, tm_list) =
--- a/src/HOL/Tools/prop_logic.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/prop_logic.ML Wed Oct 28 11:43:06 2009 +0000
@@ -37,9 +37,9 @@
val eval : (int -> bool) -> prop_formula -> bool (* semantics *)
(* propositional representation of HOL terms *)
- val prop_formula_of_term : Term.term -> int Termtab.table -> prop_formula * int Termtab.table
+ val prop_formula_of_term : term -> int Termtab.table -> prop_formula * int Termtab.table
(* HOL term representation of propositional formulae *)
- val term_of_prop_formula : prop_formula -> Term.term
+ val term_of_prop_formula : prop_formula -> term
end;
structure PropLogic : PROP_LOGIC =
--- a/src/HOL/Tools/quickcheck_generators.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/quickcheck_generators.ML Wed Oct 28 11:43:06 2009 +0000
@@ -8,7 +8,7 @@
type seed = Random_Engine.seed
val random_fun: typ -> typ -> ('a -> 'a -> bool) -> ('a -> term)
-> (seed -> ('b * (unit -> term)) * seed) -> (seed -> seed * seed)
- -> seed -> (('a -> 'b) * (unit -> Term.term)) * seed
+ -> seed -> (('a -> 'b) * (unit -> term)) * seed
val ensure_random_typecopy: string -> theory -> theory
val random_aux_specification: string -> string -> term list -> local_theory -> local_theory
val mk_random_aux_eqs: theory -> Datatype.descr -> (string * sort) list
--- a/src/HOL/Tools/recfun_codegen.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/recfun_codegen.ML Wed Oct 28 11:43:06 2009 +0000
@@ -65,9 +65,9 @@
exception EQN of string * typ * string;
-fun cycle g (xs, x : string) =
+fun cycle g x xs =
if member (op =) xs x then xs
- else Library.foldl (cycle g) (x :: xs, flat (Graph.all_paths (fst g) (x, x)));
+ else fold (cycle g) (flat (Graph.all_paths (fst g) (x, x))) (x :: xs);
fun add_rec_funs thy defs dep module eqs gr =
let
@@ -107,7 +107,7 @@
val gr1 = add_edge (dname, dep)
(new_node (dname, (SOME (EQN (s, T, "")), module, "")) gr);
val (fundef, gr2) = mk_fundef module "" true eqs' gr1 ;
- val xs = cycle gr2 ([], dname);
+ val xs = cycle gr2 dname [];
val cs = map (fn x => case get_node gr2 x of
(SOME (EQN (s, T, _)), _, _) => (s, T)
| _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
--- a/src/HOL/Tools/refute.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/refute.ML Wed Oct 28 11:43:06 2009 +0000
@@ -25,16 +25,15 @@
exception MAXVARS_EXCEEDED
- val add_interpreter : string -> (theory -> model -> arguments -> Term.term ->
+ val add_interpreter : string -> (theory -> model -> arguments -> term ->
(interpretation * model * arguments) option) -> theory -> theory
- val add_printer : string -> (theory -> model -> Term.typ ->
- interpretation -> (int -> bool) -> Term.term option) -> theory -> theory
+ val add_printer : string -> (theory -> model -> typ ->
+ interpretation -> (int -> bool) -> term option) -> theory -> theory
- val interpret : theory -> model -> arguments -> Term.term ->
+ val interpret : theory -> model -> arguments -> term ->
(interpretation * model * arguments)
- val print : theory -> model -> Term.typ -> interpretation ->
- (int -> bool) -> Term.term
+ val print : theory -> model -> typ -> interpretation -> (int -> bool) -> term
val print_model : theory -> model -> (int -> bool) -> string
(* ------------------------------------------------------------------------- *)
@@ -46,10 +45,10 @@
val get_default_params : theory -> (string * string) list
val actual_params : theory -> (string * string) list -> params
- val find_model : theory -> params -> Term.term -> bool -> unit
+ val find_model : theory -> params -> term -> bool -> unit
(* tries to find a model for a formula: *)
- val satisfy_term : theory -> (string * string) list -> Term.term -> unit
+ val satisfy_term : theory -> (string * string) list -> term -> unit
(* tries to find a model that refutes a formula: *)
val refute_term : theory -> (string * string) list -> term -> unit
val refute_goal : theory -> (string * string) list -> thm -> int -> unit
@@ -60,20 +59,18 @@
(* Additional functions used by Nitpick (to be factored out) *)
(* ------------------------------------------------------------------------- *)
- val close_form : Term.term -> Term.term
- val get_classdef : theory -> string -> (string * Term.term) option
- val norm_rhs : Term.term -> Term.term
- val get_def : theory -> string * Term.typ -> (string * Term.term) option
- val get_typedef : theory -> Term.typ -> (string * Term.term) option
- val is_IDT_constructor : theory -> string * Term.typ -> bool
- val is_IDT_recursor : theory -> string * Term.typ -> bool
- val is_const_of_class: theory -> string * Term.typ -> bool
- val monomorphic_term : Type.tyenv -> Term.term -> Term.term
- val specialize_type : theory -> (string * Term.typ) -> Term.term -> Term.term
- val string_of_typ : Term.typ -> string
- val typ_of_dtyp :
- Datatype.descr -> (Datatype.dtyp * Term.typ) list -> Datatype.dtyp
- -> Term.typ
+ val close_form : term -> term
+ val get_classdef : theory -> string -> (string * term) option
+ val norm_rhs : term -> term
+ val get_def : theory -> string * typ -> (string * term) option
+ val get_typedef : theory -> typ -> (string * term) option
+ val is_IDT_constructor : theory -> string * typ -> bool
+ val is_IDT_recursor : theory -> string * typ -> bool
+ val is_const_of_class: theory -> string * typ -> bool
+ val monomorphic_term : Type.tyenv -> term -> term
+ val specialize_type : theory -> (string * typ) -> term -> term
+ val string_of_typ : typ -> string
+ val typ_of_dtyp : Datatype.descr -> (Datatype.dtyp * typ) list -> Datatype.dtyp -> typ
end; (* signature REFUTE *)
structure Refute : REFUTE =
@@ -185,7 +182,7 @@
(* ------------------------------------------------------------------------- *)
type model =
- (Term.typ * int) list * (Term.term * interpretation) list;
+ (typ * int) list * (term * interpretation) list;
(* ------------------------------------------------------------------------- *)
(* arguments: additional arguments required during interpretation of terms *)
@@ -207,10 +204,10 @@
structure RefuteData = TheoryDataFun
(
type T =
- {interpreters: (string * (theory -> model -> arguments -> Term.term ->
+ {interpreters: (string * (theory -> model -> arguments -> term ->
(interpretation * model * arguments) option)) list,
- printers: (string * (theory -> model -> Term.typ -> interpretation ->
- (int -> bool) -> Term.term option)) list,
+ printers: (string * (theory -> model -> typ -> interpretation ->
+ (int -> bool) -> term option)) list,
parameters: string Symtab.table};
val empty = {interpreters = [], printers = [], parameters = Symtab.empty};
val copy = I;
@@ -433,9 +430,8 @@
(* (Term.indexname * Term.typ) list *)
val vars = sort_wrt (fst o fst) (map dest_Var (OldTerm.term_vars t))
in
- Library.foldl (fn (t', ((x, i), T)) =>
- (Term.all T) $ Abs (x, T, abstract_over (Var ((x, i), T), t')))
- (t, vars)
+ fold (fn ((x, i), T) => fn t' =>
+ Term.all T $ Abs (x, T, abstract_over (Var ((x, i), T), t'))) vars t
end;
(* ------------------------------------------------------------------------- *)
@@ -768,60 +764,55 @@
val _ = tracing "Adding axioms..."
val axioms = Theory.all_axioms_of thy
fun collect_this_axiom (axname, ax) axs =
- let
- val ax' = unfold_defs thy ax
- in
- if member (op aconv) axs ax' then axs
- else (tracing axname; collect_term_axioms (ax' :: axs, ax'))
- end
- (* Term.term list * Term.typ -> Term.term list *)
- and collect_sort_axioms (axs, T) =
- let
- (* string list *)
- val sort = (case T of
- TFree (_, sort) => sort
- | TVar (_, sort) => sort
- | _ => raise REFUTE ("collect_axioms", "type " ^
- Syntax.string_of_typ_global thy T ^ " is not a variable"))
- (* obtain axioms for all superclasses *)
- val superclasses = sort @ (maps (Sign.super_classes thy) sort)
- (* merely an optimization, because 'collect_this_axiom' disallows *)
- (* duplicate axioms anyway: *)
- val superclasses = distinct (op =) superclasses
- val class_axioms = maps (fn class => map (fn ax =>
- ("<" ^ class ^ ">", Thm.prop_of ax))
- (#axioms (AxClass.get_info thy class) handle ERROR _ => []))
- superclasses
- (* replace the (at most one) schematic type variable in each axiom *)
- (* by the actual type 'T' *)
- val monomorphic_class_axioms = map (fn (axname, ax) =>
- (case Term.add_tvars ax [] of
- [] =>
- (axname, ax)
- | [(idx, S)] =>
- (axname, monomorphic_term (Vartab.make [(idx, (S, T))]) ax)
- | _ =>
- raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^
- Syntax.string_of_term_global thy ax ^
- ") contains more than one type variable")))
- class_axioms
- in
- fold collect_this_axiom monomorphic_class_axioms axs
- end
- (* Term.term list * Term.typ -> Term.term list *)
- and collect_type_axioms (axs, T) =
+ let
+ val ax' = unfold_defs thy ax
+ in
+ if member (op aconv) axs ax' then axs
+ else (tracing axname; collect_term_axioms ax' (ax' :: axs))
+ end
+ and collect_sort_axioms T axs =
+ let
+ val sort =
+ (case T of
+ TFree (_, sort) => sort
+ | TVar (_, sort) => sort
+ | _ => raise REFUTE ("collect_axioms",
+ "type " ^ Syntax.string_of_typ_global thy T ^ " is not a variable"))
+ (* obtain axioms for all superclasses *)
+ val superclasses = sort @ maps (Sign.super_classes thy) sort
+ (* merely an optimization, because 'collect_this_axiom' disallows *)
+ (* duplicate axioms anyway: *)
+ val superclasses = distinct (op =) superclasses
+ val class_axioms = maps (fn class => map (fn ax =>
+ ("<" ^ class ^ ">", Thm.prop_of ax))
+ (#axioms (AxClass.get_info thy class) handle ERROR _ => []))
+ superclasses
+ (* replace the (at most one) schematic type variable in each axiom *)
+ (* by the actual type 'T' *)
+ val monomorphic_class_axioms = map (fn (axname, ax) =>
+ (case Term.add_tvars ax [] of
+ [] => (axname, ax)
+ | [(idx, S)] => (axname, monomorphic_term (Vartab.make [(idx, (S, T))]) ax)
+ | _ =>
+ raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^
+ Syntax.string_of_term_global thy ax ^
+ ") contains more than one type variable")))
+ class_axioms
+ in
+ fold collect_this_axiom monomorphic_class_axioms axs
+ end
+ and collect_type_axioms T axs =
case T of
(* simple types *)
- Type ("prop", []) => axs
- | Type ("fun", [T1, T2]) => collect_type_axioms
- (collect_type_axioms (axs, T1), T2)
+ Type ("prop", []) => axs
+ | Type ("fun", [T1, T2]) => collect_type_axioms T2 (collect_type_axioms T1 axs)
(* axiomatic type classes *)
- | Type ("itself", [T1]) => collect_type_axioms (axs, T1)
- | Type (s, Ts) =>
+ | Type ("itself", [T1]) => collect_type_axioms T1 axs
+ | Type (s, Ts) =>
(case Datatype.get_info thy s of
SOME info => (* inductive datatype *)
(* only collect relevant type axioms for the argument types *)
- Library.foldl collect_type_axioms (axs, Ts)
+ fold collect_type_axioms Ts axs
| NONE =>
(case get_typedef thy T of
SOME (axname, ax) =>
@@ -829,20 +820,19 @@
| NONE =>
(* unspecified type, perhaps introduced with "typedecl" *)
(* at least collect relevant type axioms for the argument types *)
- Library.foldl collect_type_axioms (axs, Ts)))
- (* axiomatic type classes *)
- | TFree _ => collect_sort_axioms (axs, T)
+ fold collect_type_axioms Ts axs))
(* axiomatic type classes *)
- | TVar _ => collect_sort_axioms (axs, T)
- (* Term.term list * Term.term -> Term.term list *)
- and collect_term_axioms (axs, t) =
+ | TFree _ => collect_sort_axioms T axs
+ (* axiomatic type classes *)
+ | TVar _ => collect_sort_axioms T axs
+ and collect_term_axioms t axs =
case t of
(* Pure *)
Const (@{const_name all}, _) => axs
| Const (@{const_name "=="}, _) => axs
| Const (@{const_name "==>"}, _) => axs
(* axiomatic type classes *)
- | Const (@{const_name TYPE}, T) => collect_type_axioms (axs, T)
+ | Const (@{const_name TYPE}, T) => collect_type_axioms T axs
(* HOL *)
| Const (@{const_name Trueprop}, _) => axs
| Const (@{const_name Not}, _) => axs
@@ -850,7 +840,7 @@
| Const (@{const_name True}, _) => axs
(* redundant, since 'False' is also an IDT constructor *)
| Const (@{const_name False}, _) => axs
- | Const (@{const_name undefined}, T) => collect_type_axioms (axs, T)
+ | Const (@{const_name undefined}, T) => collect_type_axioms T axs
| Const (@{const_name The}, T) =>
let
val ax = specialize_type thy (@{const_name The}, T)
@@ -865,74 +855,68 @@
in
collect_this_axiom ("Hilbert_Choice.someI", ax) axs
end
- | Const (@{const_name All}, T) => collect_type_axioms (axs, T)
- | Const (@{const_name Ex}, T) => collect_type_axioms (axs, T)
- | Const (@{const_name "op ="}, T) => collect_type_axioms (axs, T)
+ | Const (@{const_name All}, T) => collect_type_axioms T axs
+ | Const (@{const_name Ex}, T) => collect_type_axioms T axs
+ | Const (@{const_name "op ="}, T) => collect_type_axioms T axs
| Const (@{const_name "op &"}, _) => axs
| Const (@{const_name "op |"}, _) => axs
| Const (@{const_name "op -->"}, _) => axs
(* sets *)
- | Const (@{const_name Collect}, T) => collect_type_axioms (axs, T)
- | Const (@{const_name "op :"}, T) => collect_type_axioms (axs, T)
+ | Const (@{const_name Collect}, T) => collect_type_axioms T axs
+ | Const (@{const_name "op :"}, T) => collect_type_axioms T axs
(* other optimizations *)
- | Const (@{const_name Finite_Set.card}, T) => collect_type_axioms (axs, T)
+ | Const (@{const_name Finite_Set.card}, T) => collect_type_axioms T axs
| Const (@{const_name Finite_Set.finite}, T) =>
- collect_type_axioms (axs, T)
+ collect_type_axioms T axs
| Const (@{const_name HOL.less}, T as Type ("fun", [Type ("nat", []),
Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
- collect_type_axioms (axs, T)
+ collect_type_axioms T axs
| Const (@{const_name HOL.plus}, T as Type ("fun", [Type ("nat", []),
Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
- collect_type_axioms (axs, T)
+ collect_type_axioms T axs
| Const (@{const_name HOL.minus}, T as Type ("fun", [Type ("nat", []),
Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
- collect_type_axioms (axs, T)
+ collect_type_axioms T axs
| Const (@{const_name HOL.times}, T as Type ("fun", [Type ("nat", []),
Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
- collect_type_axioms (axs, T)
- | Const (@{const_name List.append}, T) => collect_type_axioms (axs, T)
- | Const (@{const_name lfp}, T) => collect_type_axioms (axs, T)
- | Const (@{const_name gfp}, T) => collect_type_axioms (axs, T)
- | Const (@{const_name fst}, T) => collect_type_axioms (axs, T)
- | Const (@{const_name snd}, T) => collect_type_axioms (axs, T)
+ collect_type_axioms T axs
+ | Const (@{const_name List.append}, T) => collect_type_axioms T axs
+ | Const (@{const_name lfp}, T) => collect_type_axioms T axs
+ | Const (@{const_name gfp}, T) => collect_type_axioms T axs
+ | Const (@{const_name fst}, T) => collect_type_axioms T axs
+ | Const (@{const_name snd}, T) => collect_type_axioms T axs
(* simply-typed lambda calculus *)
| Const (s, T) =>
if is_const_of_class thy (s, T) then
(* axiomatic type classes: add "OFCLASS(?'a::c, c_class)" *)
(* and the class definition *)
let
- val class = Logic.class_of_const s
+ val class = Logic.class_of_const s
val of_class = Logic.mk_of_class (TVar (("'a", 0), [class]), class)
- val ax_in = SOME (specialize_type thy (s, T) of_class)
+ val ax_in = SOME (specialize_type thy (s, T) of_class)
(* type match may fail due to sort constraints *)
handle Type.TYPE_MATCH => NONE
- val ax_1 = Option.map (fn ax => (Syntax.string_of_term_global thy ax, ax))
- ax_in
- val ax_2 = Option.map (apsnd (specialize_type thy (s, T)))
- (get_classdef thy class)
+ val ax_1 = Option.map (fn ax => (Syntax.string_of_term_global thy ax, ax)) ax_in
+ val ax_2 = Option.map (apsnd (specialize_type thy (s, T))) (get_classdef thy class)
in
- collect_type_axioms (fold collect_this_axiom
- (map_filter I [ax_1, ax_2]) axs, T)
+ collect_type_axioms T (fold collect_this_axiom (map_filter I [ax_1, ax_2]) axs)
end
else if is_IDT_constructor thy (s, T)
orelse is_IDT_recursor thy (s, T) then
(* only collect relevant type axioms *)
- collect_type_axioms (axs, T)
+ collect_type_axioms T axs
else
(* other constants should have been unfolded, with some *)
(* exceptions: e.g. Abs_xxx/Rep_xxx functions for *)
(* typedefs, or type-class related constants *)
(* only collect relevant type axioms *)
- collect_type_axioms (axs, T)
- | Free (_, T) => collect_type_axioms (axs, T)
- | Var (_, T) => collect_type_axioms (axs, T)
- | Bound i => axs
- | Abs (_, T, body) => collect_term_axioms
- (collect_type_axioms (axs, T), body)
- | t1 $ t2 => collect_term_axioms
- (collect_term_axioms (axs, t1), t2)
- (* Term.term list *)
- val result = map close_form (collect_term_axioms ([], t))
+ collect_type_axioms T axs
+ | Free (_, T) => collect_type_axioms T axs
+ | Var (_, T) => collect_type_axioms T axs
+ | Bound _ => axs
+ | Abs (_, T, body) => collect_term_axioms body (collect_type_axioms T axs)
+ | t1 $ t2 => collect_term_axioms t2 (collect_term_axioms t1 axs)
+ val result = map close_form (collect_term_axioms t [])
val _ = tracing " ...done."
in
result
@@ -1040,7 +1024,7 @@
fun size_of_typ T =
case AList.lookup (op =) sizes (string_of_typ T) of
SOME n => n
- | NONE => minsize
+ | NONE => minsize
in
map (fn T => (T, size_of_typ T)) xs
end;
@@ -1094,13 +1078,13 @@
case next (maxsize-minsize) 0 0 diffs of
SOME diffs' =>
(* merge with those types for which the size is fixed *)
- SOME (snd (Library.foldl_map (fn (ds, (T, _)) =>
+ SOME (fst (fold_map (fn (T, _) => fn ds =>
case AList.lookup (op =) sizes (string_of_typ T) of
(* return the fixed size *)
- SOME n => (ds, (T, n))
+ SOME n => ((T, n), ds)
(* consume the head of 'ds', add 'minsize' *)
- | NONE => (tl ds, (T, minsize + hd ds)))
- (diffs', xs)))
+ | NONE => ((T, minsize + hd ds), tl ds))
+ xs diffs'))
| NONE =>
NONE
end;
@@ -1157,8 +1141,7 @@
val _ = tracing ("Unfolded term: " ^ Syntax.string_of_term_global thy u)
val axioms = collect_axioms thy u
(* Term.typ list *)
- val types = Library.foldl (fn (acc, t') =>
- uncurry (union (op =)) (acc, (ground_types thy t'))) ([], u :: axioms)
+ val types = fold (union (op =) o ground_types thy) (u :: axioms) []
val _ = tracing ("Ground types: "
^ (if null types then "none."
else commas (map (Syntax.string_of_typ_global thy) types)))
@@ -1197,15 +1180,15 @@
val _ = tracing ("Translating term (sizes: "
^ commas (map (fn (_, n) => string_of_int n) universe) ^ ") ...")
(* translate 'u' and all axioms *)
- val ((model, args), intrs) = Library.foldl_map (fn ((m, a), t') =>
+ val (intrs, (model, args)) = fold_map (fn t' => fn (m, a) =>
let
val (i, m', a') = interpret thy m a t'
in
(* set 'def_eq' to 'true' *)
- ((m', {maxvars = #maxvars a', def_eq = true,
+ (i, (m', {maxvars = #maxvars a', def_eq = true,
next_idx = #next_idx a', bounds = #bounds a',
- wellformed = #wellformed a'}), i)
- end) ((init_model, init_args), u :: axioms)
+ wellformed = #wellformed a'}))
+ end) (u :: axioms) (init_model, init_args)
(* make 'u' either true or false, and make all axioms true, and *)
(* add the well-formedness side condition *)
val fm_u = (if negate then toFalse else toTrue) (hd intrs)
@@ -1309,10 +1292,9 @@
(* (Term.indexname * Term.typ) list *)
val vars = sort_wrt (fst o fst) (map dest_Var (OldTerm.term_vars t))
(* Term.term *)
- val ex_closure = Library.foldl (fn (t', ((x, i), T)) =>
- (HOLogic.exists_const T) $
- Abs (x, T, abstract_over (Var ((x, i), T), t')))
- (t, vars)
+ val ex_closure = fold (fn ((x, i), T) => fn t' =>
+ HOLogic.exists_const T $
+ Abs (x, T, abstract_over (Var ((x, i), T), t'))) vars t
(* Note: If 't' is of type 'propT' (rather than 'boolT'), applying *)
(* 'HOLogic.exists_const' is not type-correct. However, this is not *)
(* really a problem as long as 'find_model' still interprets the *)
@@ -1733,8 +1715,8 @@
(* create all constants of type 'T' *)
val constants = make_constants thy model T
(* interpret the 'body' separately for each constant *)
- val ((model', args'), bodies) = Library.foldl_map
- (fn ((m, a), c) =>
+ val (bodies, (model', args')) = fold_map
+ (fn c => fn (m, a) =>
let
(* add 'c' to 'bounds' *)
val (i', m', a') = interpret thy m {maxvars = #maxvars a,
@@ -1743,15 +1725,15 @@
in
(* keep the new model m' and 'next_idx' and 'wellformed', *)
(* but use old 'bounds' *)
- ((m', {maxvars = maxvars, def_eq = def_eq,
+ (i', (m', {maxvars = maxvars, def_eq = def_eq,
next_idx = #next_idx a', bounds = bounds,
- wellformed = #wellformed a'}), i')
+ wellformed = #wellformed a'}))
end)
- ((model, args), constants)
+ constants (model, args)
in
SOME (Node bodies, model', args')
end
- | t1 $ t2 =>
+ | t1 $ t2 =>
let
(* interpret 't1' and 't2' separately *)
val (intr1, model1, args1) = interpret thy model args t1
@@ -2212,14 +2194,14 @@
Node (replicate size (make_undef ds))
end
(* returns the interpretation for a constructor *)
- fun make_constr (offset, []) =
- if offset<total then
- (offset+1, Leaf ((replicate offset False) @ True ::
- (replicate (total-offset-1) False)))
+ fun make_constr [] offset =
+ if offset < total then
+ (Leaf (replicate offset False @ True ::
+ (replicate (total - offset - 1) False)), offset + 1)
else
raise REFUTE ("IDT_constructor_interpreter",
"offset >= total")
- | make_constr (offset, d::ds) =
+ | make_constr (d::ds) offset =
let
(* Term.typ *)
val dT = typ_of_dtyp descr typ_assoc d
@@ -2228,8 +2210,8 @@
(* for the IDT) *)
val terms' = canonical_terms typs' dT
(* sanity check *)
- val _ = if length terms' <>
- size_of_type thy (typs', []) dT
+ val _ =
+ if length terms' <> size_of_type thy (typs', []) dT
then
raise REFUTE ("IDT_constructor_interpreter",
"length of terms' is not equal to old size")
@@ -2239,46 +2221,52 @@
(* for the IDT) *)
val terms = canonical_terms typs dT
(* sanity check *)
- val _ = if length terms <> size_of_type thy (typs, []) dT
+ val _ =
+ if length terms <> size_of_type thy (typs, []) dT
then
raise REFUTE ("IDT_constructor_interpreter",
"length of terms is not equal to current size")
else ()
(* sanity check *)
- val _ = if length terms < length terms' then
+ val _ =
+ if length terms < length terms' then
raise REFUTE ("IDT_constructor_interpreter",
"current size is less than old size")
else ()
(* sanity check: every element of terms' must also be *)
(* present in terms *)
- val _ = if List.all (member op= terms) terms' then ()
+ val _ =
+ if List.all (member (op =) terms) terms' then ()
else
raise REFUTE ("IDT_constructor_interpreter",
"element has disappeared")
(* sanity check: the order on elements of terms' is *)
(* the same in terms, for those elements *)
- val _ = let
+ val _ =
+ let
fun search (x::xs) (y::ys) =
- if x = y then search xs ys else search (x::xs) ys
+ if x = y then search xs ys else search (x::xs) ys
| search (x::xs) [] =
- raise REFUTE ("IDT_constructor_interpreter",
- "element order not preserved")
+ raise REFUTE ("IDT_constructor_interpreter",
+ "element order not preserved")
| search [] _ = ()
in search terms' terms end
(* int * interpretation list *)
- val (new_offset, intrs) = Library.foldl_map (fn (off, t_elem) =>
- (* if 't_elem' existed at the previous depth, *)
- (* proceed recursively, otherwise map the entire *)
- (* subtree to "undefined" *)
- if t_elem mem terms' then
- make_constr (off, ds)
- else
- (off, make_undef ds)) (offset, terms)
+ val (intrs, new_offset) =
+ fold_map (fn t_elem => fn off =>
+ (* if 't_elem' existed at the previous depth, *)
+ (* proceed recursively, otherwise map the entire *)
+ (* subtree to "undefined" *)
+ if t_elem mem terms' then
+ make_constr ds off
+ else
+ (make_undef ds, off))
+ terms offset
in
- (new_offset, Node intrs)
+ (Node intrs, new_offset)
end
in
- SOME (snd (make_constr (offset, ctypes)), model, args)
+ SOME (fst (make_constr ctypes offset), model, args)
end
end
| NONE => (* body type is not an inductive datatype *)
@@ -2332,12 +2320,12 @@
else (* mconstrs_count = length params *)
let
(* interpret each parameter separately *)
- val ((model', args'), p_intrs) = Library.foldl_map (fn ((m, a), p) =>
+ val (p_intrs, (model', args')) = fold_map (fn p => fn (m, a) =>
let
val (i, m', a') = interpret thy m a p
in
- ((m', a'), i)
- end) ((model, args), params)
+ (i, (m', a'))
+ end) params (model, args)
val (typs, _) = model'
(* 'index' is /not/ necessarily the index of the IDT that *)
(* the recursion operator is associated with, but merely *)
@@ -2440,14 +2428,14 @@
end) descr
(* associate constructors with corresponding parameters *)
(* (int * (interpretation * interpretation) list) list *)
- val (p_intrs', mc_p_intrs) = Library.foldl_map
- (fn (p_intrs', (idx, c_intrs)) =>
+ val (mc_p_intrs, p_intrs') = fold_map
+ (fn (idx, c_intrs) => fn p_intrs' =>
let
val len = length c_intrs
in
- (List.drop (p_intrs', len),
- (idx, c_intrs ~~ List.take (p_intrs', len)))
- end) (p_intrs, mc_intrs)
+ ((idx, c_intrs ~~ List.take (p_intrs', len)),
+ List.drop (p_intrs', len))
+ end) mc_intrs p_intrs
(* sanity check: no 'p_intr' may be left afterwards *)
val _ = if p_intrs' <> [] then
raise REFUTE ("IDT_recursion_interpreter",
@@ -2672,15 +2660,15 @@
let
(* interpretation -> int *)
fun number_of_elements (Node xs) =
- Library.foldl (fn (n, x) =>
- if x=TT then
- n+1
- else if x=FF then
- n
- else
- raise REFUTE ("Finite_Set_card_interpreter",
- "interpretation for set type does not yield a Boolean"))
- (0, xs)
+ fold (fn x => fn n =>
+ if x = TT then
+ n + 1
+ else if x = FF then
+ n
+ else
+ raise REFUTE ("Finite_Set_card_interpreter",
+ "interpretation for set type does not yield a Boolean"))
+ xs 0
| number_of_elements (Leaf _) =
raise REFUTE ("Finite_Set_card_interpreter",
"interpretation for set type is a leaf")
@@ -2885,7 +2873,7 @@
(* of width 'size_elem' and depth 'length_list' (with 'size_list' *)
(* nodes total) *)
(* (int * (int * int)) list *)
- val (_, lenoff_lists) = Library.foldl_map (fn ((offsets, off), elem) =>
+ val (lenoff_lists, _) = fold_map (fn elem => fn (offsets, off) =>
(* corresponds to a pre-order traversal of the tree *)
let
val len = length offsets
@@ -2894,10 +2882,10 @@
in
if len < list_length then
(* go to first child node *)
- ((off :: offsets, off * size_elem), assoc)
+ (assoc, (off :: offsets, off * size_elem))
else if off mod size_elem < size_elem - 1 then
(* go to next sibling node *)
- ((offsets, off + 1), assoc)
+ (assoc, (offsets, off + 1))
else
(* go back up the stack until we find a level where we can go *)
(* to the next sibling node *)
@@ -2909,12 +2897,12 @@
[] =>
(* we're at the last node in the tree; the next value *)
(* won't be used anyway *)
- (([], 0), assoc)
+ (assoc, ([], 0))
| off'::offs' =>
(* go to next sibling node *)
- ((offs', off' + 1), assoc)
+ (assoc, (offs', off' + 1))
end
- end) (([], 0), elements)
+ end) elements ([], 0)
(* we also need the reverse association (from length/offset to *)
(* index) *)
val lenoff'_lists = map Library.swap lenoff_lists
@@ -3251,7 +3239,7 @@
print thy (typs, []) dT (List.nth (consts, n)) assignment
end) (cargs ~~ args)
in
- SOME (Library.foldl op$ (cTerm, argsTerms))
+ SOME (list_comb (cTerm, argsTerms))
end
end
| NONE => (* not an inductive datatype *)
--- a/src/HOL/Tools/res_axioms.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/res_axioms.ML Wed Oct 28 11:43:06 2009 +0000
@@ -71,7 +71,7 @@
prefix for the Skolem constant.*)
fun declare_skofuns s th =
let
- val nref = Unsynchronized.ref 0
+ val nref = Unsynchronized.ref 0 (* FIXME ??? *)
fun dec_sko (Const ("Ex",_) $ (xtp as Abs (_, T, p))) (axs, thy) =
(*Existential: declare a Skolem function, then insert into body and continue*)
let
@@ -87,7 +87,8 @@
val (c, thy') =
Sign.declare_const ((Binding.conceal (Binding.name cname), cT), NoSyn) thy
val cdef = cname ^ "_def"
- val thy'' = Theory.add_defs_i true false [(Binding.name cdef, Logic.mk_equals (c, rhs))] thy'
+ val thy'' =
+ Theory.add_defs_i true false [(Binding.name cdef, Logic.mk_equals (c, rhs))] thy'
val ax = Thm.axiom thy'' (Sign.full_bname thy'' cdef)
in dec_sko (subst_bound (list_comb (c, args), p)) (ax :: axs, thy'') end
| dec_sko (Const ("All", _) $ (Abs (a, T, p))) thx =
@@ -102,7 +103,7 @@
(*Traverse a theorem, accumulating Skolem function definitions.*)
fun assume_skofuns s th =
- let val sko_count = Unsynchronized.ref 0
+ let val sko_count = Unsynchronized.ref 0 (* FIXME ??? *)
fun dec_sko (Const ("Ex",_) $ (xtp as Abs(_,T,p))) defs =
(*Existential: declare a Skolem function, then insert into body and continue*)
let val skos = map (#1 o Logic.dest_equals) defs (*existing sko fns*)
--- a/src/HOL/Tools/res_reconstruct.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/res_reconstruct.ML Wed Oct 28 11:43:06 2009 +0000
@@ -10,17 +10,17 @@
val fix_sorts: sort Vartab.table -> term -> term
val invert_const: string -> string
val invert_type_const: string -> string
- val num_typargs: Context.theory -> string -> int
+ val num_typargs: theory -> string -> int
val make_tvar: string -> typ
val strip_prefix: string -> string -> string option
- val setup: Context.theory -> Context.theory
+ val setup: theory -> theory
(* extracting lemma list*)
val find_failure: string -> string option
val lemma_list: bool -> string ->
- string * string vector * (int * int) * Proof.context * Thm.thm * int -> string * string list
+ string * string vector * (int * int) * Proof.context * thm * int -> string * string list
(* structured proofs *)
val structured_proof: string ->
- string * string vector * (int * int) * Proof.context * Thm.thm * int -> string * string list
+ string * string vector * (int * int) * Proof.context * thm * int -> string * string list
end;
structure ResReconstruct : RES_RECONSTRUCT =
--- a/src/HOL/Tools/sat_funcs.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/sat_funcs.ML Wed Oct 28 11:43:06 2009 +0000
@@ -92,8 +92,8 @@
(* ------------------------------------------------------------------------- *)
datatype CLAUSE = NO_CLAUSE
- | ORIG_CLAUSE of Thm.thm
- | RAW_CLAUSE of Thm.thm * (int * Thm.cterm) list;
+ | ORIG_CLAUSE of thm
+ | RAW_CLAUSE of thm * (int * cterm) list;
(* ------------------------------------------------------------------------- *)
(* resolve_raw_clauses: given a non-empty list of raw clauses, we fold *)
@@ -346,11 +346,16 @@
fold Thm.weaken (rev sorted_clauses) False_thm
end
in
- case (tracing ("Invoking solver " ^ (!solver)); SatSolver.invoke_solver (!solver) fm) of
+ case
+ let val the_solver = ! solver
+ in (tracing ("Invoking solver " ^ the_solver); SatSolver.invoke_solver the_solver fm) end
+ of
SatSolver.UNSATISFIABLE (SOME (clause_table, empty_id)) => (
if !trace_sat then
tracing ("Proof trace from SAT solver:\n" ^
- "clauses: " ^ ML_Syntax.print_list (ML_Syntax.print_pair Int.toString (ML_Syntax.print_list Int.toString)) (Inttab.dest clause_table) ^ "\n" ^
+ "clauses: " ^ ML_Syntax.print_list
+ (ML_Syntax.print_pair Int.toString (ML_Syntax.print_list Int.toString))
+ (Inttab.dest clause_table) ^ "\n" ^
"empty clause: " ^ Int.toString empty_id)
else ();
if !quick_and_dirty then
@@ -388,8 +393,9 @@
val msg = "SAT solver found a countermodel:\n"
^ (commas
o map (fn (term, idx) =>
- Syntax.string_of_term_global @{theory} term ^ ": "
- ^ (case assignment idx of NONE => "arbitrary" | SOME true => "true" | SOME false => "false")))
+ Syntax.string_of_term_global @{theory} term ^ ": " ^
+ (case assignment idx of NONE => "arbitrary"
+ | SOME true => "true" | SOME false => "false")))
(Termtab.dest atom_table)
in
raise THM (msg, 0, [])
--- a/src/HOL/Tools/sat_solver.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/Tools/sat_solver.ML Wed Oct 28 11:43:06 2009 +0000
@@ -370,13 +370,13 @@
(* string * solver -> unit *)
- fun add_solver (name, new_solver) =
+ fun add_solver (name, new_solver) = CRITICAL (fn () =>
let
val the_solvers = !solvers;
val _ = if AList.defined (op =) the_solvers name
then warning ("SAT solver " ^ quote name ^ " was defined before")
else ();
- in solvers := AList.update (op =) (name, new_solver) the_solvers end;
+ in solvers := AList.update (op =) (name, new_solver) the_solvers end);
(* ------------------------------------------------------------------------- *)
(* invoke_solver: returns the solver associated with the given 'name' *)
--- a/src/HOL/ex/Predicate_Compile.thy Wed Oct 28 11:42:31 2009 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,17 +0,0 @@
-theory Predicate_Compile
-imports Complex_Main RPred
-uses
- "../Tools/Predicate_Compile/pred_compile_aux.ML"
- "../Tools/Predicate_Compile/predicate_compile_core.ML"
- "../Tools/Predicate_Compile/pred_compile_set.ML"
- "../Tools/Predicate_Compile/pred_compile_data.ML"
- "../Tools/Predicate_Compile/pred_compile_fun.ML"
- "../Tools/Predicate_Compile/pred_compile_pred.ML"
- "../Tools/Predicate_Compile/predicate_compile.ML"
- "../Tools/Predicate_Compile/pred_compile_quickcheck.ML"
-begin
-
-setup {* Predicate_Compile.setup *}
-setup {* Quickcheck.add_generator ("pred_compile", Predicate_Compile_Quickcheck.quickcheck) *}
-
-end
\ No newline at end of file
--- a/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy Wed Oct 28 11:43:06 2009 +0000
@@ -1,5 +1,5 @@
theory Predicate_Compile_Alternative_Defs
-imports Predicate_Compile
+imports Main
begin
section {* Set operations *}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Predicate_Compile_Quickcheck.thy Wed Oct 28 11:43:06 2009 +0000
@@ -0,0 +1,12 @@
+(* Author: Lukas Bulwahn, TU Muenchen *)
+
+header {* A Prototype of Quickcheck based on the Predicate Compiler *}
+
+theory Predicate_Compile_Quickcheck
+imports Main
+uses "../Tools/Predicate_Compile/predicate_compile_quickcheck.ML"
+begin
+
+setup {* Quickcheck.add_generator ("predicate_compile", Predicate_Compile_Quickcheck.quickcheck) *}
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy Wed Oct 28 11:43:06 2009 +0000
@@ -0,0 +1,118 @@
+theory Predicate_Compile_Quickcheck_ex
+imports Predicate_Compile_Quickcheck
+ Predicate_Compile_Alternative_Defs
+begin
+
+section {* Sets *}
+
+
+section {* Context Free Grammar *}
+
+datatype alphabet = a | b
+
+inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
+ "[] \<in> S\<^isub>1"
+| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
+
+theorem S\<^isub>1_sound:
+"w \<in> S\<^isub>1p \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+(*quickcheck[generator=predicate_compile, size=15]*)
+oops
+
+
+inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
+ "[] \<in> S\<^isub>2"
+| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
+| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
+| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
+| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
+| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
+
+code_pred [inductify, rpred] S\<^isub>2 .
+thm S\<^isub>2.rpred_equation
+thm A\<^isub>2.rpred_equation
+thm B\<^isub>2.rpred_equation
+
+values [random] 10 "{x. S\<^isub>2 x}"
+
+lemma "w \<in> S\<^isub>2 ==> w \<noteq> [] ==> w \<noteq> [b, a] ==> w \<in> {}"
+quickcheck[generator=predicate_compile]
+oops
+
+lemma "[x <- w. x = a] = []"
+quickcheck[generator=predicate_compile]
+oops
+
+
+lemma "length ([x \<leftarrow> w. x = a]) = (0::nat)"
+(*quickcheck[generator=predicate_compile]*)
+oops
+
+
+
+lemma
+"w \<in> S\<^isub>2 ==> length [x \<leftarrow> w. x = a] < Suc (Suc 0)"
+(*quickcheck[generator=predicate_compile]*)
+oops
+
+
+theorem S\<^isub>2_sound:
+"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+(*quickcheck[generator=predicate_compile, size=15, iterations=100]*)
+oops
+
+inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
+ "[] \<in> S\<^isub>3"
+| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
+| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
+| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
+
+code_pred [inductify] S\<^isub>3 .
+thm S\<^isub>3.equation
+
+values 10 "{x. S\<^isub>3 x}"
+
+lemma S\<^isub>3_sound:
+"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+(*quickcheck[generator=predicate_compile, size=10, iterations=1]*)
+oops
+
+
+lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
+(*quickcheck[size=10, generator = pred_compile]*)
+oops
+
+theorem S\<^isub>3_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3"
+(*quickcheck[generator=SML]*)
+(*quickcheck[generator=predicate_compile, size=10, iterations=100]*)
+oops
+
+
+inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
+ "[] \<in> S\<^isub>4"
+| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
+| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
+| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
+| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
+
+theorem S\<^isub>4_sound:
+"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+(*quickcheck[generator = predicate_compile, size=2, iterations=1]*)
+oops
+
+theorem S\<^isub>4_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
+(*quickcheck[generator = pred_compile, size=5, iterations=1]*)
+oops
+
+
+end
\ No newline at end of file
--- a/src/HOL/ex/Predicate_Compile_ex.thy Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOL/ex/Predicate_Compile_ex.thy Wed Oct 28 11:43:06 2009 +0000
@@ -2,6 +2,99 @@
imports Main Predicate_Compile_Alternative_Defs
begin
+subsection {* Basic predicates *}
+
+inductive False' :: "bool"
+
+code_pred (mode: []) False' .
+code_pred [depth_limited] False' .
+code_pred [rpred] False' .
+
+inductive EmptySet :: "'a \<Rightarrow> bool"
+
+code_pred (mode: [], [1]) EmptySet .
+
+definition EmptySet' :: "'a \<Rightarrow> bool"
+where "EmptySet' = {}"
+
+code_pred (mode: [], [1]) [inductify, show_intermediate_results] EmptySet' .
+
+inductive EmptyRel :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
+
+code_pred (mode: [], [1], [2], [1, 2]) EmptyRel .
+
+inductive EmptyClosure :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+
+code_pred (mode: [], [1], [2], [1, 2])EmptyClosure .
+thm EmptyClosure.equation
+(* TODO: inductive package is broken!
+inductive False'' :: "bool"
+where
+ "False \<Longrightarrow> False''"
+
+code_pred (mode: []) False'' .
+
+inductive EmptySet'' :: "'a \<Rightarrow> bool"
+where
+ "False \<Longrightarrow> EmptySet'' x"
+
+code_pred (mode: [1]) EmptySet'' .
+code_pred (mode: [], [1]) [inductify] EmptySet'' .
+*)
+inductive True' :: "bool"
+where
+ "True \<Longrightarrow> True'"
+
+code_pred (mode: []) True' .
+
+consts a' :: 'a
+
+inductive Fact :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+where
+"Fact a' a'"
+
+code_pred (mode: [], [1], [2], [1, 2]) Fact .
+
+inductive zerozero :: "nat * nat => bool"
+where
+ "zerozero (0, 0)"
+
+code_pred zerozero .
+code_pred [rpred, show_compilation] zerozero .
+
+subsection {* Preprocessor Inlining *}
+
+definition "equals == (op =)"
+
+inductive zerozero' where
+ "equals (x, y) (0, 0) ==> zerozero' (x, y)"
+
+code_pred (mode: [1]) zerozero' .
+
+lemma zerozero'_eq: "zerozero' == zerozero"
+proof -
+ have "zerozero' = zerozero"
+ apply (auto simp add: mem_def)
+ apply (cases rule: zerozero'.cases)
+ apply (auto simp add: equals_def intro: zerozero.intros)
+ apply (cases rule: zerozero.cases)
+ apply (auto simp add: equals_def intro: zerozero'.intros)
+ done
+ from this show "zerozero' == zerozero" by auto
+qed
+
+declare zerozero'_eq [code_pred_inline]
+
+definition "zerozero'' x == zerozero' x"
+
+text {* if preprocessing fails, zerozero'' will not have all modes. *}
+ML {* Predicate_Compile_Inline_Defs.get @{context} *}
+(* TODO: *)
+code_pred (mode: [1]) [inductify, show_intermediate_results] zerozero'' .
+
+subsection {* even predicate *}
+
inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
"even 0"
| "even n \<Longrightarrow> odd (Suc n)"
@@ -26,7 +119,6 @@
values [depth_limit = 7] "{x. even 6}"
values [depth_limit = 2] "{x. odd 7}"
values [depth_limit = 8] "{x. odd 7}"
-
values [depth_limit = 7] 10 "{n. even n}"
definition odd' where "odd' x == \<not> even x"
@@ -39,6 +131,14 @@
values [depth_limit = 2] "{x. odd' 7}"
values [depth_limit = 9] "{x. odd' 7}"
+inductive is_even :: "nat \<Rightarrow> bool"
+where
+ "n mod 2 = 0 \<Longrightarrow> is_even n"
+
+code_pred is_even .
+
+subsection {* append predicate *}
+
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"append [] xs xs"
| "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
@@ -60,7 +160,7 @@
value [code] "Predicate.the (append_1_2 [0::int, 1, 2] [3, 4, 5])"
value [code] "Predicate.the (append_3 ([]::int list))"
-subsection {* Tricky case with alternative rules *}
+text {* tricky case with alternative rules *}
inductive append2
where
@@ -78,15 +178,6 @@
from append2.cases[OF append2(1)] append2(2-3) show thesis by blast
qed
-subsection {* Tricky cases with tuples *}
-
-inductive zerozero :: "nat * nat => bool"
-where
- "zerozero (0, 0)"
-
-code_pred zerozero .
-code_pred [rpred] zerozero .
-
inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
where
"tupled_append ([], xs, xs)"
@@ -127,6 +218,8 @@
code_pred [inductify] tupled_append''' .
thm tupled_append'''.equation
+subsection {* map_ofP predicate *}
+
inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
where
"map_ofP ((a, b)#xs) a b"
@@ -135,6 +228,8 @@
code_pred (mode: [1], [1, 2], [1, 2, 3], [1, 3]) map_ofP .
thm map_ofP.equation
+subsection {* filter predicate *}
+
inductive filter1
for P
where
@@ -168,7 +263,7 @@
code_pred filter3 .
code_pred [depth_limited] filter3 .
thm filter3.depth_limited_equation
-(*code_pred [rpred] filter3 .*)
+
inductive filter4
where
"List.filter P xs = ys ==> filter4 P xs ys"
@@ -177,7 +272,7 @@
code_pred [depth_limited] filter4 .
code_pred [rpred] filter4 .
-section {* reverse *}
+subsection {* reverse predicate *}
inductive rev where
"rev [] []"
@@ -196,6 +291,8 @@
code_pred tupled_rev .
thm tupled_rev.equation
+subsection {* partition predicate *}
+
inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
for f where
"partition f [] [] []"
@@ -206,6 +303,11 @@
code_pred [depth_limited] partition .
code_pred [rpred] partition .
+values 10 "{(ys, zs). partition is_even
+ [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}"
+values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
+values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
+
inductive tupled_partition :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
for f where
"tupled_partition f ([], [], [])"
@@ -216,23 +318,13 @@
thm tupled_partition.equation
-
-inductive is_even :: "nat \<Rightarrow> bool"
-where
- "n mod 2 = 0 \<Longrightarrow> is_even n"
-
-code_pred is_even .
-
-values 10 "{(ys, zs). partition is_even
- [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}"
-values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
-values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
-
lemma [code_pred_intros]:
"r a b \<Longrightarrow> tranclp r a b"
"r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
by auto
+subsection {* transitive predicate *}
+
code_pred tranclp
proof -
case tranclp
@@ -265,7 +357,7 @@
values 20 "{(n, m). tranclp succ n m}"
*)
-subsection{* IMP *}
+subsection {* IMP *}
types
var = nat
@@ -304,7 +396,7 @@
code_pred tupled_exec .
-subsection{* CCS *}
+subsection {* CCS *}
text{* This example formalizes finite CCS processes without communication or
recursion. For simplicity, labels are natural numbers. *}
@@ -354,18 +446,15 @@
value [code] "Predicate.the (divmod_rel_1_2 1705 42)"
-section {* Executing definitions *}
+subsection {* Minimum *}
definition Min
where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)"
code_pred [inductify] Min .
-subsection {* Examples with lists *}
+subsection {* Lexicographic order *}
-subsubsection {* Lexicographic order *}
-
-thm lexord_def
code_pred [inductify] lexord .
code_pred [inductify, rpred] lexord .
thm lexord.equation
@@ -392,13 +481,7 @@
values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
values [depth_limit = 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
-(*values [random] "{xys. test_lexord xys}"*)
-(*values [depth_limit = 5 random] "{xy. lexord less_than_nat xy}"*)
-(*
-lemma "(u, v) : lexord less_than_nat ==> (x @ u, y @ v) : lexord less_than_nat"
-quickcheck[generator=pred_compile]
-oops
-*)
+
lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv
code_pred [inductify] lexn .
@@ -414,12 +497,10 @@
code_pred [inductify, rpred] lenlex .
thm lenlex.rpred_equation
-thm lists.intros
code_pred [inductify] lists .
-
thm lists.equation
-section {* AVL Tree Example *}
+subsection {* AVL Tree *}
datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
fun height :: "'a tree => nat" where
@@ -437,7 +518,6 @@
code_pred [rpred] avl .
thm avl.rpred_equation
-(*values [random] 10 "{t. avl (t::int tree)}"*)
fun set_of
where
@@ -455,10 +535,9 @@
code_pred [inductify] is_ord .
thm is_ord.equation
-code_pred [rpred] is_ord .
-thm is_ord.rpred_equation
+
-section {* Definitions about Relations *}
+subsection {* Definitions about Relations *}
code_pred [inductify] converse .
thm converse.equation
@@ -467,7 +546,7 @@
code_pred [inductify] Image .
thm Image.equation
(*TODO: *)
-ML {* Toplevel.debug := true *}
+
declare Id_on_def[unfolded UNION_def, code_pred_def]
code_pred [inductify] Id_on .
@@ -475,7 +554,7 @@
code_pred [inductify] Domain .
thm Domain.equation
code_pred [inductify] Range .
-thm sym_def
+thm Range.equation
code_pred [inductify] Field .
declare Sigma_def[unfolded UNION_def, code_pred_def]
declare refl_on_def[unfolded UNION_def, code_pred_def]
@@ -483,10 +562,6 @@
thm refl_on.equation
code_pred [inductify] total_on .
thm total_on.equation
-(*
-code_pred [inductify] sym .
-thm sym.equation
-*)
code_pred [inductify] antisym .
thm antisym.equation
code_pred [inductify] trans .
@@ -496,12 +571,13 @@
code_pred [inductify] inv_image .
thm inv_image.equation
-section {* List functions *}
+subsection {* Inverting list functions *}
code_pred [inductify] length .
+code_pred [inductify, rpred] length .
thm size_listP.equation
-code_pred [inductify, rpred] length .
thm size_listP.rpred_equation
+
values [random] 20 "{xs. size_listP (xs::nat list) (5::nat)}"
code_pred [inductify] concat .
@@ -509,7 +585,6 @@
code_pred [inductify] tl .
code_pred [inductify] last .
code_pred [inductify] butlast .
-(*code_pred [inductify] listsum .*)
code_pred [inductify] take .
code_pred [inductify] drop .
code_pred [inductify] zip .
@@ -526,15 +601,8 @@
code_pred [inductify] foldl .
code_pred [inductify] filter .
code_pred [inductify, rpred] filter .
-thm filterP.rpred_equation
-definition test where "test xs = filter (\<lambda>x. x = (1::nat)) xs"
-code_pred [inductify] test .
-thm testP.equation
-code_pred [inductify, rpred] test .
-thm testP.rpred_equation
-
-section {* Context Free Grammar *}
+subsection {* Context Free Grammar *}
datatype alphabet = a | b
@@ -553,79 +621,6 @@
values [random] 5 "{x. S\<^isub>1p x}"
-inductive is_a where
- "is_a a"
-
-inductive is_b where
- "is_b b"
-
-code_pred is_a .
-code_pred [depth_limited] is_a .
-code_pred [rpred] is_a .
-
-values [random] "{x. is_a x}"
-code_pred [depth_limited] is_b .
-code_pred [rpred] is_b .
-
-code_pred [inductify, depth_limited] filter .
-
-values [depth_limit=5] "{x. filterP is_a [a, b] x}"
-values [depth_limit=3] "{x. filterP is_b [a, b] x}"
-
-lemma "w \<in> S\<^isub>1 \<Longrightarrow> length (filter (\<lambda>x. x = a) w) = 1"
-(*quickcheck[generator=pred_compile, size=10]*)
-oops
-
-inductive test_lemma where
- "S\<^isub>1p w ==> filterP is_a w r1 ==> size_listP r1 r2 ==> filterP is_b w r3 ==> size_listP r3 r4 ==> r2 \<noteq> r4 ==> test_lemma w"
-(*
-code_pred [rpred] test_lemma .
-*)
-(*
-definition test_lemma' where
- "test_lemma' w == (w \<in> S\<^isub>1 \<and> (\<not> length [x <- w. x = a] = length [x <- w. x = b]))"
-
-code_pred [inductify, rpred] test_lemma' .
-thm test_lemma'.rpred_equation
-*)
-(*thm test_lemma'.rpred_equation*)
-(*
-values [depth_limit=3 random] "{x. S\<^isub>1 x}"
-*)
-code_pred [depth_limited] is_b .
-(*
-code_pred [inductify, depth_limited] filter .
-*)
-thm filterP.intros
-thm filterP.equation
-(*
-values [depth_limit=3] "{x. filterP is_b [a, b] x}"
-find_theorems "test_lemma'_hoaux"
-code_pred [depth_limited] test_lemma'_hoaux .
-thm test_lemma'_hoaux.depth_limited_equation
-values [depth_limit=2] "{x. test_lemma'_hoaux b}"
-inductive test1 where
- "\<not> test_lemma'_hoaux x ==> test1 x"
-code_pred test1 .
-code_pred [depth_limited] test1 .
-thm test1.depth_limited_equation
-thm test_lemma'_hoaux.depth_limited_equation
-thm test1.intros
-
-values [depth_limit=2] "{x. test1 b}"
-
-thm filterP.intros
-thm filterP.depth_limited_equation
-values [depth_limit=3] "{x. filterP test_lemma'_hoaux [a, b] x}"
-values [depth_limit=4 random] "{w. test_lemma w}"
-values [depth_limit=4 random] "{w. test_lemma' w}"
-*)
-(*
-theorem S\<^isub>1_sound:
-"w \<in> S\<^isub>1p \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=pred_compile, size=15]
-oops
-*)
inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
"[] \<in> S\<^isub>2"
| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
@@ -641,12 +636,6 @@
values [random] 10 "{x. S\<^isub>2 x}"
-theorem S\<^isub>2_sound:
-"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-(*quickcheck[generator=SML]*)
-(*quickcheck[generator=pred_compile, size=15, iterations=1]*)
-oops
-
inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
"[] \<in> S\<^isub>3"
| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
@@ -659,31 +648,6 @@
thm S\<^isub>3.equation
values 10 "{x. S\<^isub>3 x}"
-(*
-theorem S\<^isub>3_sound:
-"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=pred_compile, size=10, iterations=1]
-oops
-*)
-lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
-(*quickcheck[size=10, generator = pred_compile]*)
-oops
-(*
-inductive test
-where
- "length [x \<leftarrow> w. a = x] = length [x \<leftarrow> w. x = b] ==> test w"
-ML {* @{term "[x \<leftarrow> w. x = a]"} *}
-code_pred (inductify_all) test .
-
-thm test.equation
-*)
-(*
-theorem S\<^isub>3_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3"
-(*quickcheck[generator=SML]*)
-quickcheck[generator=pred_compile, size=10, iterations=100]
-oops
-*)
inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
"[] \<in> S\<^isub>4"
@@ -693,25 +657,8 @@
| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
-(*
-theorem S\<^isub>4_sound:
-"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator = pred_compile, size=2, iterations=1]
-oops
-*)
-theorem S\<^isub>4_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
-(*quickcheck[generator = pred_compile, size=5, iterations=1]*)
-oops
-theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete:
-"w \<in> S\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-"w \<in> A\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1"
-"w \<in> B\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1"
-(*quickcheck[generator = pred_compile, size=5, iterations=1]*)
-oops
-
-section {* Lambda *}
+subsection {* Lambda *}
datatype type =
Atom nat
@@ -728,7 +675,6 @@
"[]\<langle>i\<rangle> = None"
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
-
inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
where
"nth_el' (x # xs) 0 x"
@@ -738,8 +684,7 @@
where
Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
| Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
-(* | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U" *)
- | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
+ | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
primrec
lift :: "[dB, nat] => dB"
@@ -763,15 +708,4 @@
| appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
| abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
-lemma "Gamma \<turnstile> t : T \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> Gamma \<turnstile> t' : T"
-(*quickcheck[generator = pred_compile, size = 10, iterations = 1]*)
-oops
-
-lemma filter_eq_ConsD:
- "filter P ys = x#xs \<Longrightarrow>
- \<exists>us vs. ys = ts @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
-(*quickcheck[generator = pred_compile]*)
-oops
-
-
end
\ No newline at end of file
--- a/src/HOLCF/IOA/ABP/Check.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOLCF/IOA/ABP/Check.ML Wed Oct 28 11:43:06 2009 +0000
@@ -15,8 +15,8 @@
fun check(extacts,intacts,string_of_a,startsI,string_of_s,
nexts,hom,transA,startsS) =
let fun check_s(s,unchecked,checked) =
- let fun check_sa(unchecked,a) =
- let fun check_sas(unchecked,t) =
+ let fun check_sa a unchecked =
+ let fun check_sas t unchecked =
(if a mem extacts then
(if transA(hom s,a,hom t) then ( )
else (writeln("Error: Mapping of Externals!");
@@ -29,8 +29,8 @@
string_of_a a; writeln"";
string_of_s t;writeln"";writeln"" ));
if t mem checked then unchecked else insert (op =) t unchecked)
- in Library.foldl check_sas (unchecked,nexts s a) end;
- val unchecked' = Library.foldl check_sa (unchecked,extacts @ intacts)
+ in fold check_sas (nexts s a) unchecked end;
+ val unchecked' = fold check_sa (extacts @ intacts) unchecked
in (if s mem startsI then
(if hom(s) mem startsS then ()
else writeln("Error: At start states!"))
--- a/src/HOLCF/Tools/Domain/domain_axioms.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOLCF/Tools/Domain/domain_axioms.ML Wed Oct 28 11:43:06 2009 +0000
@@ -17,7 +17,7 @@
end;
-structure Domain_Axioms :> DOMAIN_AXIOMS =
+structure Domain_Axioms : DOMAIN_AXIOMS =
struct
open Domain_Library;
@@ -218,13 +218,13 @@
("bisim_def",
%%:(comp_dname^"_bisim")==mk_lam("R", foldr1 mk_conj (mapn one_comp 0 eqs)));
- fun add_one (thy,(dnam,axs,dfs)) =
- thy |> Sign.add_path dnam
- |> add_defs_infer dfs
- |> add_axioms_infer axs
- |> Sign.parent_path;
+ fun add_one (dnam, axs, dfs) =
+ Sign.add_path dnam
+ #> add_defs_infer dfs
+ #> add_axioms_infer axs
+ #> Sign.parent_path;
- val thy = Library.foldl add_one (thy', mapn (calc_axioms comp_dname eqs) 0 eqs);
+ val thy = fold add_one (mapn (calc_axioms comp_dname eqs) 0 eqs) thy';
in thy |> Sign.add_path comp_dnam
|> add_defs_infer (bisim_def::(if length eqs>1 then [copy_def] else []))
--- a/src/HOLCF/Tools/cont_consts.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOLCF/Tools/cont_consts.ML Wed Oct 28 11:43:06 2009 +0000
@@ -1,5 +1,4 @@
(* Title: HOLCF/Tools/cont_consts.ML
- ID: $Id$
Author: Tobias Mayr, David von Oheimb, and Markus Wenzel
HOLCF version of consts: handle continuous function types in mixfix
@@ -12,7 +11,7 @@
val add_consts_i: (binding * typ * mixfix) list -> theory -> theory
end;
-structure ContConsts :> CONT_CONSTS =
+structure ContConsts: CONT_CONSTS =
struct
@@ -29,18 +28,23 @@
| change_arrow n (Type(_,[S,T])) = Type ("fun",[S,change_arrow (n-1) T])
| change_arrow _ _ = sys_error "cont_consts: change_arrow";
-fun trans_rules name2 name1 n mx = let
- fun argnames _ 0 = []
- | argnames c n = chr c::argnames (c+1) (n-1);
- val vnames = argnames (ord "A") n;
- val extra_parse_rule = Syntax.ParseRule (Constant name2, Constant name1);
- in [Syntax.ParsePrintRule (Syntax.mk_appl (Constant name2) (map Variable vnames),
- Library.foldl (fn (t,arg) => (Syntax.mk_appl (Constant "Rep_CFun")
- [t,Variable arg]))
- (Constant name1,vnames))]
- @(case mx of InfixName _ => [extra_parse_rule]
- | InfixlName _ => [extra_parse_rule]
- | InfixrName _ => [extra_parse_rule] | _ => []) end;
+fun trans_rules name2 name1 n mx =
+ let
+ fun argnames _ 0 = []
+ | argnames c n = chr c::argnames (c+1) (n-1);
+ val vnames = argnames (ord "A") n;
+ val extra_parse_rule = Syntax.ParseRule (Constant name2, Constant name1);
+ in
+ [Syntax.ParsePrintRule
+ (Syntax.mk_appl (Constant name2) (map Variable vnames),
+ fold (fn arg => fn t => Syntax.mk_appl (Constant "Rep_CFun") [t, Variable arg])
+ vnames (Constant name1))] @
+ (case mx of
+ InfixName _ => [extra_parse_rule]
+ | InfixlName _ => [extra_parse_rule]
+ | InfixrName _ => [extra_parse_rule]
+ | _ => [])
+ end;
(* transforming infix/mixfix declarations of constants with type ...->...
@@ -59,7 +63,8 @@
(const_binding mx syn, T, InfixrName (Binding.name_of syn, p))
| fix_mixfix decl = decl;
-fun transform decl = let
+fun transform decl =
+ let
val (c, T, mx) = fix_mixfix decl;
val c1 = Binding.name_of c;
val c2 = "_cont_" ^ c1;
--- a/src/HOLCF/Tools/pcpodef.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/HOLCF/Tools/pcpodef.ML Wed Oct 28 11:43:06 2009 +0000
@@ -90,6 +90,7 @@
fun make_cpo admissible (type_def, below_def, set_def) theory =
let
+ (* FIXME fold_rule might fold user input inintentionally *)
val admissible' = fold_rule (the_list set_def) admissible;
val cpo_thms = map (Thm.transfer theory) [type_def, below_def, admissible'];
val theory' = theory
@@ -112,6 +113,7 @@
fun make_pcpo UU_mem (type_def, below_def, set_def) theory =
let
+ (* FIXME fold_rule might fold user input inintentionally *)
val UU_mem' = fold_rule (the_list set_def) UU_mem;
val pcpo_thms = map (Thm.transfer theory) [type_def, below_def, UU_mem'];
val theory' = theory
--- a/src/Provers/Arith/fast_lin_arith.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Provers/Arith/fast_lin_arith.ML Wed Oct 28 11:43:06 2009 +0000
@@ -702,12 +702,11 @@
result
end;
-fun add_datoms (dats : (bool * term) list, ((lhs,_,_,rhs,_,d) : LA_Data.decomp, _)) :
- (bool * term) list =
+fun add_datoms ((lhs,_,_,rhs,_,d) : LA_Data.decomp, _) (dats : (bool * term) list) =
union_bterm (map (pair d o fst) lhs) (union_bterm (map (pair d o fst) rhs) dats);
fun discr (initems : (LA_Data.decomp * int) list) : bool list =
- map fst (Library.foldl add_datoms ([],initems));
+ map fst (fold add_datoms initems []);
fun refutes ctxt params show_ex :
(typ list * (LA_Data.decomp * int) list) list -> injust list -> injust list option =
--- a/src/Provers/order.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Provers/order.ML Wed Oct 28 11:43:06 2009 +0000
@@ -697,9 +697,9 @@
dfs_visit along with u form a connected component. We
collect all the connected components together in a
list, which is what is returned. *)
- Library.foldl (fn (comps,u) =>
+ fold (fn u => fn comps =>
if been_visited u then comps
- else ((u :: dfs_visit (transpose (op aconv) G) u) :: comps)) ([],(!finish))
+ else (u :: dfs_visit (transpose (op aconv) G) u) :: comps) (!finish) []
end;
--- a/src/Provers/splitter.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Provers/splitter.ML Wed Oct 28 11:43:06 2009 +0000
@@ -26,9 +26,10 @@
signature SPLITTER =
sig
(* somewhat more internal functions *)
- val cmap_of_split_thms : thm list -> (string * (typ * term * thm * typ * int) list) list
- val split_posns : (string * (typ * term * thm * typ * int) list) list -> theory -> typ list -> term ->
- (thm * (typ * typ * int list) list * int list * typ * term) list (* first argument is a "cmap", returns a list of "split packs" *)
+ val cmap_of_split_thms: thm list -> (string * (typ * term * thm * typ * int) list) list
+ val split_posns: (string * (typ * term * thm * typ * int) list) list ->
+ theory -> typ list -> term -> (thm * (typ * typ * int list) list * int list * typ * term) list
+ (* first argument is a "cmap", returns a list of "split packs" *)
(* the "real" interface, providing a number of tactics *)
val split_tac : thm list -> int -> tactic
val split_inside_tac: thm list -> int -> tactic
@@ -57,37 +58,33 @@
fun split_format_err () = error "Wrong format for split rule";
-(* thm -> (string * typ) * bool *)
fun split_thm_info thm = case concl_of (Data.mk_eq thm) of
Const("==", _) $ (Var _ $ t) $ c => (case strip_comb t of
(Const p, _) => (p, case c of (Const (s, _) $ _) => s = const_not | _ => false)
| _ => split_format_err ())
| _ => split_format_err ();
-(* thm list -> (string * (typ * term * thm * typ * int) list) list *)
fun cmap_of_split_thms thms =
let
val splits = map Data.mk_eq thms
- fun add_thm (cmap, thm) =
- (case concl_of thm of _$(t as _$lhs)$_ =>
- (case strip_comb lhs of (Const(a,aT),args) =>
- let val info = (aT,lhs,thm,fastype_of t,length args)
- in case AList.lookup (op =) cmap a of
- SOME infos => AList.update (op =) (a, info::infos) cmap
- | NONE => (a,[info])::cmap
- end
- | _ => split_format_err())
- | _ => split_format_err())
+ fun add_thm thm cmap =
+ (case concl_of thm of _ $ (t as _ $ lhs) $ _ =>
+ (case strip_comb lhs of (Const(a,aT),args) =>
+ let val info = (aT,lhs,thm,fastype_of t,length args)
+ in case AList.lookup (op =) cmap a of
+ SOME infos => AList.update (op =) (a, info::infos) cmap
+ | NONE => (a,[info])::cmap
+ end
+ | _ => split_format_err())
+ | _ => split_format_err())
in
- Library.foldl add_thm ([], splits)
+ fold add_thm splits []
end;
(* ------------------------------------------------------------------------- *)
(* mk_case_split_tac *)
(* ------------------------------------------------------------------------- *)
-(* (int * int -> order) -> thm list -> int -> tactic * <split_posns> *)
-
fun mk_case_split_tac order =
let
@@ -154,7 +151,7 @@
(* add all loose bound variables in t to list is *)
-fun add_lbnos (is,t) = add_loose_bnos (t,0,is);
+fun add_lbnos t is = add_loose_bnos (t, 0, is);
(* check if the innermost abstraction that needs to be removed
has a body of type T; otherwise the expansion thm will fail later on
@@ -194,7 +191,7 @@
fun mk_split_pack (thm, T: typ, T', n, ts, apsns, pos, TB, t) =
if n > length ts then []
else let val lev = length apsns
- val lbnos = Library.foldl add_lbnos ([],Library.take(n,ts))
+ val lbnos = fold add_lbnos (Library.take (n, ts)) []
val flbnos = List.filter (fn i => i < lev) lbnos
val tt = incr_boundvars (~lev) t
in if null flbnos then
@@ -225,13 +222,11 @@
local
exception MATCH
in
- (* Context.theory -> Type.tyenv * (Term.typ * Term.typ) -> Type.tyenv *)
fun typ_match sg (tyenv, TU) = (Sign.typ_match sg TU tyenv)
handle Type.TYPE_MATCH => raise MATCH
- (* Context.theory -> Term.typ list * Term.term * Term.term -> bool *)
+
fun fomatch sg args =
let
- (* Type.tyenv -> Term.typ list * Term.term * Term.term -> Type.tyenv *)
fun mtch tyinsts = fn
(Ts, Var(_,T), t) =>
typ_match sg (tyinsts, (T, fastype_of1(Ts,t)))
@@ -247,10 +242,8 @@
mtch (mtch tyinsts (Ts,f,g)) (Ts, t, u)
| _ => raise MATCH
in (mtch Vartab.empty args; true) handle MATCH => false end;
-end (* local *)
+end;
-(* (string * (Term.typ * Term.term * Thm.thm * Term.typ * int) list) list -> Context.theory -> Term.typ list -> Term.term ->
- (Thm.thm * (Term.typ * Term.typ * int list) list * int list * Term.typ * Term.term) list *)
fun split_posns (cmap : (string * (typ * term * thm * typ * int) list) list) sg Ts t =
let
val T' = fastype_of1 (Ts, t);
@@ -260,20 +253,21 @@
| posns Ts pos apsns t =
let
val (h, ts) = strip_comb t
- fun iter((i, a), t) = (i+1, (posns Ts (i::pos) apsns t) @ a);
- val a = case h of
- Const(c, cT) =>
- let fun find [] = []
- | find ((gcT, pat, thm, T, n)::tups) =
- let val t2 = list_comb (h, Library.take (n, ts))
- in if Sign.typ_instance sg (cT, gcT)
- andalso fomatch sg (Ts,pat,t2)
- then mk_split_pack(thm,T,T',n,ts,apsns,pos,type_of1(Ts,t2),t2)
- else find tups
- end
- in find (these (AList.lookup (op =) cmap c)) end
- | _ => []
- in snd(Library.foldl iter ((0, a), ts)) end
+ fun iter t (i, a) = (i+1, (posns Ts (i::pos) apsns t) @ a);
+ val a =
+ case h of
+ Const(c, cT) =>
+ let fun find [] = []
+ | find ((gcT, pat, thm, T, n)::tups) =
+ let val t2 = list_comb (h, Library.take (n, ts))
+ in if Sign.typ_instance sg (cT, gcT)
+ andalso fomatch sg (Ts,pat,t2)
+ then mk_split_pack(thm,T,T',n,ts,apsns,pos,type_of1(Ts,t2),t2)
+ else find tups
+ end
+ in find (these (AList.lookup (op =) cmap c)) end
+ | _ => []
+ in snd (fold iter ts (0, a)) end
in posns Ts [] [] t end;
fun nth_subgoal i thm = List.nth (prems_of thm, i-1);
@@ -343,8 +337,8 @@
(Logic.strip_assums_concl (Thm.prop_of thm'))));
val cert = cterm_of (Thm.theory_of_thm state);
val cntxt = mk_cntxt_splitthm t tt TB;
- val abss = Library.foldl (fn (t, T) => Abs ("", T, t));
- in cterm_instantiate [(cert P, cert (abss (cntxt, Ts)))] thm'
+ val abss = fold (fn T => fn t => Abs ("", T, t));
+ in cterm_instantiate [(cert P, cert (abss Ts cntxt))] thm'
end;
@@ -355,8 +349,6 @@
i : number of subgoal the tactic should be applied to
*****************************************************************************)
-(* thm list -> int -> tactic *)
-
fun split_tac [] i = no_tac
| split_tac splits i =
let val cmap = cmap_of_split_thms splits
@@ -379,9 +371,9 @@
in (split_tac, exported_split_posns) end; (* mk_case_split_tac *)
-val (split_tac, split_posns) = mk_case_split_tac int_ord;
+val (split_tac, split_posns) = mk_case_split_tac int_ord;
-val (split_inside_tac, _) = mk_case_split_tac (rev_order o int_ord);
+val (split_inside_tac, _) = mk_case_split_tac (rev_order o int_ord);
(*****************************************************************************
@@ -389,7 +381,7 @@
splits : list of split-theorems to be tried
****************************************************************************)
-fun split_asm_tac [] = K no_tac
+fun split_asm_tac [] = K no_tac
| split_asm_tac splits =
let val cname_list = map (fst o fst o split_thm_info) splits;
@@ -431,25 +423,28 @@
(* addsplits / delsplits *)
-fun string_of_typ (Type (s, Ts)) = (if null Ts then ""
- else enclose "(" ")" (commas (map string_of_typ Ts))) ^ s
+fun string_of_typ (Type (s, Ts)) =
+ (if null Ts then "" else enclose "(" ")" (commas (map string_of_typ Ts))) ^ s
| string_of_typ _ = "_";
fun split_name (name, T) asm = "split " ^
(if asm then "asm " else "") ^ name ^ " :: " ^ string_of_typ T;
fun ss addsplits splits =
- let fun addsplit (ss,split) =
- let val (name,asm) = split_thm_info split
- in Simplifier.addloop (ss, (split_name name asm,
- (if asm then split_asm_tac else split_tac) [split])) end
- in Library.foldl addsplit (ss,splits) end;
+ let
+ fun addsplit split ss =
+ let
+ val (name, asm) = split_thm_info split
+ val tac = (if asm then split_asm_tac else split_tac) [split]
+ in Simplifier.addloop (ss, (split_name name asm, tac)) end
+ in fold addsplit splits ss end;
fun ss delsplits splits =
- let fun delsplit(ss,split) =
- let val (name,asm) = split_thm_info split
- in Simplifier.delloop (ss, split_name name asm)
- end in Library.foldl delsplit (ss,splits) end;
+ let
+ fun delsplit split ss =
+ let val (name, asm) = split_thm_info split
+ in Simplifier.delloop (ss, split_name name asm) end
+ in fold delsplit splits ss end;
(* attributes *)
@@ -471,7 +466,8 @@
(* theory setup *)
val setup =
- Attrib.setup @{binding split} (Attrib.add_del split_add split_del) "declare case split rule" #>
+ Attrib.setup @{binding split}
+ (Attrib.add_del split_add split_del) "declare case split rule" #>
Method.setup @{binding split}
(Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (CHANGED_PROP o gen_split_tac ths))))
"apply case split rule";
--- a/src/Pure/General/file.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/General/file.ML Wed Oct 28 11:43:06 2009 +0000
@@ -90,10 +90,10 @@
Unsynchronized.ref (Symtab.empty: {time_stamp: string, id: string} Symtab.table);
in
-fun check_cache (path, time) = CRITICAL (fn () =>
+fun check_cache (path, time) =
(case Symtab.lookup (! ident_cache) path of
NONE => NONE
- | SOME {time_stamp, id} => if time_stamp = time then SOME id else NONE));
+ | SOME {time_stamp, id} => if time_stamp = time then SOME id else NONE);
fun update_cache (path, (time, id)) = CRITICAL (fn () =>
Unsynchronized.change ident_cache
--- a/src/Pure/General/print_mode.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/General/print_mode.ML Wed Oct 28 11:43:06 2009 +0000
@@ -35,7 +35,7 @@
let val modes =
(case Thread.getLocal tag of
SOME (SOME modes) => modes
- | _ => NAMED_CRITICAL "print_mode" (fn () => ! print_mode))
+ | _ => ! print_mode)
in subtract (op =) [input, internal] modes end;
fun print_mode_active mode = member (op =) (print_mode_value ()) mode;
--- a/src/Pure/Isar/isar_cmd.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/Isar/isar_cmd.ML Wed Oct 28 11:43:06 2009 +0000
@@ -333,14 +333,14 @@
Toplevel.keep (ProofContext.print_abbrevs o Toplevel.context_of);
val print_facts = Toplevel.unknown_context o Toplevel.keep (fn state =>
- ProofContext.setmp_verbose
+ ProofContext.setmp_verbose_CRITICAL
ProofContext.print_lthms (Toplevel.context_of state));
val print_configs = Toplevel.unknown_context o Toplevel.keep (fn state =>
Attrib.print_configs (Toplevel.context_of state));
val print_theorems_proof = Toplevel.keep (fn state =>
- ProofContext.setmp_verbose
+ ProofContext.setmp_verbose_CRITICAL
ProofContext.print_lthms (Proof.context_of (Toplevel.proof_of state)));
val print_theorems_theory = Toplevel.keep (fn state =>
@@ -431,10 +431,10 @@
(* print proof context contents *)
val print_binds = Toplevel.unknown_context o Toplevel.keep (fn state =>
- ProofContext.setmp_verbose ProofContext.print_binds (Toplevel.context_of state));
+ ProofContext.setmp_verbose_CRITICAL ProofContext.print_binds (Toplevel.context_of state));
val print_cases = Toplevel.unknown_context o Toplevel.keep (fn state =>
- ProofContext.setmp_verbose ProofContext.print_cases (Toplevel.context_of state));
+ ProofContext.setmp_verbose_CRITICAL ProofContext.print_cases (Toplevel.context_of state));
(* print theorems, terms, types etc. *)
--- a/src/Pure/Isar/isar_document.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/Isar/isar_document.ML Wed Oct 28 11:43:06 2009 +0000
@@ -119,13 +119,13 @@
in
fun change_states f = NAMED_CRITICAL "Isar" (fn () => Unsynchronized.change global_states f);
-fun get_states () = NAMED_CRITICAL "Isar" (fn () => ! global_states);
+fun get_states () = ! global_states;
fun change_commands f = NAMED_CRITICAL "Isar" (fn () => Unsynchronized.change global_commands f);
-fun get_commands () = NAMED_CRITICAL "Isar" (fn () => ! global_commands);
+fun get_commands () = ! global_commands;
fun change_documents f = NAMED_CRITICAL "Isar" (fn () => Unsynchronized.change global_documents f);
-fun get_documents () = NAMED_CRITICAL "Isar" (fn () => ! global_documents);
+fun get_documents () = ! global_documents;
fun init () = NAMED_CRITICAL "Isar" (fn () =>
(global_states := Symtab.empty;
--- a/src/Pure/Isar/outer_keyword.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/Isar/outer_keyword.ML Wed Oct 28 11:43:06 2009 +0000
@@ -121,8 +121,8 @@
in
-fun get_commands () = CRITICAL (fn () => ! global_commands);
-fun get_lexicons () = CRITICAL (fn () => ! global_lexicons);
+fun get_commands () = ! global_commands;
+fun get_lexicons () = ! global_lexicons;
fun change_commands f = CRITICAL (fn () => Unsynchronized.change global_commands f);
fun change_lexicons f = CRITICAL (fn () => Unsynchronized.change global_lexicons f);
--- a/src/Pure/Isar/outer_syntax.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/Isar/outer_syntax.ML Wed Oct 28 11:43:06 2009 +0000
@@ -101,8 +101,8 @@
(* access current syntax *)
-fun get_commands () = CRITICAL (fn () => ! global_commands);
-fun get_markups () = CRITICAL (fn () => ! global_markups);
+fun get_commands () = ! global_commands;
+fun get_markups () = ! global_markups;
fun get_command () = Symtab.lookup (get_commands ());
fun get_syntax () = CRITICAL (fn () => (OuterKeyword.get_lexicons (), get_command ()));
--- a/src/Pure/Isar/proof_context.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/Isar/proof_context.ML Wed Oct 28 11:43:06 2009 +0000
@@ -122,7 +122,7 @@
val add_abbrev: string -> binding * term -> Proof.context -> (term * term) * Proof.context
val revert_abbrev: string -> string -> Proof.context -> Proof.context
val verbose: bool Unsynchronized.ref
- val setmp_verbose: ('a -> 'b) -> 'a -> 'b
+ val setmp_verbose_CRITICAL: ('a -> 'b) -> 'a -> 'b
val print_syntax: Proof.context -> unit
val print_abbrevs: Proof.context -> unit
val print_binds: Proof.context -> unit
@@ -1200,7 +1200,7 @@
val verbose = Unsynchronized.ref false;
fun verb f x = if ! verbose then f (x ()) else [];
-fun setmp_verbose f x = setmp_CRITICAL verbose true f x;
+fun setmp_verbose_CRITICAL f x = setmp_CRITICAL verbose true f x;
(* local syntax *)
--- a/src/Pure/Isar/toplevel.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/Isar/toplevel.ML Wed Oct 28 11:43:06 2009 +0000
@@ -552,7 +552,7 @@
local val hooks = Unsynchronized.ref ([]: (transition -> state -> state -> unit) list) in
fun add_hook f = CRITICAL (fn () => Unsynchronized.change hooks (cons f));
-fun get_hooks () = CRITICAL (fn () => ! hooks);
+fun get_hooks () = ! hooks;
end;
--- a/src/Pure/ML/ml_context.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/ML/ml_context.ML Wed Oct 28 11:43:06 2009 +0000
@@ -54,7 +54,7 @@
(* theorem bindings *)
-val stored_thms: thm list Unsynchronized.ref = Unsynchronized.ref [];
+val stored_thms: thm list Unsynchronized.ref = Unsynchronized.ref []; (* FIXME via context!? *)
fun ml_store sel (name, ths) =
let
@@ -195,6 +195,7 @@
fun eval_in ctxt verbose pos txt =
Context.setmp_thread_data (Option.map Context.Proof ctxt) (fn () => eval verbose pos txt) ();
+(* FIXME not thread-safe -- reference can be accessed directly *)
fun evaluate ctxt verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () =>
let
val _ = r := NONE;
--- a/src/Pure/Proof/extraction.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/Proof/extraction.ML Wed Oct 28 11:43:06 2009 +0000
@@ -716,8 +716,9 @@
quote name ^ " has no computational content")
in (Reconstruct.reconstruct_proof thy prop prf, vs) end;
- val defs = Library.foldl (fn (defs, (prf, vs)) =>
- fst (extr 0 defs vs [] [] [] prf)) ([], map prep_thm thms);
+ val defs =
+ fold (fn (prf, vs) => fn defs => fst (extr 0 defs vs [] [] [] prf))
+ (map prep_thm thms) [];
fun add_def (s, (vs, ((t, u), (prf, _)))) thy =
(case Sign.const_type thy (extr_name s vs) of
--- a/src/Pure/Proof/reconstruct.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/Proof/reconstruct.ML Wed Oct 28 11:43:06 2009 +0000
@@ -40,11 +40,11 @@
fun mk_var env Ts T =
let val (env', v) = Envir.genvar "a" (env, rev Ts ---> T)
- in (env', list_comb (v, map Bound (length Ts - 1 downto 0))) end;
+ in (list_comb (v, map Bound (length Ts - 1 downto 0)), env') end;
-fun mk_tvar (Envir.Envir {maxidx, tenv, tyenv}, s) =
- (Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv},
- TVar (("'t", maxidx + 1), s));
+fun mk_tvar S (Envir.Envir {maxidx, tenv, tyenv}) =
+ (TVar (("'t", maxidx + 1), S),
+ Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv});
val mk_abs = fold (fn T => fn u => Abs ("", T, u));
@@ -73,14 +73,14 @@
| infer_type thy env Ts vTs (t as Free (s, T)) =
if T = dummyT then (case Symtab.lookup vTs s of
NONE =>
- let val (env', T) = mk_tvar (env, [])
+ let val (T, env') = mk_tvar [] env
in (Free (s, T), T, Symtab.update_new (s, T) vTs, env') end
| SOME T => (Free (s, T), T, vTs, env))
else (t, T, vTs, env)
| infer_type thy env Ts vTs (Var _) = error "reconstruct_proof: internal error"
| infer_type thy env Ts vTs (Abs (s, T, t)) =
let
- val (env', T') = if T = dummyT then mk_tvar (env, []) else (env, T);
+ val (T', env') = if T = dummyT then mk_tvar [] env else (T, env);
val (t', U, vTs', env'') = infer_type thy env' (T' :: Ts) vTs t
in (Abs (s, T', t'), T' --> U, vTs', env'') end
| infer_type thy env Ts vTs (t $ u) =
@@ -90,7 +90,7 @@
in (case chaseT env2 T of
Type ("fun", [U', V]) => (t' $ u', V, vTs2, unifyT thy env2 U U')
| _ =>
- let val (env3, V) = mk_tvar (env2, [])
+ let val (V, env3) = mk_tvar [] env2
in (t' $ u', V, vTs2, unifyT thy env3 T (U --> V)) end)
end
| infer_type thy env Ts vTs (t as Bound i) = ((t, nth Ts i, vTs, env)
@@ -99,21 +99,24 @@
fun cantunify thy (t, u) = error ("Non-unifiable terms:\n" ^
Syntax.string_of_term_global thy t ^ "\n\n" ^ Syntax.string_of_term_global thy u);
-fun decompose thy Ts (env, p as (t, u)) =
- let fun rigrig (a, T) (b, U) uT ts us = if a <> b then cantunify thy p
- else apsnd flat (Library.foldl_map (decompose thy Ts) (uT env T U, ts ~~ us))
- in case pairself (strip_comb o Envir.head_norm env) p of
+fun decompose thy Ts (p as (t, u)) env =
+ let
+ fun rigrig (a, T) (b, U) uT ts us =
+ if a <> b then cantunify thy p
+ else apfst flat (fold_map (decompose thy Ts) (ts ~~ us) (uT env T U))
+ in
+ case pairself (strip_comb o Envir.head_norm env) p of
((Const c, ts), (Const d, us)) => rigrig c d (unifyT thy) ts us
| ((Free c, ts), (Free d, us)) => rigrig c d (unifyT thy) ts us
| ((Bound i, ts), (Bound j, us)) =>
rigrig (i, dummyT) (j, dummyT) (K o K) ts us
| ((Abs (_, T, t), []), (Abs (_, U, u), [])) =>
- decompose thy (T::Ts) (unifyT thy env T U, (t, u))
+ decompose thy (T::Ts) (t, u) (unifyT thy env T U)
| ((Abs (_, T, t), []), _) =>
- decompose thy (T::Ts) (env, (t, incr_boundvars 1 u $ Bound 0))
+ decompose thy (T::Ts) (t, incr_boundvars 1 u $ Bound 0) env
| (_, (Abs (_, T, u), [])) =>
- decompose thy (T::Ts) (env, (incr_boundvars 1 t $ Bound 0, u))
- | _ => (env, [(mk_abs Ts t, mk_abs Ts u)])
+ decompose thy (T::Ts) (incr_boundvars 1 t $ Bound 0, u) env
+ | _ => ([(mk_abs Ts t, mk_abs Ts u)], env)
end;
fun make_constraints_cprf thy env cprf =
@@ -125,7 +128,7 @@
in
(prop, prf, cs, Pattern.unify thy (t', u') env, vTs)
handle Pattern.Pattern =>
- let val (env', cs') = decompose thy [] (env, (t', u'))
+ let val (cs', env') = decompose thy [] (t', u') env
in (prop, prf, cs @ cs', env', vTs) end
| Pattern.Unif =>
cantunify thy (Envir.norm_term env t', Envir.norm_term env u')
@@ -135,10 +138,10 @@
let
val tvars = OldTerm.term_tvars prop;
val tfrees = OldTerm.term_tfrees prop;
- val (env', Ts) =
+ val (Ts, env') =
(case opTs of
- NONE => Library.foldl_map mk_tvar (env, map snd tvars @ map snd tfrees)
- | SOME Ts => (env, Ts));
+ NONE => fold_map mk_tvar (map snd tvars @ map snd tfrees) env
+ | SOME Ts => (Ts, env));
val prop' = subst_atomic_types (map TVar tvars @ map TFree tfrees ~~ Ts)
(forall_intr_vfs prop) handle Library.UnequalLengths =>
error ("Wrong number of type arguments for " ^
@@ -152,8 +155,10 @@
handle Subscript => error ("mk_cnstrts: bad variable index " ^ string_of_int i))
| mk_cnstrts env Ts Hs vTs (Abst (s, opT, cprf)) =
let
- val (env', T) = (case opT of
- NONE => mk_tvar (env, []) | SOME T => (env, T));
+ val (T, env') =
+ (case opT of
+ NONE => mk_tvar [] env
+ | SOME T => (T, env));
val (t, prf, cnstrts, env'', vTs') =
mk_cnstrts env' (T::Ts) (map (incr_boundvars 1) Hs) vTs cprf;
in (Const ("all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, SOME T, prf),
@@ -167,7 +172,7 @@
end
| mk_cnstrts env Ts Hs vTs (AbsP (s, NONE, cprf)) =
let
- val (env', t) = mk_var env Ts propT;
+ val (t, env') = mk_var env Ts propT;
val (u, prf, cnstrts, env'', vTs') = mk_cnstrts env' Ts (t::Hs) vTs cprf;
in (Logic.mk_implies (t, u), AbsP (s, SOME t, prf), cnstrts, env'', vTs')
end
@@ -178,7 +183,7 @@
add_cnstrt Ts t' (prf1 %% prf2) (cnstrts' @ cnstrts)
env'' vTs'' (u, u')
| (t, prf1, cnstrts', env'', vTs'') =>
- let val (env''', v) = mk_var env'' Ts propT
+ let val (v, env''') = mk_var env'' Ts propT
in add_cnstrt Ts v (prf1 %% prf2) (cnstrts' @ cnstrts)
env''' vTs'' (t, Logic.mk_implies (u, v))
end)
@@ -192,7 +197,7 @@
in (betapply (f, t'), prf % SOME t', cnstrts, env3, vTs2)
end
| (u, prf, cnstrts, env2, vTs2) =>
- let val (env3, v) = mk_var env2 Ts (U --> propT);
+ let val (v, env3) = mk_var env2 Ts (U --> propT);
in
add_cnstrt Ts (v $ t') (prf % SOME t') cnstrts env3 vTs2
(u, Const ("all", (U --> propT) --> propT) $ v)
@@ -202,14 +207,14 @@
(case head_norm (mk_cnstrts env Ts Hs vTs cprf) of
(Const ("all", Type ("fun", [Type ("fun", [T, _]), _])) $ f,
prf, cnstrts, env', vTs') =>
- let val (env'', t) = mk_var env' Ts T
+ let val (t, env'') = mk_var env' Ts T
in (betapply (f, t), prf % SOME t, cnstrts, env'', vTs')
end
| (u, prf, cnstrts, env', vTs') =>
let
- val (env1, T) = mk_tvar (env', []);
- val (env2, v) = mk_var env1 Ts (T --> propT);
- val (env3, t) = mk_var env2 Ts T
+ val (T, env1) = mk_tvar [] env';
+ val (v, env2) = mk_var env1 Ts (T --> propT);
+ val (t, env3) = mk_var env2 Ts T
in
add_cnstrt Ts (v $ t) (prf % SOME t) cnstrts env3 vTs'
(u, Const ("all", (T --> propT) --> propT) $ v)
@@ -267,7 +272,7 @@
(Pattern.unify thy (tn1, tn2) env, ps) handle Pattern.Unif =>
cantunify thy (tn1, tn2)
else
- let val (env', cs') = decompose thy [] (env, (tn1, tn2))
+ let val (cs', env') = decompose thy [] (tn1, tn2) env
in if cs' = [(tn1, tn2)] then
apsnd (cons (false, (tn1, tn2), vs)) (search env ps)
else search env' (map (fn q => (true, q, vs)) cs' @ ps)
--- a/src/Pure/System/isabelle_process.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/System/isabelle_process.ML Wed Oct 28 11:43:06 2009 +0000
@@ -61,9 +61,6 @@
else message_pos ts
| _ => NONE);
-fun output out_stream s = NAMED_CRITICAL "IO" (fn () =>
- (TextIO.output (out_stream, s); TextIO.output (out_stream, "\n")));
-
in
fun message _ _ "" = ()
@@ -74,14 +71,16 @@
Position.properties_of (Position.thread_data ())
|> Position.default_properties pos;
val txt = clean_string [Symbol.STX] body;
- in output out_stream (special ch ^ message_props props ^ txt ^ special_end) end;
+ val msg = special ch ^ message_props props ^ txt ^ special_end ^ "\n";
+ in TextIO.output (out_stream, msg) end;
fun init_message out_stream =
let
val pid = (Markup.pidN, process_id ());
val session = (Markup.sessionN, List.last (Session.id ()) handle List.Empty => "unknown");
val text = Session.welcome ();
- in output out_stream (special "A" ^ message_props [pid, session] ^ text ^ special_end) end;
+ val msg = special "A" ^ message_props [pid, session] ^ text ^ special_end ^ "\n";
+ in TextIO.output (out_stream, msg) end;
end;
--- a/src/Pure/System/isar.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/System/isar.ML Wed Oct 28 11:43:06 2009 +0000
@@ -47,10 +47,10 @@
| edit n (st, hist) = edit (n - 1) (f st hist);
in edit count (! global_state, ! global_history) end);
-fun state () = NAMED_CRITICAL "Isar" (fn () => ! global_state);
+fun state () = ! global_state;
-fun exn () = NAMED_CRITICAL "Isar" (fn () => ! global_exn);
-fun set_exn exn = NAMED_CRITICAL "Isar" (fn () => global_exn := exn);
+fun exn () = ! global_exn;
+fun set_exn exn = global_exn := exn;
end;
--- a/src/Pure/Thy/html.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/Thy/html.ML Wed Oct 28 11:43:06 2009 +0000
@@ -267,7 +267,7 @@
(* document *)
val charset = Unsynchronized.ref "ISO-8859-1";
-fun with_charset s = setmp_noncritical charset s; (* FIXME *)
+fun with_charset s = setmp_noncritical charset s; (* FIXME unreliable *)
fun begin_document title =
let val cs = ! charset in
--- a/src/Pure/Thy/thy_load.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/Thy/thy_load.ML Wed Oct 28 11:43:06 2009 +0000
@@ -41,16 +41,22 @@
fun show_path () = map Path.implode (! load_path);
-fun del_path s = CRITICAL (fn () =>
- Unsynchronized.change load_path (remove (op =) (Path.explode s)));
-fun add_path s = CRITICAL (fn () =>
- (del_path s; Unsynchronized.change load_path (cons (Path.explode s))));
-fun path_add s = CRITICAL (fn () =>
+fun del_path s =
+ CRITICAL (fn () => Unsynchronized.change load_path (remove (op =) (Path.explode s)));
+
+fun add_path s =
+ CRITICAL (fn () => (del_path s; Unsynchronized.change load_path (cons (Path.explode s))));
+
+fun path_add s =
+ CRITICAL (fn () =>
(del_path s; Unsynchronized.change load_path (fn path => path @ [Path.explode s])));
+
fun reset_path () = load_path := [Path.current];
fun with_paths ss f x =
- CRITICAL (fn () => setmp_CRITICAL load_path (! load_path @ map Path.explode ss) f x);
+ CRITICAL (fn () =>
+ setmp_CRITICAL load_path (! load_path @ map Path.explode ss) f x);
+
fun with_path s f x = with_paths [s] f x;
fun search_path dir path =
--- a/src/Pure/codegen.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/codegen.ML Wed Oct 28 11:43:06 2009 +0000
@@ -105,7 +105,7 @@
val quiet_mode = Unsynchronized.ref true;
fun message s = if !quiet_mode then () else writeln s;
-val mode = Unsynchronized.ref ([] : string list);
+val mode = Unsynchronized.ref ([] : string list); (* FIXME proper functional argument *)
val margin = Unsynchronized.ref 80;
@@ -928,7 +928,7 @@
[str "(result ())"],
str ");"]) ^
"\n\nend;\n";
- val _ = NAMED_CRITICAL "codegen" (fn () =>
+ val _ = NAMED_CRITICAL "codegen" (fn () => (* FIXME ??? *)
ML_Context.eval_in (SOME ctxt) false Position.none s);
in !eval_result end;
in e () end;
--- a/src/Pure/meta_simplifier.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Pure/meta_simplifier.ML Wed Oct 28 11:43:06 2009 +0000
@@ -1139,8 +1139,8 @@
end
and mut_impc [] concl [] [] prems' rrss' asms' eqns ss changed k =
- transitive1 (Library.foldl (fn (eq2, (eq1, prem)) => transitive1 eq1
- (Option.map (disch false prem) eq2)) (NONE, eqns ~~ prems'))
+ transitive1 (fold (fn (eq1, prem) => fn eq2 => transitive1 eq1
+ (Option.map (disch false prem) eq2)) (eqns ~~ prems') NONE)
(if changed > 0 then
mut_impc (rev prems') concl (rev rrss') (rev asms')
[] [] [] [] ss ~1 changed
--- a/src/Tools/eqsubst.ML Wed Oct 28 11:42:31 2009 +0000
+++ b/src/Tools/eqsubst.ML Wed Oct 28 11:43:06 2009 +0000
@@ -269,7 +269,7 @@
(* FOR DEBUGGING...
type trace_subst_errT = int (* subgoal *)
* thm (* thm with all goals *)
- * (Thm.cterm list (* certified free var placeholders for vars *)
+ * (cterm list (* certified free var placeholders for vars *)
* thm) (* trivial thm of goal concl *)
(* possible matches/unifiers *)
* thm (* rule *)