new code equation for ntrancl that allows computation of the transitive closure of sets on infinite types as well
--- a/src/HOL/Transitive_Closure.thy Fri Jan 27 17:02:08 2012 +0100
+++ b/src/HOL/Transitive_Closure.thy Fri Jan 27 17:22:05 2012 +0100
@@ -1003,7 +1003,7 @@
unfolding ntrancl_def by auto
qed
-lemma ntrancl_Suc [simp, code]:
+lemma ntrancl_Suc [simp]:
"ntrancl (Suc n) R = ntrancl n R O (Id \<union> R)"
proof
{
@@ -1034,6 +1034,10 @@
unfolding ntrancl_def by fastforce
qed
+lemma [code]:
+ "ntrancl (Suc n) r = (let r' = ntrancl n r in r' Un r' O r)"
+unfolding Let_def by auto
+
lemma finite_trancl_ntranl:
"finite R \<Longrightarrow> trancl R = ntrancl (card R - 1) R"
by (cases "card R") (auto simp add: trancl_finite_eq_rel_pow rel_pow_empty ntrancl_def)