A new theory: a model of bytecode verification.
authornipkow
Tue, 28 Sep 1999 16:36:12 +0200
changeset 7626 5997f35954d7
parent 7625 94b2a50e69a5
child 7627 6b0709a2f6c7
A new theory: a model of bytecode verification.
src/HOL/BCV/DFAandWTI.ML
src/HOL/BCV/DFAandWTI.thy
src/HOL/BCV/DFAimpl.ML
src/HOL/BCV/DFAimpl.thy
src/HOL/BCV/Fixpoint.ML
src/HOL/BCV/Fixpoint.thy
src/HOL/BCV/Machine.ML
src/HOL/BCV/Machine.thy
src/HOL/BCV/Orders.ML
src/HOL/BCV/Orders.thy
src/HOL/BCV/Orders0.ML
src/HOL/BCV/Orders0.thy
src/HOL/BCV/Plus.ML
src/HOL/BCV/Plus.thy
src/HOL/BCV/README.html
src/HOL/BCV/ROOT.ML
src/HOL/BCV/SemiLattice.ML
src/HOL/BCV/SemiLattice.thy
src/HOL/BCV/Types.ML
src/HOL/BCV/Types.thy
src/HOL/BCV/Types0.ML
src/HOL/BCV/Types0.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/DFAandWTI.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,79 @@
+(*  Title:      HOL/BCV/DFAandWTI.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+*)
+
+Goalw [wti_is_fix_step_def]
+ "[| wti_is_fix_step step wti succs n L; \
+\    sos : listsn n (option L); p < n |] ==> \
+\ wti p sos = stable step succs p sos";
+by(Asm_simp_tac 1);
+qed "wti_is_fix_stepD";
+
+Goalw [is_dfa_def]
+ "[| is_dfa dfa step succs n L; sos : listsn n (option L) |] ==> \
+\ dfa sos = (? tos : listsn (size sos) (option L). \
+\             sos <= tos & (!p < size tos. stable step succs p tos))";
+by(Asm_simp_tac 1);
+qed "is_dfaD";
+
+Goalw [welltyping_def]
+ "[| wti_is_fix_step step wti succs n L; is_dfa dfa step succs n L; \
+\    sos : listsn n (option L) |] ==> \
+\ dfa sos = (? tos : listsn n (option L). sos <= tos & welltyping wti tos)";
+bd is_dfaD 1;
+ba 1;
+by(Asm_full_simp_tac 1);
+by(blast_tac (claset() addDs [wti_is_fix_stepD]) 1);
+qed "dfa_iff_welltyping";
+
+Goal
+ "x:M ==> replicate n x : listsn n M";
+by(induct_tac "n" 1);
+by(Simp_tac 1);
+by(asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff,Bex_def]) 1);
+qed "replicate_in_listsn";
+Addsimps [replicate_in_listsn];
+
+Goal "!ys. replicate n x <= ys = (length ys = n & (!y:set ys. x <= y))";
+by(induct_tac "n" 1);
+ by(Force_tac 1);
+by(asm_full_simp_tac (simpset() addsimps [Cons_le_iff,length_Suc_conv]) 1);
+by(Force_tac 1);
+qed_spec_mp "replicate_le_conv";
+AddIffs [replicate_le_conv];
+
+Goal
+ "[| wti_is_fix_step step wti succs n L; is_dfa dfa step succs n L; \
+\    0 < n; init : (option L) |] ==> \
+\ dfa (init # replicate (n-1) None) = \
+\ (? tos : listsn n (option L). init <= tos!0 & welltyping wti tos)";
+by(subgoal_tac "? m. n = m+1" 1);
+ by(res_inst_tac [("x","n-1")] exI 2);
+ by(arith_tac 2);
+by(Clarify_tac 1);
+by(Asm_full_simp_tac 1);
+bd dfa_iff_welltyping 1;
+  ba 1;
+ be trans 2;
+ by(asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff,Bex_def]) 1);
+by(asm_full_simp_tac (simpset() addsimps [Cons_le_iff]) 1);
+br iffI 1;
+ by(Clarify_tac 1);
+ br bexI 1;
+  br conjI 1;
+   ba 2;
+  ba 2;
+ by(Asm_simp_tac 1);
+by(Clarify_tac 1);
+by(asm_full_simp_tac (simpset() addsimps [in_listsn_Suc_iff,Bex_def]) 1);
+by(Clarify_tac 1);
+by(asm_full_simp_tac (simpset() addsimps []) 1);
+br exI 1;
+br conjI 1;
+ br conjI 2;
+  ba 3;
+ by(Blast_tac 1);
+by(Force_tac 1);
+qed "dfa_iff_welltyping0";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/DFAandWTI.thy	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,41 @@
+(*  Title:      HOL/BCV/DFAandWTI.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+The relationship between dataflow analysis and a welltyped-insruction predicate.
+*)
+
+DFAandWTI = SemiLattice +
+
+types 's os = 's option list
+
+constdefs
+
+ stable :: (nat => 's => 's option) => (nat => nat list) => nat => 's os => bool
+"stable step succs p sos ==
+ !s. sos!p = Some s --> (? t. step p s = Some(t) &
+                              (!q:set(succs p). Some t <= sos!q))"
+
+ wti_is_fix_step ::
+   "(nat => 's => 's option)
+    => (nat => 's os => bool)
+    => (nat => nat list)
+    => nat => 's set => bool"
+"wti_is_fix_step step wti succs n L == !sos p.
+   sos : listsn n (option L) & p < n -->
+   wti p sos = stable step succs p sos"
+
+ welltyping :: (nat => 's os => bool) => 's os => bool
+"welltyping wti tos == !p<size(tos). wti p tos"
+
+ is_dfa ::  ('s os => bool)
+            => (nat => 's => 's option)
+            => (nat => nat list)
+            => nat => 's set => bool
+"is_dfa dfa step succs n L == !sos : listsn n (option L).
+   dfa sos =
+   (? tos : listsn n (option L).
+      sos <= tos & (!p < n. stable step succs p tos))"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/DFAimpl.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,336 @@
+(*  Title:      HOL/BCV/DFAimpl.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+*)
+
+(** merges **)
+
+Goal "!sos. size(merges t ps sos) = size sos";
+by(induct_tac "ps" 1);
+by(Auto_tac);
+qed_spec_mp "length_merges";
+Addsimps [length_merges];
+
+Goal
+ "!xs. xs : listsn n (option A) --> x : A --> (!p : set ps. p<n) --> \
+\      semilat A --> merges x ps xs : listsn n (option A)";
+by(induct_tac "ps" 1);
+by(Auto_tac);
+qed_spec_mp "merges_preserves_type";
+Addsimps [merges_preserves_type];
+(*AddSIs [merges_preserves_type];*)
+
+Goal "semilat A ==> !xs. xs : listsn n (option A) --> x:A --> (!p:set ps. p<n) \
+\ --> xs <= merges x ps xs";
+by(induct_tac "ps" 1);
+ by(Simp_tac 1);
+by(Simp_tac 1);
+by(Clarify_tac 1);
+br order_trans 1;
+ be list_update_incr 1;
+   by(Blast_tac 2);
+  ba 2;
+ be (Some_in_option RS iffD2) 1;
+by(blast_tac (claset() addSIs [listsnE_set]
+          addIs [semilat_plus,listsnE_length RS nth_in]) 1);
+qed_spec_mp "merges_incr";
+
+(* is redundant but useful *)
+Goal "[| xs!i = Some x; xs : listsn n (option A); i < n |] ==> x:A";
+bd listsnE_nth_in 1;
+ba 1;
+by(Asm_full_simp_tac 1);
+qed "listsn_optionE_in";
+(*Addsimps [listsn_optionE_in];*)
+
+Goal
+ "[| semilat L |] ==> \
+\ (!xs. xs : listsn n (option L) --> x:L --> (!p:set ps. p<n) --> \
+\      (merges x ps xs = xs) = (!p:set ps. Some x <= xs!p))";
+by(induct_tac "ps" 1);
+ by(Asm_simp_tac 1);
+by(Clarsimp_tac 1);
+by(rename_tac "p ps xs" 1);
+br iffI 1;
+ br context_conjI 1;
+  by(subgoal_tac "xs[p := Some x + xs!p] <= xs" 1);
+   by(EVERY[etac subst 2, rtac merges_incr 2]);
+      ba 2;
+     by(Force_tac 2);
+    ba 2;
+   ba 2;
+  by(exhaust_tac "xs!p" 1);
+   by(asm_full_simp_tac (simpset() addsimps [list_update_le_conv]) 1);
+  by(asm_full_simp_tac (simpset() addsimps
+      [list_update_le_conv,listsn_optionE_in]) 1);
+ by(Clarify_tac 1);
+ by(rotate_tac ~3 1);
+ by(asm_full_simp_tac (simpset() addsimps
+      [le_iff_plus_unchanged RS iffD1,listsn_optionE_in,
+       list_update_same_conv RS iffD2]) 1);
+by(Clarify_tac 1);
+by(asm_simp_tac (simpset() addsimps
+      [le_iff_plus_unchanged RS iffD1,listsn_optionE_in,
+       list_update_same_conv RS iffD2]) 1);
+qed_spec_mp "merges_same_conv";
+
+
+Goalw [le_list,plus_option]
+ "xs <= ys --> xs : listsn n (option L) --> ys : listsn n (option L) --> \
+\ p < n --> ys!p = Some y --> x <= y --> x : L --> semilat L --> \
+\ xs[p := Some x + xs!p] <= ys";
+by(simp_tac (simpset()  addsimps [nth_list_update]
+                        addsplits [option.split]) 1);
+by(Clarify_tac 1);
+by(rotate_tac 3 1);
+by(force_tac (claset() addEs [listsn_optionE_in] addIs [semilat_lub],
+              simpset()) 1);
+qed_spec_mp "list_update_le_listI";
+
+Goal
+ "[| semilat(L); t:L; tos : listsn n (option L); \
+\    !p. p:set ps --> Some t <= tos!p; \
+\    !p. p:set ps --> p<n |] ==> \
+\ set qs <= set ps  --> \
+\ (!sos. sos : listsn n (option L) & sos <= tos --> merges t qs sos <= tos)";
+by(induct_tac "qs" 1);
+ by(Asm_simp_tac 1);
+by(force_tac (claset(),simpset() addsimps [list_update_le_listI]) 1);
+val lemma = result();
+
+Goal
+ "[| semilat(L); t:L; \
+\    !p. p:set ps --> Some t <= tos!p; \
+\    !p. p:set ps --> p<n; \
+\    sos <= tos; sos : listsn n (option L); tos : listsn n (option L) |] \
+\ ==> merges t ps sos <= tos";
+by(blast_tac (claset() addDs [lemma]) 1);
+qed "merges_pres_le_ub";
+
+
+(** next **)
+
+Goalw [is_next_def]
+ "[| is_next next; next step succs sos = None; succs_bounded succs n; \
+\    sos : listsn n S |] ==> \
+\ ? p<n. ? s. sos!p = Some s & step p s = None";
+by(subgoal_tac "n=size sos" 1);
+by(Blast_tac 1);
+by(Asm_simp_tac 1);
+qed "next_None";
+
+Goalw [is_next_def]
+ "[| is_next next; succs_bounded succs n; sos : listsn n S |] ==> \
+\ next step succs sos = Some sos --> \
+\ (!p<n. !s. sos!p = Some s --> (? t. \
+\         step p s = Some(t) & merges t (succs p) sos = sos))";
+by(subgoal_tac "n=size sos" 1);
+by(Blast_tac 1);
+by(Asm_simp_tac 1);
+qed "next_Some1";
+
+Goalw [is_next_def]
+ "[| is_next next; next step succs sos = Some sos'; sos' ~= sos; \
+\    succs_bounded succs n; sos : listsn n S |] ==> \
+\ ? p<n. ? s. sos!p = Some s & (? t. \
+\     step p s = Some(t) & merges t (succs p) sos = sos')";
+by(subgoal_tac "n=size sos" 1);
+by(Blast_tac 1);
+by(Asm_simp_tac 1);
+qed "next_Some2";
+
+Goal
+ "[| is_next next; succs_bounded succs n; sos : listsn n S |] ==> \
+\ (next step succs sos = Some sos) = \
+\ (!p<n. !s. sos!p = Some s --> (? t. \
+\         step p s = Some(t) & merges t (succs p) sos = sos))";
+br iffI 1;
+ by(asm_simp_tac (simpset() addsimps [next_Some1]) 1);
+by(exhaust_tac "next step succs sos" 1);
+ by(best_tac (claset() addSDs [next_None] addss simpset()) 1);
+by(rename_tac "sos'" 1);
+by(case_tac "sos' = sos" 1);
+ by(Blast_tac 1);
+by(best_tac (claset() addSDs [next_Some2] addss simpset()) 1);
+qed "next_Some1_eq";
+
+Addsimps [next_Some1_eq];
+
+Goalw [step_pres_type_def]
+ "[| step_pres_type step n L; s:L; p<n; step p s = Some(t) |] ==> t:L";
+by(Blast_tac 1);
+qed "step_pres_typeD";
+
+Goalw [succs_bounded_def]
+ "[| succs_bounded succs n; p < n; q : set(succs p) |] ==> q < n";
+by(Blast_tac 1);
+qed "succs_boundedD";
+
+Goal
+ "[| is_next next; semilat A; \
+\    step_pres_type step n A; succs_bounded succs n;\
+\    sos : listsn n (option A) |] ==> \
+\ next step succs sos : option (listsn n (option A))";
+by(exhaust_tac "next step succs sos" 1);
+ by(ALLGOALS Asm_simp_tac);
+by(rename_tac "sos'" 1);
+by(case_tac "sos' = sos" 1);
+ by(Asm_simp_tac 1);
+by(blast_tac (claset() addIs [step_pres_typeD,succs_boundedD,listsn_optionE_in] addSIs [merges_preserves_type] addDs [next_Some2]) 1);
+qed_spec_mp "next_preserves_type";
+
+Goal
+ "[| is_next next; semilat A; \
+\    step_pres_type step n A; succs_bounded succs n; \
+\    next step succs xs = Some ys; xs : listsn n (option A) |] ==> xs <= ys";
+by(case_tac "ys = xs" 1);
+ by(Asm_full_simp_tac 1);
+by(blast_tac (claset() addSIs [merges_incr] addIs [listsn_optionE_in]
+       addDs [step_pres_typeD,succs_boundedD,next_Some2]) 1);
+qed_spec_mp "next_incr";
+
+val lemma = (Unify.trace_bound, Unify.search_bound);
+Unify.trace_bound := 50;
+Unify.search_bound := 50;
+
+Goalw [is_next_def]
+ "is_next (%step succs sos. itnext (size sos) step succs sos)";
+by(Clarify_tac 1);
+by(etac thin_rl 1);
+by(res_inst_tac [("n","length sos")] nat_induct 1);
+ by(Asm_full_simp_tac 1);
+by(asm_full_simp_tac (simpset() addsimps [less_Suc_eq,Let_def]
+                                addsplits [option.split])1);
+by(Blast_tac 1);
+qed "is_next_itnext";
+
+Unify.trace_bound := !(fst lemma);
+Unify.search_bound := !(snd lemma);
+
+(** fix step **)
+
+Goalw [step_mono_None_def]
+ "[| step_mono_None step n L; s : L; p < n; s <= t; step p s = None |] ==> \
+\ step p t = None";
+by(Blast_tac 1);
+qed "step_mono_NoneD";
+
+Goalw [step_mono_def]
+ "[| step_mono step n L; s : L; p < n; s <= t; step p s = Some(u) |] ==> \
+\ !v. step p t = Some(v) --> u <= v";
+by(Blast_tac 1);
+qed "step_monoD";
+
+Goalw [stable_def]
+"[| is_next next; semilat L; sos : listsn n (option L); \
+\   step_pres_type step n L; succs_bounded succs n |] \
+\ ==> (next step succs sos = Some sos) = (!p<n. stable step succs p sos)";
+by(Asm_simp_tac 1);
+br iffI 1;
+ by(Clarify_tac 1);
+ by(etac allE 1 THEN mp_tac 1);
+ by(etac allE 1 THEN mp_tac 1);
+ by(Clarify_tac 1);
+ bd(merges_same_conv RS iffD1)1;
+     ba 4;
+    ba 1;
+   by(blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
+  by(blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
+ by(Blast_tac 1);
+by(Clarify_tac 1);
+by(etac allE 1 THEN mp_tac 1);
+by(Asm_full_simp_tac 1);
+by(Clarify_tac 1);
+by(Asm_simp_tac 1);
+by(blast_tac (claset() addSIs [merges_same_conv RS iffD2]
+     addIs [step_pres_typeD,succs_boundedD,listsn_optionE_in]) 1);
+qed "fixpoint_next_iff_stable";
+
+Goal
+ "[| semilat L; is_next next; succs_bounded succs n; \
+\    step_pres_type step n L; step_mono_None step n L; step_mono step n L; \
+\    tos:listsn n (option L); next step succs tos = Some tos; \
+\    sos:listsn n (option L); sos <= tos |] \
+\ ==> ? sos'. next step succs sos = Some sos' & sos' <= tos";
+by(subgoal_tac
+   "!p<n. !s. sos!p = Some s --> (? u. \
+\             step p s = Some(u) & (!q:set(succs p). Some u<=tos!q))" 1);
+ by(exhaust_tac "next step succs sos" 1);
+  bd next_None 1;
+     ba 1;
+    ba 1;
+   ba 1;
+  by(Force_tac 1);
+ by(rename_tac "sos'" 1);
+ by(case_tac "sos' = sos" 1);
+  by(Blast_tac 1);
+ bd next_Some2 1;
+    by(EVERY1[atac, atac, atac, atac]);
+ by(Clarify_tac 1);
+ by(etac allE 1 THEN mp_tac 1);
+ by(etac allE 1 THEN mp_tac 1);
+ by(Clarify_tac 1);
+ by(EVERY1[rtac exI, rtac conjI, atac]);
+ br merges_pres_le_ub 1;
+       ba 1;
+      by(blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
+     by(Asm_full_simp_tac 1);
+    by(blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
+   by(REPEAT(atac 1));
+by(Clarify_tac 1);
+by(exhaust_tac "tos!p" 1);
+ by(force_tac (claset() addDs [le_listD],simpset()) 1);
+by(rename_tac "t" 1);
+by(subgoal_tac "s <= t" 1);
+ by(force_tac (claset() addDs [le_listD],simpset()) 2);
+by(exhaust_tac "step p s" 1);
+ bd step_mono_NoneD 1;
+     ba 4;
+    by(blast_tac (claset() addIs [listsn_optionE_in]) 1);
+   ba 1;
+  ba 1;
+ by(Force_tac 1);
+bd step_monoD 1;
+    ba 4;
+   by(blast_tac (claset() addIs [listsn_optionE_in]) 1);
+  ba 1;
+ ba 1;
+by(Asm_full_simp_tac 1);
+by(etac allE 1 THEN mp_tac 1);
+by(etac allE 1 THEN mp_tac 1);
+by(Clarify_tac 1);
+by(Asm_full_simp_tac 1);
+by(forward_tac[merges_same_conv RS iffD1]1);
+    ba 4;
+   ba 1;
+  by(blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
+ by(blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
+by(blast_tac (claset() addIs [order_trans]) 1);
+val lemma = result();
+
+Goalw [is_dfa_def]
+ "[| semilat L; acc L; is_next next; \
+\    step_pres_type step n L; succs_bounded succs n; \
+\    step_mono_None step n L; step_mono step n L |] ==> \
+\ is_dfa (%sos. fix(next step succs, sos)) step succs n L";
+by(Clarify_tac 1);
+by(stac fix_iff_has_fixpoint 1);
+     by(etac (acc_option RS acc_listsn) 1);
+    by(blast_tac (claset() addIs [lemma]) 1);
+   by(blast_tac (claset() addIs [next_preserves_type]) 1);
+  by(blast_tac (claset() addIs [next_incr]) 1);
+ ba 1;
+by(asm_simp_tac (simpset() delsimps [next_Some1_eq] addsimps [fixpoint_next_iff_stable]) 1);
+qed "is_dfa_fix_next";
+
+Goal
+ "[| semilat L; acc L; is_next next; \
+\    step_pres_type step n L; succs_bounded succs n; \
+\    step_mono_None step n L; step_mono step n L; \
+\    wti_is_fix_step step wti succs n L; \
+\    sos : listsn n (option L) |] ==> \
+\ fix(next step succs, sos) = \
+\ (? tos:listsn n (option L). sos<=tos & welltyping wti tos)";
+by(blast_tac (claset() addSDs [dfa_iff_welltyping] addSIs [is_dfa_fix_next]) 1);
+qed "fix_next_iff_welltyping";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/DFAimpl.thy	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,60 @@
+(*  Title:      HOL/BCV/DFAimpl.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+Implementations of data-flow analysis.
+*)
+
+DFAimpl = DFAandWTI + Fixpoint +
+
+consts merges :: "('s::plus) => nat list => 's os => 's os"
+primrec
+"merges a []     s = s"
+"merges a (p#ps) s = merges a ps (s[p := Some(a) + s!p])"
+
+constdefs
+ succs_bounded :: "(nat => nat list) => nat => bool"
+"succs_bounded succs n == !p<n. !q:set(succs p). q<n"
+
+ is_next :: "((nat => 's => ('s::plus)option) => (nat => nat list)
+              => 's os => 's os option)
+             => bool"
+"is_next next == !step succs sos sos'.
+   succs_bounded succs (size sos) -->
+   (next step succs sos = None -->
+      (? p<size sos. ? s. sos!p = Some s & step p s = None))  &
+   (next step succs sos = Some sos -->
+      (!p<size sos. !s. sos!p = Some s --> (? t.
+          step p s = Some(t) & merges t (succs p) sos = sos))) &
+   (next step succs sos = Some sos' & sos' ~= sos -->
+      (? p<size sos. ? s. sos!p = Some s & (? t.
+          step p s = Some(t) & merges t (succs p) sos = sos')))"
+
+ step_pres_type :: "(nat => 's => 's option) => nat => 's set => bool"
+"step_pres_type step n L == !s:L. !p<n. !t. step p s = Some(t) --> t:L"
+
+ step_mono_None :: "(nat => 's => 's option) => nat => 's set => bool"
+"step_mono_None step n L == !s p t.
+   s : L & p < n & s <= t & step p s = None --> step p t = None"
+
+ step_mono :: "(nat => 's => 's option) => nat => 's set => bool"
+"step_mono step n L == !s p t u.
+   s : L & p < n & s <= t & step p s = Some(u)
+   --> (!v. step p t = Some(v) --> u <= v)"
+
+consts
+itnext :: nat => (nat => 's => 's option) => (nat => nat list)
+          => 's os => ('s::plus) os option
+primrec
+"itnext 0       step succs sos = Some sos"
+"itnext (Suc p) step succs sos =
+   (case sos!p of
+      None => itnext p step succs sos
+    | Some s => (case step p s of None => None
+                 | Some t => let sos' = merges t (succs p) sos
+                             in if sos' = sos
+                                then itnext p step succs sos
+                                else Some sos'))"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Fixpoint.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,142 @@
+(*  Title:      HOL/BCV/Fixpoint.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+*)
+
+goalw_cterm [] (cterm_of (sign_of thy) (HOLogic.mk_Trueprop (hd fix.tcs)));
+by(asm_full_simp_tac (simpset() addsimps [wf_iff_no_infinite_down_chain]
+                                addsplits [split_split]) 1);
+by(Clarify_tac 1);
+br ccontr 1;
+by(Asm_full_simp_tac 1);
+by(res_inst_tac[("x","0")] allE 1);
+ba 1;
+by(Clarify_tac 1);
+by(rename_tac "next s0" 1);
+by(subgoal_tac "!i. fst(f i) = next" 1);
+by(eres_inst_tac[("x","%i. snd(f i)")] allE 1);
+by(Asm_full_simp_tac 1);
+by(Clarify_tac 1);
+by(eres_inst_tac[("x","i")] allE 1);
+by(eres_inst_tac[("x","i")] allE 1);
+by(Force_tac 1);
+br allI 1;
+by(induct_tac "i" 1);
+by(Asm_simp_tac 1);
+by(eres_inst_tac[("x","n")] allE 1);
+by(Auto_tac);
+val fix_tc = result();
+
+Goalw [wf_iff_no_infinite_down_chain RS eq_reflection]
+ "[| wf{(t,s). s:A & next s = Some t & t ~= s}; \
+\    !a:A. next a : option A; s:A |] ==> \
+\ fix(next,s) = (case next s of None => False \
+\                | Some t => if t=s then True else fix(next,t))";
+by(stac (fix_tc RS (hd fix.rules)) 1);
+by(Simp_tac 1);
+by(Clarify_tac 1);
+by(subgoal_tac "!i. f i : A" 1);
+ by(Blast_tac 1);
+br allI 1;
+by(induct_tac "i" 1);
+ by(Asm_simp_tac 1);
+by(Auto_tac);
+qed "fix_unfold";
+
+(* Thm: next has fixpoint above s iff fix(next,s) *)
+
+Goal "[| x = Some y; x : option A |] ==> y : A";
+by(Blast_tac 1);
+val lemma = result();
+
+Goal
+"[| acc L; \
+\   !t:L. !s:L. next t = Some t & s <= t --> (? u. next s = Some u & u <= t); \
+\   !a:L. next a : option L; !a:L. !b. next a = Some b --> a <= b |] \
+\ ==> s:L --> fix(next,s) = (? t:L. s <= t & next t = Some t)";
+by(subgoal_tac "wf{(y,x::'a::order). x : L & next x = Some y & y ~= x}" 1);
+ by(full_simp_tac (simpset() addsimps [acc_def]) 2);
+ be wf_subset 2;
+ by(simp_tac (simpset() addsimps [order_less_le]) 2);
+ by(blast_tac (claset() addDs [lemma]) 2);
+by(res_inst_tac [("a","s")] wf_induct 1);
+ ba 1;
+by(Clarify_tac 1);
+by(Full_simp_tac 1);
+by(stac fix_unfold 1);
+   ba 1;
+  ba 1;
+ ba 1;
+by(split_tac [option.split] 1);
+br conjI 1;
+ by(thin_tac "!a:L. !b. next a = Some b --> a <= b" 1);
+ by(Force_tac 1);
+by(Clarify_tac 1);
+by(split_tac [split_if] 1);
+br conjI 1;
+ by(Blast_tac 1);
+by(Clarify_tac 1);
+by(EVERY[etac allE 1, etac impE 1, etac impE 2, etac trans 3]);
+  by(Blast_tac 1);
+ be lemma 1;
+ by(Blast_tac 1);
+br iffI 1;
+ by(blast_tac (claset() addIs [order_trans]) 1);
+by(Clarify_tac 1);
+by(EVERY[rtac bexI 1, atac 2, rtac conjI 1, atac 2]);
+by(thin_tac "!a:L. !b. next a = Some b --> a <= b" 1);
+by(Force_tac 1);
+qed_spec_mp "fix_iff_has_fixpoint";
+
+
+(* This lemma looks more pleasing to the eye, because of the monotonicity
+assumption for next, instead of the strange assumption above.
+However, function next as defined in DFAimpl is NOT monotone because
+None is not required to be detected as early as possible. Thus the following
+does not hold: sos <= tos & next sos = None ==> next tos = None
+
+Goal
+"[| wf{(y,x::'a::order). x : L & next x = Some y & y ~= x}; \
+\   !t:L. !s:L. s <= t & next s = None --> next t = None; \
+\   !t:L. !s:L. !u. s <= t & next s = Some u --> \
+\                   next t = None | (? v. next t = Some v & u <= v); \
+\   !a:L. next a : option L; !a:L. !b. next a = Some b --> a <= b |] \
+\ ==> s:L --> fix(next,s) = (? t:L. s <= t & next t = Some t)";
+by(res_inst_tac [("a","s")] wf_induct 1);
+ ba 1;
+by(Clarify_tac 1);
+by(Full_simp_tac 1);
+by(stac fix_unfold 1);
+   ba 1;
+  ba 1;
+ ba 1;
+by(split_tac [option.split] 1);
+br conjI 1;
+ by(blast_tac (claset() addDs [sym RS trans]) 1);
+by(Clarify_tac 1);
+by(split_tac [split_if] 1);
+br conjI 1;
+ by(Blast_tac 1);
+by(Clarify_tac 1);
+by(EVERY[etac allE 1, etac impE 1, etac impE 2, etac trans 3]);
+  by(Blast_tac 1);
+ be lemma 1;
+ by(Blast_tac 1);
+br iffI 1;
+ by(subgoal_tac "next a ~= None" 1);
+  by(Clarify_tac 1);
+  by(EVERY[etac ballE 1, etac ballE 1, etac allE 1, etac impE 1,etac disjE 2]);
+      br conjI 1;
+       ba 1;
+      ba 1;
+     by(blast_tac (claset() addDs [sym RS trans]) 1);
+    by(blast_tac (claset() addIs [order_trans]) 1);
+   by(blast_tac (claset() addIs [order_trans]) 1);
+  by(Blast_tac 1);
+ br notI 1;
+ by(Clarify_tac 1);
+ by(blast_tac (claset() addDs [sym RS trans,lemma]) 1);
+by(blast_tac (claset() addDs [sym RS trans]) 1);
+qed_spec_mp "fix_iff_has_fixpoint";
+*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Fixpoint.thy	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,23 @@
+(*  Title:      HOL/BCV/Fixpoint.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+Searching for a fixpoint.
+*)
+
+Fixpoint = SemiLattice +
+
+consts fix :: "('a => 'a option) * 'a => bool"
+recdef fix
+ "{((nxt,t),(nxs,s)) . nxt=nxs &
+      ~(? f. f 0 = s & (!i. nxt (f i) = Some(f(i+1)) & f(i+1) ~= f i)) &
+      nxs s = Some t & t ~= s}"
+"fix(next,s) =
+  (if (? f. f 0 = s & (!i. next (f i) = Some(f(i+1)) & f(i+1) ~= f i))
+   then arbitrary
+   else case next s of None => False |
+                       Some t => if t=s then True else fix(next,t))"
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Machine.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,117 @@
+(*  Title:      HOL/BCV/Machine.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+*)
+
+Addsimps [Let_def];
+
+Addsimps [stype_def];
+
+
+(*** Machine specific ***)
+
+Goal "!p. p<size(bs) --> maxreg(bs!p) < maxregs(bs)";
+by(induct_tac "bs" 1);
+by(auto_tac (claset(), simpset() addsimps [nth_Cons,less_max_iff_disj]
+                                 addsplits [nat.split]));
+qed_spec_mp "maxreg_less_maxregs";
+
+Goalw [le_typ] "(t <= Ctyp) = (t = Ctyp)";
+by(Simp_tac 1);
+qed "le_Ctyp_conv";
+
+Goalw [le_typ] "(t ~= Top) = (t = Atyp | t = Btyp | t = Ctyp)";
+by(induct_tac "t" 1);
+by(ALLGOALS Simp_tac);
+qed "le_Top_conv";
+
+Goalw [step_pres_type_def,listsn_def,exec_def,stype_def]
+ "step_pres_type (%p. exec (bs!p)) (length bs) (stype bs)";
+by(force_tac (claset(), simpset() addsplits [instr.split_asm,split_if_asm]) 1);
+qed "exec_pres_type";
+
+Goalw [wti_is_fix_step_def,stable_def,wt_instr_def,exec_def,succs_def]
+ "wti_is_fix_step (%p. exec (bs!p)) (%u. wt_instr (bs ! u) u) \
+\     (%p. succs (bs ! p) p) (length bs) (stype bs)";
+by(force_tac (claset() addDs [maxreg_less_maxregs,
+                              le_Top_conv RS iffD1,le_Ctyp_conv RS iffD1],
+              simpset() addsplits [option.split,instr.split]) 1);
+qed "wt_instr_is_fix_exec";
+
+
+Goalw [step_mono_None_def,exec_def,succs_def,le_list]
+ "step_mono_None (%p. exec (bs!p)) (length bs) (stype bs)";
+by(Clarify_tac 1);
+by(forward_tac [maxreg_less_maxregs] 1);
+by(split_asm_tac [instr.split_asm] 1);
+
+by(ALLGOALS (asm_full_simp_tac (simpset() addsplits [split_if_asm])));
+
+by(rotate_tac 1 1);
+by(Asm_full_simp_tac 1);
+by(blast_tac (claset() addIs [order_trans]) 1);
+
+by(rotate_tac 1 1);
+by(Asm_full_simp_tac 1);
+by(blast_tac (claset() addIs [order_trans]) 1);
+
+by(rotate_tac 1 1);
+by(Asm_full_simp_tac 1);
+by(blast_tac (claset() addIs [order_trans]) 1);
+
+by(rotate_tac 1 1);
+by(Asm_full_simp_tac 1);
+by(subgoal_tac "s!nat1 <= t!nat1" 1);
+by(Blast_tac 2);
+by(subgoal_tac "s!nat2 <= t!nat2" 1);
+by(Blast_tac 2);
+be thin_rl 1;
+by(asm_full_simp_tac (simpset() addsimps [le_typ])1);
+by(exhaust_tac "t!nat1" 1);
+by(ALLGOALS Asm_full_simp_tac);
+by(exhaust_tac "t!nat2" 1);
+by(ALLGOALS Asm_full_simp_tac);
+by(exhaust_tac "t!nat2" 1);
+by(ALLGOALS Asm_full_simp_tac);
+by(exhaust_tac "s!nat1" 1);
+by(ALLGOALS Asm_full_simp_tac);
+by(exhaust_tac "t!nat2" 1);
+by(ALLGOALS Asm_full_simp_tac);
+
+qed "exec_mono_None";
+
+Goalw [step_mono_def,exec_def,succs_def]
+ "step_mono (%p. exec (bs!p)) (length bs) (stype bs)";
+by(Clarify_tac 1);
+by(forward_tac [maxreg_less_maxregs] 1);
+by(split_asm_tac [instr.split_asm] 1);
+by(ALLGOALS
+   (fast_tac (claset() addIs [list_update_le_cong,le_listD]
+               addss (simpset() addsplits [split_if_asm]))));
+
+qed_spec_mp "exec_mono_Some";
+
+Goalw [stype_def] "semilat(stype bs)";
+by(Blast_tac 1);
+qed "lat_stype";
+
+Goalw [stype_def] "acc(stype bs)";
+by(Simp_tac 1);
+qed "acc_stype";
+
+Delsimps [stype_def];
+
+Goal
+ "[| is_next next; \
+\    succs_bounded (%p. succs (bs!p) p) (size bs); \
+\    sos : listsn (size bs) (option(stype bs)) |] ==> \
+\ fix(next (%p. exec (bs!p)) (%p. succs (bs!p) p), sos) = \
+\ (? tos : listsn (size bs) (option(stype bs)). \
+\      sos <= tos & welltyping (%p. wt_instr (bs!p) p) tos)";
+by(simp_tac (simpset() delsimps [not_None_eq]) 1);
+by(REPEAT(ares_tac [fix_next_iff_welltyping,exec_pres_type,
+                    exec_mono_None,exec_mono_Some,
+                    wt_instr_is_fix_exec,lat_stype,acc_stype] 1));
+
+qed "fix_iff_welltyped";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Machine.thy	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,96 @@
+(*  Title:      HOL/BCV/Machine.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+A register machine.
+*)
+
+Machine = Types + DFAimpl +
+
+datatype ('a,'b,'c)instr =
+    Move nat nat
+  | Aload 'a nat | Bload 'b nat | Cload 'c nat
+  | Aop nat nat nat | Bop nat nat nat | Cop nat nat nat
+  | Comp nat nat nat
+  | Halt
+
+consts maxreg :: ('a,'b,'c)instr => nat
+       maxregs :: ('a,'b,'c)instr list => nat
+primrec
+"maxreg(Move m n) = max m n"
+"maxreg(Aload x n) = n"
+"maxreg(Bload x n) = n"
+"maxreg(Cload x n) = n"
+"maxreg(Aop i j k) = max i (max j k)"
+"maxreg(Bop i j k) = max i (max j k)"
+"maxreg(Cop i j k) = max i (max j k)"
+"maxreg(Comp i j l) = max i j"
+"maxreg Halt = 0"
+
+primrec
+"maxregs [] = 0"
+"maxregs (x#xs) = max (maxreg x+1) (maxregs xs)"
+
+constdefs
+
+ wt_instr :: ('a,'b,'c)instr => nat => typ list option list => bool
+"wt_instr b p T ==
+ case T!p of
+   None => True
+ | Some ts =>
+ case b of
+   Move i j  => Some(ts[j:=ts!i]) <= T!(p+1)
+ | Aload x i => Some(ts[i:=Atyp]) <= T!(p+1)
+ | Bload x i => Some(ts[i:=Btyp]) <= T!(p+1)
+ | Cload x i => Some(ts[i:=Ctyp]) <= T!(p+1)
+ | Aop i j k => ts!i <= Atyp & ts!j <= Atyp & Some(ts[k:=Atyp]) <= T!(p+1)
+ | Bop i j k => ts!i <= Btyp & ts!j <= Btyp & Some(ts[k:=Btyp]) <= T!(p+1)
+ | Cop i j k => ts!i <= Ctyp & ts!j <= Ctyp & Some(ts[k:=Ctyp]) <= T!(p+1)
+ | Comp i j q => (ts!i <= ts!j & ts!j ~= Top | ts!j <= ts!i & ts!i ~= Top) &
+                 T!p <= T!(p+1) & T!p <= T!q
+ | Halt =>      True"
+
+ stype :: ('a,'b,'c)instr list => typ list set
+"stype bs == listsn (maxregs bs) UNIV"
+
+constdefs
+ succs :: "('a,'b,'c)instr => nat => nat list"
+"succs instr p ==
+  case instr of
+    Move i j  => [p+1]
+  | Aload x i => [p+1]
+  | Bload x i => [p+1]
+  | Cload x i => [p+1]
+  | Aop i j k => [p+1]
+  | Bop i j k => [p+1]
+  | Cop i j k => [p+1]
+  | Comp i j q => [p+1,q]
+  | Halt       => []"
+
+ exec :: "('a,'b,'c)instr => typ list => typ list option"
+"exec instr ts ==
+  case instr of
+    Move i j  => Some(ts[j := ts!i])
+  | Aload x i => Some(ts[i := Atyp])
+  | Bload x i => Some(ts[i := Btyp])
+  | Cload x i => Some(ts[i := Ctyp])
+  | Aop i j k => if ts!i <= Atyp & ts!j <= Atyp
+                 then Some(ts[k := Atyp])
+                 else None
+  | Bop i j k => if ts!i <= Btyp & ts!j <= Btyp
+                 then Some(ts[k := Btyp])
+                 else None
+  | Cop i j k => if ts!i <= Ctyp & ts!j <= Ctyp
+                 then Some(ts[k := Ctyp])
+                 else None
+  | Comp i j q => if ts!i <= ts!j & ts!j ~= Top | ts!j <= ts!i & ts!i ~= Top
+                  then Some ts
+                  else None
+  | Halt       => Some ts"
+(*
+  | Bop i j k => if ts!i <= Btyp & ts!j <= Btyp
+                 then Some(ts[k := ts!i+ts!j])
+                 else None
+*)
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Orders.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,177 @@
+(*  Title:      HOL/BCV/Orders.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+*)
+
+Delrules[le_maxI1, le_maxI2];
+
+(** option **)
+section "option";
+
+Goalw [option_def] "None : option A";
+by(Simp_tac 1);
+qed "None_in_option";
+AddIffs [None_in_option];
+
+Goalw [option_def] "(Some x : option A) = (x:A)";
+by(Auto_tac);
+qed "Some_in_option";
+AddIffs [Some_in_option];
+
+Goalw [less_option,le_option]
+ "Some x < opt = (? y. opt = Some y & x < (y::'a::order))";
+by(simp_tac (simpset() addsimps [order_less_le] addsplits [option.split]) 1);
+qed_spec_mp "Some_less_conv";
+AddIffs [Some_less_conv];
+
+Goalw [less_option,le_option] "None < x = (? a. x = Some a)";
+by(asm_simp_tac (simpset() addsplits [option.split]) 1);
+qed_spec_mp "None_less_iff";
+AddIffs [None_less_iff];
+
+
+Goalw [acc_def] "acc A ==> acc(option A)";
+by(asm_full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1);
+by(Clarify_tac 1);
+by(case_tac "? a. Some a : Q" 1);
+ by(eres_inst_tac [("x","{a . Some a : Q}")] allE 1);
+ by(Blast_tac 1);
+by(exhaust_tac "x" 1);
+ by(Fast_tac 1);
+by(Blast_tac 1);
+qed "acc_option";
+Addsimps [acc_option];
+AddSIs [acc_option];
+
+
+(** list **)
+section "list";
+
+Goalw [le_list] "~ [] <= x#xs";
+by(Simp_tac 1);
+qed "Nil_notle_Cons";
+
+Goalw [le_list] "~ x#xs <= []";
+by(Simp_tac 1);
+qed "Cons_notle_Nil";
+
+AddIffs [Nil_notle_Cons,Cons_notle_Nil];
+
+Goalw [le_list] "x#xs <= y#ys = (x <= y & xs <= ys)";
+br iffI 1;
+ by(Asm_full_simp_tac 1);
+ by(Clarify_tac 1);
+ br conjI 1;
+  by(eres_inst_tac [("x","0")] allE 1);
+  by(Asm_full_simp_tac 1);
+ by(Clarify_tac 1);
+ by(eres_inst_tac [("x","Suc i")] allE 1);
+ by(Asm_full_simp_tac 1);
+by(Asm_full_simp_tac 1);
+by(Clarify_tac 1);
+by(exhaust_tac "i" 1);
+by(ALLGOALS Asm_simp_tac);
+qed "Cons_le_Cons";
+AddIffs [Cons_le_Cons];
+
+Goal "(x#xs <= ys) = (? z zs. ys = z#zs & x <= z & xs <= zs)";
+by(exhaust_tac "ys" 1);
+by(ALLGOALS Asm_simp_tac);
+qed "Cons_le_iff";
+
+Goal "(xs <= y#ys) = (? z zs. xs = z#zs & z <= y & zs <= ys)";
+by(exhaust_tac "xs" 1);
+by(ALLGOALS Asm_simp_tac);
+qed "le_Cons_iff";
+
+Goalw [less_list]
+ "x#xs < y#ys = (x < (y::'a::order) & xs <= ys  |  x = y & xs < ys)";
+by(simp_tac (simpset() addsimps [order_less_le]) 1);
+by(Blast_tac 1);
+qed "Cons_less_Cons";
+AddIffs [Cons_less_Cons];
+
+Goalw [le_list]
+ "[| i<size xs; xs <= ys; x <= y |] ==> xs[i:=x] <= ys[i:=y]";
+by(asm_full_simp_tac (simpset() addsimps [nth_list_update]) 1);
+qed "list_update_le_cong";
+
+Goal "!i < length xs. (xs[i := x] <= xs) = ((x::'a::order) <= xs!i)";
+by(induct_tac "xs" 1);
+ by(Simp_tac 1);
+by(asm_simp_tac (simpset() addsplits [nat.split]) 1);
+qed_spec_mp "list_update_le_conv";
+
+Goalw [listsn_def] "xs : listsn n A ==> length xs = n";
+by(Blast_tac 1);
+qed "listsnE_length";
+Addsimps [listsnE_length];
+
+Goalw [listsn_def] "xs : listsn n A ==> set xs <= A";
+by(Blast_tac 1);
+qed "listsnE_set";
+Addsimps [listsnE_set];
+
+Goalw [listsn_def] "listsn 0 A = {[]}";
+by(Auto_tac); 
+qed "listsn_0";
+Addsimps [listsn_0];
+
+Goalw [listsn_def]
+ "(xs : listsn (Suc n) A) = (? y:A. ? ys:listsn n A. xs = y#ys)";
+by(exhaust_tac "xs" 1);
+by(Auto_tac);
+qed "in_listsn_Suc_iff";
+
+
+Goal "? a. a:A ==> ? xs. xs : listsn n A";
+by(induct_tac "n" 1);
+ by(Simp_tac 1);
+by(asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff]) 1);
+by(Blast_tac 1);
+qed "listsn_not_empty";
+
+
+Goal "!i n. length xs = n --> set xs <= A --> i < n --> (xs!i) : A";
+by(induct_tac "xs" 1);
+ by(Simp_tac 1);
+by(asm_full_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
+qed_spec_mp "nth_in";
+Addsimps [nth_in];
+
+Goal "[| xs : listsn n A; i < n |] ==> (xs!i) : A";
+by(blast_tac (HOL_cs addIs [nth_in,listsnE_length,listsnE_set]) 1);
+qed "listsnE_nth_in";
+
+Goalw [listsn_def]
+ "[| xs : listsn n A; x:A; i < n |] ==> xs[i := x] : listsn n A";
+by(Asm_full_simp_tac 1);
+by(blast_tac (claset() addDs [set_list_update_subset RS subsetD]) 1);
+qed "list_update_in_listsn";
+Addsimps [list_update_in_listsn];
+AddSIs [list_update_in_listsn];
+
+
+Goalw [acc_def] "acc(A) ==> acc(listsn n A)";
+by(asm_full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1);
+by(induct_tac "n" 1);
+ by(simp_tac (simpset() addcongs [conj_cong]) 1);
+by(simp_tac (simpset() addsimps [in_listsn_Suc_iff]) 1);
+by(Clarify_tac 1);
+by(rename_tac "M m" 1);
+by(case_tac "? x xs. x:A & xs:listsn n A & x#xs : M" 1);
+ be thin_rl 2;
+ be thin_rl 2;
+ by(Blast_tac 2);
+by(eres_inst_tac [("x","{a. a:A & (? xs. xs:listsn n A & a#xs:M)}")] allE 1);
+be impE 1;
+ by(Blast_tac 1);
+by(thin_tac "? x xs. ?P x xs" 1);
+by(Clarify_tac 1);
+by(rename_tac "maxA xs" 1);
+by(eres_inst_tac [("x","{xs. xs:listsn n A & maxA#xs : M}")] allE 1);
+by(Blast_tac 1);
+qed "acc_listsn";
+Addsimps [acc_listsn];
+AddSIs [acc_listsn];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Orders.thy	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,28 @@
+(*  Title:      HOL/BCV/Orders.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+Orderings and some sets.
+*)
+
+Orders = Orders0 +
+
+instance option :: (order)order (le_option_refl,le_option_trans,
+                                 le_option_antisym,less_le_option)
+instance list :: (order)order (le_list_refl,le_list_trans,
+                               le_list_antisym,less_le_list)
+instance "*" :: (order,order)order
+                (le_prod_refl,le_prod_trans,le_prod_antisym,less_le_prod)
+
+constdefs
+ acc :: "'a::order set => bool"
+"acc A == wf{(y,x) . x:A & y:A & x < y}"
+
+ option :: 'a set => 'a option set
+"option A == insert None {x . ? y:A. x = Some y}"
+
+ listsn :: nat => 'a set => 'a list set
+"listsn n A == {xs. length xs = n & set xs <= A}"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Orders0.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,129 @@
+(*  Title:      HOL/BCV/Orders0.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+*)
+
+(** option **)
+section "option";
+
+Goalw [le_option] "(o1::('a::order)option) <= o1";
+by(simp_tac (simpset() addsplits [option.split]) 1);
+qed "le_option_refl";
+
+Goalw [le_option] "(o1::('a::order)option) <= o2 --> o2<=o3 --> o1<=o3";
+by(simp_tac (simpset() addsplits [option.split]) 1);
+by(blast_tac (claset() addIs [order_trans]) 1);
+qed_spec_mp "le_option_trans";
+
+Goalw [le_option] "(o1::('a::order)option) <= o2 --> o2<=o1 --> o1=o2";
+by(simp_tac (simpset() addsplits [option.split]) 1);
+by(blast_tac (claset() addIs [order_antisym]) 1);
+qed_spec_mp "le_option_antisym";
+
+Goalw [less_option] "(o1::('a::ord)option) < o2 = (o1<=o2 & o1 ~= o2)";
+br refl 1;
+qed "less_le_option";
+
+Goalw [le_option] "Some x <= opt = (? y. opt = Some y & x <= y)";
+by(simp_tac (simpset() addsplits [option.split]) 1);
+qed_spec_mp "Some_le_conv";
+AddIffs [Some_le_conv];
+
+Goalw [le_option] "None <= opt";
+by(simp_tac (simpset() addsplits [option.split]) 1);
+qed "None_le";
+AddIffs [None_le];
+
+
+(** list **)
+section "list";
+
+Goalw [le_list] "[| xs <= ys; p < size xs |] ==> xs!p <= ys!p";
+by(Blast_tac 1);
+qed "le_listD";
+
+Goalw [le_list] "([] <= ys) = (ys = [])";
+by(Simp_tac 1);
+qed "Nil_le_conv";
+AddIffs [Nil_le_conv];
+
+Goalw [le_list] "(xs::('a::order)list) <= xs";
+by(induct_tac "xs" 1);
+by(Auto_tac);
+qed "le_list_refl";
+
+Goalw [le_list]
+ "!ys zs.(xs::('a::order)list) <= ys --> ys <= zs --> xs <= zs";
+by(induct_tac "xs" 1);
+ by(Simp_tac 1);
+br allI 1;
+by(exhaust_tac "ys" 1);
+ by(hyp_subst_tac 1);
+ by(Simp_tac 1);
+br allI 1;
+by(exhaust_tac "zs" 1);
+ by(hyp_subst_tac 1);
+ by(Simp_tac 1);
+by(hyp_subst_tac 1);
+by(simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
+by(Clarify_tac 1);
+br conjI 1;
+ by(REPEAT(EVERY1[etac allE, etac conjE, etac impE, rtac refl]));
+ by(blast_tac (claset() addIs [order_trans]) 1);
+by(Clarify_tac 1);
+by(EVERY[etac allE 1, etac allE 1, etac impE 1, etac impE 2]);
+  br conjI 1;
+  ba 1;
+  by(Blast_tac 1);
+ br conjI 1;
+ ba 1;
+ by(Blast_tac 1);
+by(Asm_full_simp_tac 1);
+qed_spec_mp "le_list_trans";
+
+Goalw [le_list]
+ "!ys. (xs::('a::order)list) <= ys --> ys <= xs --> xs = ys";
+by(induct_tac "xs" 1);
+ by(Simp_tac 1);
+br allI 1;
+by(exhaust_tac "ys" 1);
+ by(hyp_subst_tac 1);
+ by(Simp_tac 1);
+by(hyp_subst_tac 1);
+by(simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
+by(Clarify_tac 1);
+br conjI 1;
+ by(blast_tac (claset() addIs [order_antisym]) 1);
+by(Asm_full_simp_tac 1);
+qed_spec_mp "le_list_antisym";
+
+Goalw [less_list] "(xs::('a::ord)list) < ys = (xs<=ys & xs ~= ys)";
+br refl 1;
+qed "less_le_list";
+
+(** product **)
+section "product";
+
+Goalw [le_prod] "(p1::('a::order * 'b::order)) <= p1";
+by(Simp_tac 1);
+qed "le_prod_refl";
+
+Goalw [le_prod]
+ "[| (p1::('a::order * 'b::order)) <= p2; p2<=p3 |] ==> p1<=p3";
+by(blast_tac (claset() addIs [order_trans]) 1);
+qed "le_prod_trans";
+
+Goalw [le_prod]
+ "[| (p1::('a::order * 'b::order)) <= p2; p2 <= p1 |] ==> p1 = p2";
+by(blast_tac (claset() addIs [order_antisym,trans,surjective_pairing,sym]) 1);
+qed_spec_mp "le_prod_antisym";
+
+Goalw [less_prod] "(p1::('a::order * 'b::order)) < p2 = (p1<=p2 & p1 ~= p2)";
+br refl 1;
+qed "less_le_prod";
+
+Goalw [le_prod] "((a,b) <= (c,d)) = (a <= c & b <= d)";
+by(Simp_tac 1);
+qed "le_prod_iff";
+AddIffs [le_prod_iff];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Orders0.thy	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,26 @@
+(*  Title:      HOL/BCV/Orders0.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+Orderings on various types.
+*)
+
+Orders0 = Main +
+
+instance option :: (ord)ord
+instance list   :: (ord)ord
+instance "*" :: (ord,ord)ord
+
+defs
+le_option "o1 <= o2 == case o2 of None => o1=None |
+                         Some y => (case o1 of None => True | Some x => x<=y)"
+less_option "(o1::('a::ord)option) < o2 == o1 <= o2 & o1 ~= o2"
+
+le_list "xs <= ys == size xs = size ys & (!i. i<size xs --> xs!i <= ys!i)"
+less_list "(xs::('a::ord)list) < ys == xs <= ys & xs ~= ys"
+
+le_prod "p1 <= p2 == fst p1 <= fst p2 & snd p1 <= snd p2"
+less_prod "(p1::('a::ord * 'b::ord)) < p2 == p1 <= p2 & p1 ~= p2"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Plus.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,48 @@
+(*  Title:      HOL/BCV/Plus.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+*)
+
+(** option **)
+
+Goalw [plus_option] "x+None = x";
+by(simp_tac (simpset() addsplits [option.split]) 1);
+qed "plus_None";
+Addsimps [plus_None];
+
+Goalw [plus_option] "None+x = x";
+by(simp_tac (simpset() addsplits [option.split]) 1);
+qed "None_plus";
+Addsimps [None_plus];
+
+Goalw [plus_option] "Some x + Some y = Some(x+y)";
+by(Simp_tac 1);
+qed "Some_plus_Some";
+Addsimps [Some_plus_Some];
+
+Goalw [plus_option] "? y. Some x + opt = Some y";
+by(simp_tac (simpset() addsplits [option.split]) 1);
+qed "plus_option_Some_Some";
+
+(** list **)
+
+Goal "list_plus xs [] = xs";
+by(induct_tac "xs" 1);
+ by(Simp_tac 1);
+by(asm_simp_tac (simpset() addsplits [list.split]) 1);
+qed "list_plus_Nil2";
+Addsimps [list_plus_Nil2];
+
+Goal "!ys. length(list_plus xs ys) = max(length xs) (length ys)";
+by(induct_tac "xs" 1);
+ by(Simp_tac 1);
+by(asm_simp_tac (simpset() addsplits [list.split]) 1);
+qed_spec_mp "length_list_plus";
+Addsimps [length_list_plus];
+
+Goalw [plus_list]
+ "length(ts+us) = max (length ts) (length us)";
+by(Simp_tac 1);
+qed "length_plus_list";
+Addsimps [length_plus_list];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Plus.thy	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,26 @@
+(*  Title:      HOL/BCV/Plus.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+Supremum operation on various domains.
+*)
+
+Plus = Orders +
+
+instance option :: (plus)plus
+instance list   :: (plus)plus
+instance "*"    :: (plus,plus)plus
+
+consts list_plus :: "('a::plus)list => 'a list => 'a list"
+primrec
+"list_plus [] ys = ys"
+"list_plus (x#xs) ys = (case ys of [] => x#xs | z#zs => (x+z)#list_plus xs zs)"
+defs
+plus_option "o1 + o2 == case o1 of None => o2
+                        | Some x => (case o2 of None => o1
+                                     | Some y => Some(x+y))"
+plus_list "op + == list_plus"
+plus_prod "op + == %(a,b) (c,d). (a+c,b+d)"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/README.html	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,16 @@
+<HTML><HEAD><TITLE>HOL/BCV/README</TITLE></HEAD>
+<BODY>
+
+<H1>Towards Verified Bytecode Verifiers</H1>
+
+This theory defines a very abstract and generic model of bytecode
+verification, i.e. data-flow analysis for assembly languages with subtypes.
+
+<P>
+
+A paper describing the theory is found here:
+<A HREF = "http://www4.in.tum.de/~nipkow/pubs/bcv.html">
+Towards Verified Bytecode Verifiers</A>.
+
+</BODY>
+</HTML>
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/ROOT.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,12 @@
+(*  Title:      HOL/BCV/ROOT.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+A simple model of dataflow (sub)type analysis of instruction sequences,
+aka `bytecode verification'.
+*)
+
+writeln"Root file for HOL/BCV";
+
+time_use_thy "Machine";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/SemiLattice.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,155 @@
+(*  Title:      HOL/BCV/SemiLattice.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+*)
+
+Goalw [semilat_def] "[| x:A; y:A; semilat A |] ==> x <= x+y";
+by(Blast_tac 1);
+qed "semilat_ub1";
+
+Goalw [semilat_def] "[| x:A; y:A; semilat A |] ==> y <= x+y";
+by(Blast_tac 1);
+qed "semilat_ub2";
+
+Goalw [semilat_def] "[| x:A; y:A; z:A; x <= z; y <= z; semilat A |] ==> x+y <= z";
+by(Blast_tac 1);
+qed "semilat_lub";
+
+Goalw [semilat_def] "[| x:A; y:A; semilat A |] ==> x+y : A";
+by(Blast_tac 1);
+qed "semilat_plus";
+
+Addsimps [semilat_ub1,semilat_ub2,semilat_lub,semilat_plus];
+
+Goal "[| x:A; y:A; semilat A |] ==> (x <= y) = (x+y = y)";
+br iffI 1;
+ by(REPEAT(ares_tac [order_antisym,semilat_lub,order_refl,semilat_ub2] 1));
+be subst 1;
+by(Asm_simp_tac 1);
+qed "le_iff_plus_unchanged";
+
+Goal "[| x:A; y:A; semilat A |] ==> (x <= y) = (y = x+y)";
+br iffI 1;
+ by(REPEAT(ares_tac [order_antisym,semilat_lub,order_refl,semilat_ub2] 1));
+be ssubst 1;
+by(Asm_simp_tac 1);
+qed "le_iff_plus_unchanged2";
+
+Goal "[| x:A; y:A; z:A; semilat A; x <= y | x <= z |] ==> x <= y+z";
+by(blast_tac (claset() addIs [order_trans,semilat_ub1,semilat_ub2]) 1);
+qed "plus_mono";
+
+Goal "[| x:A; semilat A |] ==> x+x = x";
+by(REPEAT(ares_tac [order_refl RSN (4,le_iff_plus_unchanged RS iffD1)] 1));
+qed "semilat_idemp";
+Addsimps [semilat_idemp];
+
+Goal "[| x:A; y:A; z:A; semilat A |] ==> (x+y)+z = x+(y+z)";
+by(blast_tac (HOL_cs addSIs [order_refl,order_antisym,semilat_lub,semilat_plus]
+                     addIs [plus_mono]) 1);
+qed "semilat_assoc";
+
+Goal "[| x:A; y:A; semilat A |] ==> x+(x+y) = x+y";
+by(asm_simp_tac (simpset() addsimps [semilat_assoc RS sym]) 1);
+qed "semilat_assoc_idemp";
+Addsimps [semilat_assoc_idemp];
+
+Goal "[| x:A; y:A; z:A; semilat A |] ==> (x+y <= z) = (x <= z & y <= z)";
+by(blast_tac (claset() addIs [semilat_ub1,semilat_ub2,semilat_lub,order_trans]) 1);
+qed "plus_le_conv";
+Addsimps [plus_le_conv];
+
+
+(** option **)
+
+Goal "semilat A ==> semilat (option A)";
+by(simp_tac (simpset() addsimps [semilat_def,le_option,plus_option]
+                       addsplits [option.split]) 1);
+by(blast_tac (claset() addIs [semilat_ub1,semilat_ub2,semilat_lub,semilat_plus]) 1);
+qed "semilat_option";
+Addsimps [semilat_option];
+AddSIs [semilat_option];
+
+(** list **)
+Goalw [plus_list] "[] + [] = []";
+by(Simp_tac 1);
+qed "plus_Nil_Nil";
+Addsimps [plus_Nil_Nil];
+
+Goalw [plus_list] "(x#xs) + (y#ys) = (x+y)#(xs+ys)";
+by(Simp_tac 1);
+qed "plus_Cons_Cons";
+Addsimps [plus_Cons_Cons];
+
+Goal
+ "!xs ys i. length xs = n--> length ys = n--> i<n--> (xs+ys)!i = xs!i + ys!i";
+by(induct_tac "n" 1);
+ by(Simp_tac 1);
+by(Clarify_tac 1);
+by(exhaust_tac "xs" 1);
+ by(Asm_full_simp_tac 1);
+by(exhaust_tac "ys" 1);
+ by(Asm_full_simp_tac 1);
+by(asm_full_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
+qed_spec_mp "nth_plus";
+Addsimps [nth_plus];
+
+Goalw [le_list]
+ "semilat A ==> !xs:listsn n A. !ys:listsn n A. xs <= xs+ys";
+by(asm_simp_tac (simpset() addsimps [le_max_iff_disj]) 1);
+by(Clarify_tac 1);
+by(blast_tac (HOL_cs addIs [nth_in,nth_plus,listsnE_length,listsnE_set,semilat_ub1]) 1);
+qed_spec_mp "plus_list_ub1";
+
+Goalw [le_list]
+ "semilat A ==> !xs:listsn n A. !ys:listsn n A. ys <= xs+ys";
+by(asm_simp_tac (simpset() addsimps [le_max_iff_disj]) 1);
+by(Clarify_tac 1);
+by(blast_tac (HOL_cs addIs [nth_in,nth_plus,listsnE_length,listsnE_set,semilat_ub2]) 1);
+qed_spec_mp "plus_list_ub2";
+
+Goalw [le_list]
+ "semilat A ==> !xs:listsn n A. !ys:listsn n A. !zs:listsn n A. \
+\              xs <= zs & ys <= zs --> xs+ys <= zs";
+by(asm_simp_tac (simpset() addsimps [le_max_iff_disj] addcongs [conj_cong]) 1);
+by(Clarify_tac 1);
+by(blast_tac (HOL_cs addIs [nth_in,nth_plus,listsnE_length,listsnE_set,semilat_lub]) 1);
+qed_spec_mp "plus_list_lub";
+
+Goalw [listsn_def]
+ "semilat A ==> !xs:listsn n A. !ys:listsn n A. xs+ys : listsn n A";
+by(induct_tac "n" 1);
+ by(Simp_tac 1);
+by(Asm_full_simp_tac 1);
+by(Clarify_tac 1);
+by(exhaust_tac "xs" 1);
+ by(Asm_full_simp_tac 1);
+by(exhaust_tac "ys" 1);
+ by(Asm_full_simp_tac 1);
+by(Asm_full_simp_tac 1);
+by(blast_tac (claset() addIs [semilat_plus]) 1);
+qed_spec_mp "plus_list_closed";
+
+Goal "semilat A ==> semilat (listsn n A)";
+by(simp_tac (HOL_basic_ss addsimps [semilat_def]) 1);
+by(asm_simp_tac (HOL_ss addsimps [plus_list_ub1,plus_list_ub2,plus_list_lub,plus_list_closed]) 1);
+qed "semilat_listsn";
+Addsimps [semilat_listsn];
+AddSIs [semilat_listsn];
+
+Goalw [le_list]
+ "!i xs. xs : listsn n A --> x:A --> semilat A --> i<n --> xs <= xs[i := x + xs!i]";
+by(Simp_tac 1);
+by(induct_tac "n" 1);
+ by(Simp_tac 1);
+by(asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff]) 1);
+by(Clarify_tac 1);
+by(asm_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
+qed_spec_mp "list_update_incr";
+
+(* product *)
+
+Goalw [semilat_def,plus_prod] "[| semilat A; semilat B |] ==> semilat (A Times B)";
+by(Asm_simp_tac 1);
+qed "semilat_Times";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/SemiLattice.thy	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,18 @@
+(*  Title:      HOL/BCV/SemiLattice.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+Semilattices
+*)
+
+SemiLattice = Plus +
+
+constdefs
+ semilat :: "('a::{order,plus} set) => bool"
+"semilat A == (!x:A. !y:A. x <= x+y)  &
+          (!x:A. !y:A. y <= x+y)  &
+          (!x:A. !y:A. !z:A. x <= z & y <= z --> x+y <= z) &
+          (!x:A. !y:A. x+y : A)"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Types.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,31 @@
+(*  Title:      HOL/BCV/Types.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+*)
+
+Goalw [semilat_def,le_typ,plus_typ] "semilat (UNIV::typ set)";
+by(Auto_tac);
+qed "semilat_typ";
+AddIffs [semilat_typ];
+
+Goal "{(x,y::'a::order). y<x}^+ = {(x,y::'a::order). y<x}";
+by(Auto_tac);
+ be trancl_induct 1;
+ by(Blast_tac 1);
+ by(blast_tac (claset() addIs [order_less_trans]) 1);
+by(blast_tac (claset() addIs [r_into_trancl]) 1);
+qed "trancl_order1_conv";
+Addsimps [trancl_order1_conv];
+
+Goalw [acyclic_def] "acyclic{(x,y::'a::order). y<x}";
+by(Simp_tac 1);
+qed "acyclic_order1";
+AddIffs [acyclic_order1];
+
+Goalw [acc_def] "acc(UNIV::typ set)";
+br finite_acyclic_wf 1;
+ by(fast_tac (claset() addIs [finite_SigmaI RSN (2,finite_subset)]) 1);
+by(blast_tac (claset() addIs [acyclic_subset]) 1);
+qed "acc_UNIV_typ";
+AddIffs [acc_UNIV_typ];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Types.thy	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,13 @@
+(*  Title:      HOL/BCV/Types.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+Types for the register machine
+*)
+
+Types = Types0 +
+
+instance typ :: order (le_typ_refl,le_typ_trans,le_typ_antisym,less_le_typ)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Types0.ML	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,33 @@
+(*  Title:      HOL/BCV/Types0.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+*)
+
+Goalw [le_typ] "(t::typ) <= t";
+by(Blast_tac 1);
+qed "le_typ_refl";
+
+Goalw [le_typ] "[| (t1::typ) <= t2; t2<=t3 |] ==> t1<=t3";
+by(Blast_tac 1);
+qed "le_typ_trans";
+
+Goalw [le_typ] "[| (t1::typ) <= t2; t2<=t1 |] ==> t1=t2";
+by(Blast_tac 1);
+qed "le_typ_antisym";
+
+Goalw [less_typ] "(t1::typ) < t2 = (t1<=t2 & t1 ~= t2)";
+br refl 1;
+qed "less_le_typ";
+
+Goal "(UNIV::typ set) = {Atyp,Btyp,Ctyp,Top}";
+by(Auto_tac);
+by(rename_tac "t" 1);
+by(exhaust_tac "t" 1);
+by(Auto_tac);
+qed "UNIV_typ_enum";
+
+Goal "finite(UNIV::typ set)";
+by(simp_tac (simpset() addsimps [UNIV_typ_enum]) 1);
+qed "finite_UNIV_typ";
+AddIffs [finite_UNIV_typ];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BCV/Types0.thy	Tue Sep 28 16:36:12 1999 +0200
@@ -0,0 +1,21 @@
+(*  Title:      HOL/BCV/Types0.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1999 TUM
+
+Types for the register machine
+*)
+
+Types0 = SemiLattice +
+
+datatype typ = Atyp | Btyp | Ctyp | Top
+
+instance typ :: ord
+instance typ :: plus
+
+defs
+le_typ "t1 <= t2 == (t1=t2) | (t1=Atyp & t2=Btyp) | (t2=Top)"
+less_typ "(t1::typ) < t2 == t1 <= t2 & t1 ~= t2"
+plus_typ "t1 + t2 == if t1<=t2 then t2 else if t2 <= t1 then t1 else Top"
+
+end