--- a/ex/PL.ML Fri Jun 17 14:15:38 1994 +0200
+++ b/ex/PL.ML Fri Jun 17 14:16:50 1994 +0200
@@ -135,7 +135,7 @@
(** The function eval **)
-val pl_ss = set_ss addsimps [pl_rec_var,pl_rec_false,pl_rec_imp] @ PL0.simps;
+val pl_ss = set_ss addsimps [pl_rec_var,pl_rec_false,pl_rec_imp] @ PL0.pl.simps;
goalw PL.thy [eval_def] "tt[false] = False";
by (simp_tac pl_ss 1);
@@ -212,7 +212,7 @@
(*This formulation is required for strong induction hypotheses*)
goal PL.thy "hyps(p,tt) |- if(tt[p], p, p->false)";
by (rtac (expand_if RS iffD2) 1);
-by(PL0.induct_tac "p" 1);
+by(PL0.pl.induct_tac "p" 1);
by (ALLGOALS (simp_tac (pl_ss addsimps [conseq_I, conseq_H])));
by (fast_tac (set_cs addIs [weaken_left_Un1, weaken_left_Un2,
weaken_right, imp_false]
@@ -249,7 +249,7 @@
(*For the case hyps(p,t)-insert(#v,Y) |- p;
we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *)
goal PL.thy "hyps(p, t-{v}) <= insert((#v)->false, hyps(p,t)-{#v})";
-by (PL0.induct_tac "p" 1);
+by (PL0.pl.induct_tac "p" 1);
by (simp_tac pl_ss 1);
by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1);
by (simp_tac pl_ss 1);
@@ -259,7 +259,7 @@
(*For the case hyps(p,t)-insert(#v -> false,Y) |- p;
we also have hyps(p,t)-{(#v)->false} <= hyps(p, insert(v,t)) *)
goal PL.thy "hyps(p, insert(v,t)) <= insert(#v, hyps(p,t)-{(#v)->false})";
-by (PL0.induct_tac "p" 1);
+by (PL0.pl.induct_tac "p" 1);
by (simp_tac pl_ss 1);
by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1);
by (simp_tac pl_ss 1);
@@ -279,7 +279,7 @@
(*The set hyps(p,t) is finite, and elements have the form #v or #v->false;
could probably prove the stronger hyps(p,t) : Fin(hyps(p,{}) Un hyps(p,nat))*)
goal PL.thy "hyps(p,t) : Fin(UN v:{x.True}. {#v, (#v)->false})";
-by (PL0.induct_tac "p" 1);
+by (PL0.pl.induct_tac "p" 1);
by (ALLGOALS (simp_tac (pl_ss setloop (split_tac [expand_if])) THEN'
fast_tac (set_cs addSIs [Fin_0I, Fin_insertI, Fin_UnI])));
val hyps_finite = result();
--- a/ex/PL0.thy Fri Jun 17 14:15:38 1994 +0200
+++ b/ex/PL0.thy Fri Jun 17 14:16:50 1994 +0200
@@ -13,4 +13,6 @@
false :: "'a pl"
"->" :: "['a pl,'a pl] => 'a pl" (infixr 90)
var :: "'a => 'a pl" ("#_")
+datatype
+ 'a pl = false | var('a) | "op->" ('a pl,'a pl)
end