| author | desharna | 
| Mon, 10 Jun 2024 08:34:09 +0200 | |
| changeset 80317 | be2e772e0adf | 
| parent 67613 | ce654b0e6d69 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 42150 | 1 | (* Title: HOL/MicroJava/DFA/Semilat.thy | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 2 | Author: Tobias Nipkow | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 3 | Copyright 2000 TUM | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 4 | *) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 5 | |
| 61361 | 6 | chapter \<open>Bytecode Verifier \label{cha:bv}\<close>
 | 
| 58886 | 7 | |
| 61361 | 8 | section \<open>Semilattices\<close> | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 9 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 10 | theory Semilat | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
63258diff
changeset | 11 | imports Main "HOL-Library.While_Combinator" | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 12 | begin | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 13 | |
| 42463 | 14 | type_synonym 'a ord = "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 15 | type_synonym 'a binop = "'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 16 | type_synonym 'a sl = "'a set \<times> 'a ord \<times> 'a binop" | |
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 17 | |
| 62145 | 18 | definition lesub :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" | 
| 19 | where "lesub x r y \<longleftrightarrow> r x y" | |
| 20 | ||
| 21 | definition lesssub :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" | |
| 22 | where "lesssub x r y \<longleftrightarrow> lesub x r y \<and> x \<noteq> y" | |
| 23 | ||
| 24 | definition plussub :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c"
 | |
| 25 | where "plussub x f y = f x y" | |
| 26 | ||
| 61994 | 27 | notation (ASCII) | 
| 35251 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
35109diff
changeset | 28 |   "lesub"  ("(_ /<='__ _)" [50, 1000, 51] 50) and
 | 
| 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
35109diff
changeset | 29 |   "lesssub"  ("(_ /<'__ _)" [50, 1000, 51] 50) and
 | 
| 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
35109diff
changeset | 30 |   "plussub"  ("(_ /+'__ _)" [65, 1000, 66] 65)
 | 
| 62145 | 31 | |
| 61994 | 32 | notation | 
| 35251 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
35109diff
changeset | 33 |   "lesub"  ("(_ /\<sqsubseteq>\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and
 | 
| 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
35109diff
changeset | 34 |   "lesssub"  ("(_ /\<sqsubset>\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and
 | 
| 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
35109diff
changeset | 35 |   "plussub"  ("(_ /\<squnion>\<^bsub>_\<^esub> _)" [65, 0, 66] 65)
 | 
| 62145 | 36 | |
| 35355 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 37 | (* allow \<sub> instead of \<bsub>..\<esub> *) | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 38 | abbreviation (input) | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 39 |   lesub1 :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubseteq>\<^sub>_ _)" [50, 1000, 51] 50)
 | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 40 | where "x \<sqsubseteq>\<^sub>r y == x \<sqsubseteq>\<^bsub>r\<^esub> y" | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 41 | |
| 35355 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 42 | abbreviation (input) | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 43 |   lesssub1 :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubset>\<^sub>_ _)" [50, 1000, 51] 50)
 | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 44 | where "x \<sqsubset>\<^sub>r y == x \<sqsubset>\<^bsub>r\<^esub> y" | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 45 | |
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 46 | abbreviation (input) | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 47 |   plussub1 :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /\<squnion>\<^sub>_ _)" [65, 1000, 66] 65)
 | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35251diff
changeset | 48 | where "x \<squnion>\<^sub>f y == x \<squnion>\<^bsub>f\<^esub> y" | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 49 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35251diff
changeset | 50 | definition ord :: "('a \<times> 'a) set \<Rightarrow> 'a ord" where
 | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 51 | "ord r \<equiv> \<lambda>x y. (x,y) \<in> r" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 52 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35251diff
changeset | 53 | definition order :: "'a ord \<Rightarrow> bool" where | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 54 | "order r \<equiv> (\<forall>x. x \<sqsubseteq>\<^sub>r x) \<and> (\<forall>x y. x \<sqsubseteq>\<^sub>r y \<and> y \<sqsubseteq>\<^sub>r x \<longrightarrow> x=y) \<and> (\<forall>x y z. x \<sqsubseteq>\<^sub>r y \<and> y \<sqsubseteq>\<^sub>r z \<longrightarrow> x \<sqsubseteq>\<^sub>r z)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 55 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35251diff
changeset | 56 | definition top :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool" where | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 57 | "top r T \<equiv> \<forall>x. x \<sqsubseteq>\<^sub>r T" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 58 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35251diff
changeset | 59 | definition acc :: "'a ord \<Rightarrow> bool" where | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 60 |   "acc r \<equiv> wf {(y,x). x \<sqsubset>\<^sub>r y}"
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 61 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35251diff
changeset | 62 | definition closed :: "'a set \<Rightarrow> 'a binop \<Rightarrow> bool" where | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 63 | "closed A f \<equiv> \<forall>x\<in>A. \<forall>y\<in>A. x \<squnion>\<^sub>f y \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 64 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35251diff
changeset | 65 | definition semilat :: "'a sl \<Rightarrow> bool" where | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 66 | "semilat \<equiv> \<lambda>(A,r,f). order r \<and> closed A f \<and> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 67 | (\<forall>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y) \<and> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 68 | (\<forall>x\<in>A. \<forall>y\<in>A. y \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y) \<and> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 69 | (\<forall>x\<in>A. \<forall>y\<in>A. \<forall>z\<in>A. x \<sqsubseteq>\<^sub>r z \<and> y \<sqsubseteq>\<^sub>r z \<longrightarrow> x \<squnion>\<^sub>f y \<sqsubseteq>\<^sub>r z)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 70 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35251diff
changeset | 71 | definition is_ub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
 | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 72 | "is_ub r x y u \<equiv> (x,u)\<in>r \<and> (y,u)\<in>r" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 73 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35251diff
changeset | 74 | definition is_lub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
 | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 75 | "is_lub r x y u \<equiv> is_ub r x y u \<and> (\<forall>z. is_ub r x y z \<longrightarrow> (u,z)\<in>r)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 76 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35251diff
changeset | 77 | definition some_lub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 78 | "some_lub r x y \<equiv> SOME z. is_lub r x y z" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 79 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 80 | locale Semilat = | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 81 | fixes A :: "'a set" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 82 | fixes r :: "'a ord" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 83 | fixes f :: "'a binop" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 84 | assumes semilat: "semilat (A, r, f)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 85 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 86 | lemma order_refl [simp, intro]: "order r \<Longrightarrow> x \<sqsubseteq>\<^sub>r x" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 87 | (*<*) by (unfold order_def) (simp (no_asm_simp)) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 88 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 89 | lemma order_antisym: "\<lbrakk> order r; x \<sqsubseteq>\<^sub>r y; y \<sqsubseteq>\<^sub>r x \<rbrakk> \<Longrightarrow> x = y" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 90 | (*<*) by (unfold order_def) (simp (no_asm_simp)) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 91 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 92 | lemma order_trans: "\<lbrakk> order r; x \<sqsubseteq>\<^sub>r y; y \<sqsubseteq>\<^sub>r z \<rbrakk> \<Longrightarrow> x \<sqsubseteq>\<^sub>r z" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 93 | (*<*) by (unfold order_def) blast (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 94 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 95 | lemma order_less_irrefl [intro, simp]: "order r \<Longrightarrow> \<not> x \<sqsubset>\<^sub>r x" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 96 | (*<*) by (unfold order_def lesssub_def) blast (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 97 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 98 | lemma order_less_trans: "\<lbrakk> order r; x \<sqsubset>\<^sub>r y; y \<sqsubset>\<^sub>r z \<rbrakk> \<Longrightarrow> x \<sqsubset>\<^sub>r z" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 99 | (*<*) by (unfold order_def lesssub_def) blast (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 100 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 101 | lemma topD [simp, intro]: "top r T \<Longrightarrow> x \<sqsubseteq>\<^sub>r T" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 102 | (*<*) by (simp add: top_def) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 103 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 104 | lemma top_le_conv [simp]: "\<lbrakk> order r; top r T \<rbrakk> \<Longrightarrow> (T \<sqsubseteq>\<^sub>r x) = (x = T)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 105 | (*<*) by (blast intro: order_antisym) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 106 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 107 | lemma semilat_Def: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 108 | "semilat(A,r,f) \<equiv> order r \<and> closed A f \<and> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 109 | (\<forall>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y) \<and> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 110 | (\<forall>x\<in>A. \<forall>y\<in>A. y \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y) \<and> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 111 | (\<forall>x\<in>A. \<forall>y\<in>A. \<forall>z\<in>A. x \<sqsubseteq>\<^sub>r z \<and> y \<sqsubseteq>\<^sub>r z \<longrightarrow> x \<squnion>\<^sub>f y \<sqsubseteq>\<^sub>r z)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 112 | (*<*) by (unfold semilat_def) clarsimp (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 113 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 114 | lemma (in Semilat) orderI [simp, intro]: "order r" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 115 | (*<*) using semilat by (simp add: semilat_Def) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 116 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 117 | lemma (in Semilat) closedI [simp, intro]: "closed A f" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 118 | (*<*) using semilat by (simp add: semilat_Def) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 119 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 120 | lemma closedD: "\<lbrakk> closed A f; x\<in>A; y\<in>A \<rbrakk> \<Longrightarrow> x \<squnion>\<^sub>f y \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 121 | (*<*) by (unfold closed_def) blast (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 122 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 123 | lemma closed_UNIV [simp]: "closed UNIV f" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 124 | (*<*) by (simp add: closed_def) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 125 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 126 | lemma (in Semilat) closed_f [simp, intro]: "\<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> x \<squnion>\<^sub>f y \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 127 | (*<*) by (simp add: closedD [OF closedI]) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 128 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 129 | lemma (in Semilat) refl_r [intro, simp]: "x \<sqsubseteq>\<^sub>r x" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 130 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 131 | lemma (in Semilat) antisym_r [intro?]: "\<lbrakk> x \<sqsubseteq>\<^sub>r y; y \<sqsubseteq>\<^sub>r x \<rbrakk> \<Longrightarrow> x = y" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 132 | (*<*) by (rule order_antisym) auto (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 133 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 134 | lemma (in Semilat) trans_r [trans, intro?]: "\<lbrakk>x \<sqsubseteq>\<^sub>r y; y \<sqsubseteq>\<^sub>r z\<rbrakk> \<Longrightarrow> x \<sqsubseteq>\<^sub>r z" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 135 | (*<*) by (auto intro: order_trans) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 136 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 137 | lemma (in Semilat) ub1 [simp, intro?]: "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> x \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y" | 
| 63258 | 138 | (*<*) using semilat by (simp add: semilat_Def) (*>*) | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 139 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 140 | lemma (in Semilat) ub2 [simp, intro?]: "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> y \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y" | 
| 63258 | 141 | (*<*) using semilat by (simp add: semilat_Def) (*>*) | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 142 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 143 | lemma (in Semilat) lub [simp, intro?]: | 
| 58860 | 144 | "\<lbrakk> x \<sqsubseteq>\<^sub>r z; y \<sqsubseteq>\<^sub>r z; x \<in> A; y \<in> A; z \<in> A \<rbrakk> \<Longrightarrow> x \<squnion>\<^sub>f y \<sqsubseteq>\<^sub>r z" | 
| 63258 | 145 | (*<*) using semilat by (simp add: semilat_Def) (*>*) | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 146 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 147 | lemma (in Semilat) plus_le_conv [simp]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 148 | "\<lbrakk> x \<in> A; y \<in> A; z \<in> A \<rbrakk> \<Longrightarrow> (x \<squnion>\<^sub>f y \<sqsubseteq>\<^sub>r z) = (x \<sqsubseteq>\<^sub>r z \<and> y \<sqsubseteq>\<^sub>r z)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 149 | (*<*) by (blast intro: ub1 ub2 lub order_trans) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 150 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 151 | lemma (in Semilat) le_iff_plus_unchanged: "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> (x \<sqsubseteq>\<^sub>r y) = (x \<squnion>\<^sub>f y = y)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 152 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 153 | apply (rule iffI) | 
| 46226 | 154 | apply (blast intro: antisym_r lub ub2) | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 155 | apply (erule subst) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 156 | apply simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 157 | done | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 158 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 159 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 160 | lemma (in Semilat) le_iff_plus_unchanged2: "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> (x \<sqsubseteq>\<^sub>r y) = (y \<squnion>\<^sub>f x = y)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 161 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 162 | apply (rule iffI) | 
| 46226 | 163 | apply (blast intro: order_antisym lub ub1) | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 164 | apply (erule subst) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 165 | apply simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 166 | done | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 167 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 168 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 169 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 170 | lemma (in Semilat) plus_assoc [simp]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 171 | assumes a: "a \<in> A" and b: "b \<in> A" and c: "c \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 172 | shows "a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c) = a \<squnion>\<^sub>f b \<squnion>\<^sub>f c" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 173 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 174 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 175 | from a b have ab: "a \<squnion>\<^sub>f b \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 176 | from this c have abc: "(a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 177 | from b c have bc: "b \<squnion>\<^sub>f c \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 178 | from a this have abc': "a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c) \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 179 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 180 | show ?thesis | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 181 | proof | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 182 | show "a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c) \<sqsubseteq>\<^sub>r (a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 183 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 184 | from a b have "a \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f b" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 185 | also from ab c have "\<dots> \<sqsubseteq>\<^sub>r \<dots> \<squnion>\<^sub>f c" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 186 | finally have "a<": "a \<sqsubseteq>\<^sub>r (a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c" . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 187 | from a b have "b \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f b" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 188 | also from ab c have "\<dots> \<sqsubseteq>\<^sub>r \<dots> \<squnion>\<^sub>f c" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 189 | finally have "b<": "b \<sqsubseteq>\<^sub>r (a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c" . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 190 | from ab c have "c<": "c \<sqsubseteq>\<^sub>r (a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 191 | from "b<" "c<" b c abc have "b \<squnion>\<^sub>f c \<sqsubseteq>\<^sub>r (a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 192 | from "a<" this a bc abc show ?thesis .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 193 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 194 | show "(a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 195 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 196 | from b c have "b \<sqsubseteq>\<^sub>r b \<squnion>\<^sub>f c" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 197 | also from a bc have "\<dots> \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f \<dots>" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 198 | finally have "b<": "b \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c)" . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 199 | from b c have "c \<sqsubseteq>\<^sub>r b \<squnion>\<^sub>f c" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 200 | also from a bc have "\<dots> \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f \<dots>" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 201 | finally have "c<": "c \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c)" . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 202 | from a bc have "a<": "a \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c)" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 203 | from "a<" "b<" a b abc' have "a \<squnion>\<^sub>f b \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c)" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 204 | from this "c<" ab c abc' show ?thesis .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 205 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 206 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 207 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 208 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 209 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 210 | lemma (in Semilat) plus_com_lemma: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 211 | "\<lbrakk>a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> a \<squnion>\<^sub>f b \<sqsubseteq>\<^sub>r b \<squnion>\<^sub>f a" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 212 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 213 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 214 | assume a: "a \<in> A" and b: "b \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 215 | from b a have "a \<sqsubseteq>\<^sub>r b \<squnion>\<^sub>f a" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 216 | moreover from b a have "b \<sqsubseteq>\<^sub>r b \<squnion>\<^sub>f a" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 217 | moreover note a b | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 218 | moreover from b a have "b \<squnion>\<^sub>f a \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 219 | ultimately show ?thesis .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 220 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 221 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 222 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 223 | lemma (in Semilat) plus_commutative: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 224 | "\<lbrakk>a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> a \<squnion>\<^sub>f b = b \<squnion>\<^sub>f a" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 225 | (*<*) by(blast intro: order_antisym plus_com_lemma) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 226 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 227 | lemma is_lubD: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 228 | "is_lub r x y u \<Longrightarrow> is_ub r x y u \<and> (\<forall>z. is_ub r x y z \<longrightarrow> (u,z) \<in> r)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 229 | (*<*) by (simp add: is_lub_def) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 230 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 231 | lemma is_ubI: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 232 | "\<lbrakk> (x,u) \<in> r; (y,u) \<in> r \<rbrakk> \<Longrightarrow> is_ub r x y u" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 233 | (*<*) by (simp add: is_ub_def) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 234 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 235 | lemma is_ubD: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 236 | "is_ub r x y u \<Longrightarrow> (x,u) \<in> r \<and> (y,u) \<in> r" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 237 | (*<*) by (simp add: is_ub_def) (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 238 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 239 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 240 | lemma is_lub_bigger1 [iff]: | 
| 67613 | 241 | "is_lub (r\<^sup>*) x y y = ((x,y)\<in>r\<^sup>*)" | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 242 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 243 | apply (unfold is_lub_def is_ub_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 244 | apply blast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 245 | done | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 246 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 247 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 248 | lemma is_lub_bigger2 [iff]: | 
| 67613 | 249 | "is_lub (r\<^sup>*) x y x = ((y,x)\<in>r\<^sup>*)" | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 250 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 251 | apply (unfold is_lub_def is_ub_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 252 | apply blast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 253 | done | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 254 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 255 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 256 | lemma extend_lub: | 
| 67613 | 257 | "\<lbrakk> single_valued r; is_lub (r\<^sup>*) x y u; (x',x) \<in> r \<rbrakk> | 
| 258 | \<Longrightarrow> \<exists>v. is_lub (r\<^sup>*) x' y v" | |
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 259 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 260 | apply (unfold is_lub_def is_ub_def) | 
| 67613 | 261 | apply (case_tac "(y,x) \<in> r\<^sup>*") | 
| 262 | apply (case_tac "(y,x') \<in> r\<^sup>*") | |
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 263 | apply blast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 264 | apply (blast elim: converse_rtranclE dest: single_valuedD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 265 | apply (rule exI) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 266 | apply (rule conjI) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 267 | apply (blast intro: converse_rtrancl_into_rtrancl dest: single_valuedD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 268 | apply (blast intro: rtrancl_into_rtrancl converse_rtrancl_into_rtrancl | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 269 | elim: converse_rtranclE dest: single_valuedD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 270 | done | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 271 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 272 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 273 | lemma single_valued_has_lubs [rule_format]: | 
| 67613 | 274 | "\<lbrakk>single_valued r; (x,u) \<in> r\<^sup>*\<rbrakk> \<Longrightarrow> (\<forall>y. (y,u) \<in> r\<^sup>* \<longrightarrow> | 
| 275 | (\<exists>z. is_lub (r\<^sup>*) x y z))" | |
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 276 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 277 | apply (erule converse_rtrancl_induct) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 278 | apply clarify | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 279 | apply (erule converse_rtrancl_induct) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 280 | apply blast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 281 | apply (blast intro: converse_rtrancl_into_rtrancl) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 282 | apply (blast intro: extend_lub) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 283 | done | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 284 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 285 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 286 | lemma some_lub_conv: | 
| 67613 | 287 | "\<lbrakk>acyclic r; is_lub (r\<^sup>*) x y u\<rbrakk> \<Longrightarrow> some_lub (r\<^sup>*) x y = u" | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 288 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 289 | apply (unfold some_lub_def is_lub_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 290 | apply (rule someI2) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 291 | apply assumption | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 292 | apply (blast intro: antisymD dest!: acyclic_impl_antisym_rtrancl) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 293 | done | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 294 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 295 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 296 | lemma is_lub_some_lub: | 
| 67613 | 297 | "\<lbrakk>single_valued r; acyclic r; (x,u)\<in>r\<^sup>*; (y,u)\<in>r\<^sup>*\<rbrakk> | 
| 298 | \<Longrightarrow> is_lub (r\<^sup>*) x y (some_lub (r\<^sup>*) x y)" | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42463diff
changeset | 299 | (*<*) by (fastforce dest: single_valued_has_lubs simp add: some_lub_conv) (*>*) | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 300 | |
| 61361 | 301 | subsection\<open>An executable lub-finder\<close> | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 302 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35251diff
changeset | 303 | definition exec_lub :: "('a * 'a) set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a binop" where
 | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 304 | "exec_lub r f x y \<equiv> while (\<lambda>z. (x,z) \<notin> r\<^sup>*) f y" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 305 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 306 | lemma exec_lub_refl: "exec_lub r f T T = T" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 307 | by (simp add: exec_lub_def while_unfold) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 308 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 309 | lemma acyclic_single_valued_finite: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 310 | "\<lbrakk>acyclic r; single_valued r; (x,y) \<in> r\<^sup>*\<rbrakk> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 311 |   \<Longrightarrow> finite (r \<inter> {a. (x, a) \<in> r\<^sup>*} \<times> {b. (b, y) \<in> r\<^sup>*})"
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 312 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 313 | apply(erule converse_rtrancl_induct) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 314 |  apply(rule_tac B = "{}" in finite_subset)
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 315 | apply(simp only:acyclic_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 316 | apply(blast intro:rtrancl_into_trancl2 rtrancl_trancl_trancl) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 317 | apply simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 318 | apply(rename_tac x x') | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 319 | apply(subgoal_tac "r \<inter> {a. (x,a) \<in> r\<^sup>*} \<times> {b. (b,y) \<in> r\<^sup>*} =
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 320 |                    insert (x,x') (r \<inter> {a. (x', a) \<in> r\<^sup>*} \<times> {b. (b, y) \<in> r\<^sup>*})")
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 321 | apply simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 322 | apply(blast intro:converse_rtrancl_into_rtrancl | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 323 | elim:converse_rtranclE dest:single_valuedD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 324 | done | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 325 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 326 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 327 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 328 | lemma exec_lub_conv: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 329 | "\<lbrakk> acyclic r; \<forall>x y. (x,y) \<in> r \<longrightarrow> f x = y; is_lub (r\<^sup>*) x y u \<rbrakk> \<Longrightarrow> | 
| 58860 | 330 | exec_lub r f x y = u" | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 331 | (*<*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 332 | apply(unfold exec_lub_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 333 | apply(rule_tac P = "\<lambda>z. (y,z) \<in> r\<^sup>* \<and> (z,u) \<in> r\<^sup>*" and | 
| 67613 | 334 |                r = "(r \<inter> {(a,b). (y,a) \<in> r\<^sup>* \<and> (b,u) \<in> r\<^sup>*})\<inverse>" in while_rule)
 | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 335 | apply(blast dest: is_lubD is_ubD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 336 | apply(erule conjE) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 337 | apply(erule_tac z = u in converse_rtranclE) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 338 | apply(blast dest: is_lubD is_ubD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 339 | apply(blast dest:rtrancl_into_rtrancl) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 340 | apply(rename_tac s) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 341 | apply(subgoal_tac "is_ub (r\<^sup>*) x y s") | 
| 58860 | 342 | prefer 2 apply(simp add:is_ub_def) | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 343 | apply(subgoal_tac "(u, s) \<in> r\<^sup>*") | 
| 58860 | 344 | prefer 2 apply(blast dest:is_lubD) | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 345 | apply(erule converse_rtranclE) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 346 | apply blast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 347 | apply(simp only:acyclic_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 348 | apply(blast intro:rtrancl_into_trancl2 rtrancl_trancl_trancl) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 349 | apply(rule finite_acyclic_wf) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 350 | apply simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 351 | apply(erule acyclic_single_valued_finite) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 352 | apply(blast intro:single_valuedI) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 353 | apply(simp add:is_lub_def is_ub_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 354 | apply simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 355 | apply(erule acyclic_subset) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 356 | apply blast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 357 | apply simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 358 | apply(erule conjE) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 359 | apply(erule_tac z = u in converse_rtranclE) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 360 | apply(blast dest: is_lubD is_ubD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 361 | apply(blast dest:rtrancl_into_rtrancl) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 362 | done | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 363 | (*>*) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 364 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 365 | lemma is_lub_exec_lub: | 
| 67613 | 366 | "\<lbrakk> single_valued r; acyclic r; (x,u)\<in>r\<^sup>*; (y,u)\<in>r\<^sup>*; \<forall>x y. (x,y) \<in> r \<longrightarrow> f x = y \<rbrakk> | 
| 367 | \<Longrightarrow> is_lub (r\<^sup>*) x y (exec_lub r f x y)" | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42463diff
changeset | 368 | (*<*) by (fastforce dest: single_valued_has_lubs simp add: exec_lub_conv) (*>*) | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 369 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 370 | end |