| author | wenzelm | 
| Sun, 02 Nov 2014 16:59:40 +0100 | |
| changeset 58875 | ab1c65b015c3 | 
| parent 42150 | b0c0638c4aad | 
| child 58886 | 8a6cac7c7247 | 
| permissions | -rw-r--r-- | 
| 42150 | 1 | (* Title: HOL/MicroJava/DFA/SemilatAlg.thy | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 2 | Author: Gerwin Klein | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 3 | Copyright 2002 Technische Universitaet Muenchen | 
| 
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 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 6 | header {* \isaheader{More on Semilattices} *}
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 7 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 8 | theory SemilatAlg | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 9 | imports Typing_Framework Product | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 10 | begin | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 11 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35109diff
changeset | 12 | definition lesubstep_type :: "(nat \<times> 's) list \<Rightarrow> 's ord \<Rightarrow> (nat \<times> 's) list \<Rightarrow> bool" | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35109diff
changeset | 13 |                     ("(_ /<=|_| _)" [50, 0, 51] 50) where
 | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 14 | "x <=|r| y \<equiv> \<forall>(p,s) \<in> set x. \<exists>s'. (p,s') \<in> set y \<and> s <=_r s'" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 15 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35109diff
changeset | 16 | primrec plusplussub :: "'a list \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" ("(_ /++'__ _)" [65, 1000, 66] 65) where
 | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 17 | "[] ++_f y = y" | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35109diff
changeset | 18 | | "(x#xs) ++_f y = xs ++_f (x +_f y)" | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 19 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35109diff
changeset | 20 | definition bounded :: "'s step_type \<Rightarrow> nat \<Rightarrow> bool" where | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 21 | "bounded step n == !p<n. !s. !(q,t):set(step p s). q<n" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 22 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35109diff
changeset | 23 | definition pres_type :: "'s step_type \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool" where | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 24 | "pres_type step n A == \<forall>s\<in>A. \<forall>p<n. \<forall>(q,s')\<in>set (step p s). s' \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 25 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35109diff
changeset | 26 | definition mono :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool" where | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 27 | "mono r step n A == | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 28 | \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> step p s <=|r| step p t" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 29 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 30 | lemma pres_typeD: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 31 | "\<lbrakk> pres_type step n A; s\<in>A; p<n; (q,s')\<in>set (step p s) \<rbrakk> \<Longrightarrow> s' \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 32 | by (unfold pres_type_def, blast) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 33 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 34 | lemma monoD: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 35 | "\<lbrakk> mono r step n A; p < n; s\<in>A; s <=_r t \<rbrakk> \<Longrightarrow> step p s <=|r| step p t" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 36 | by (unfold mono_def, blast) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 37 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 38 | lemma boundedD: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 39 | "\<lbrakk> bounded step n; p < n; (q,t) : set (step p xs) \<rbrakk> \<Longrightarrow> q < n" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 40 | by (unfold bounded_def, blast) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 41 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 42 | lemma lesubstep_type_refl [simp, intro]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 43 | "(\<And>x. x <=_r x) \<Longrightarrow> x <=|r| x" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 44 | by (unfold lesubstep_type_def) auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 45 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 46 | lemma lesub_step_typeD: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 47 | "a <=|r| b \<Longrightarrow> (x,y) \<in> set a \<Longrightarrow> \<exists>y'. (x, y') \<in> set b \<and> y <=_r y'" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 48 | by (unfold lesubstep_type_def) blast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 49 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 50 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 51 | lemma list_update_le_listI [rule_format]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 52 | "set xs <= A \<longrightarrow> set ys <= A \<longrightarrow> xs <=[r] ys \<longrightarrow> p < size xs \<longrightarrow> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 53 | x <=_r ys!p \<longrightarrow> semilat(A,r,f) \<longrightarrow> x\<in>A \<longrightarrow> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 54 | xs[p := x +_f xs!p] <=[r] ys" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 55 | apply (unfold Listn.le_def lesub_def semilat_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 56 | apply (simp add: list_all2_conv_all_nth nth_list_update) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 57 | done | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 58 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 59 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 60 | lemma plusplus_closed: assumes "semilat (A, r, f)" shows | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 61 | "\<And>y. \<lbrakk> set x \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> x ++_f y \<in> A" (is "PROP ?P") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 62 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 63 | interpret Semilat A r f using assms by (rule Semilat.intro) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 64 | show "PROP ?P" proof (induct x) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 65 | show "\<And>y. y \<in> A \<Longrightarrow> [] ++_f y \<in> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 66 | fix y x xs | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 67 | assume y: "y \<in> A" and xs: "set (x#xs) \<subseteq> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 68 | assume IH: "\<And>y. \<lbrakk> set xs \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> xs ++_f y \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 69 | from xs obtain x: "x \<in> A" and xs': "set xs \<subseteq> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 70 | from x y have "(x +_f y) \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 71 | with xs' have "xs ++_f (x +_f y) \<in> A" by (rule IH) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 72 | thus "(x#xs) ++_f y \<in> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 73 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 74 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 75 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 76 | lemma (in Semilat) pp_ub2: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 77 | "\<And>y. \<lbrakk> set x \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> y <=_r x ++_f y" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 78 | proof (induct x) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 79 | from semilat show "\<And>y. y <=_r [] ++_f y" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 80 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 81 | fix y a l | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 82 | assume y: "y \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 83 | assume "set (a#l) \<subseteq> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 84 | then obtain a: "a \<in> A" and x: "set l \<subseteq> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 85 | assume "\<And>y. \<lbrakk>set l \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> y <=_r l ++_f y" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 86 | hence IH: "\<And>y. y \<in> A \<Longrightarrow> y <=_r l ++_f y" using x . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 87 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 88 | from a y have "y <=_r a +_f y" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 89 | also from a y have "a +_f y \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 90 | hence "(a +_f y) <=_r l ++_f (a +_f y)" by (rule IH) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 91 | finally have "y <=_r l ++_f (a +_f y)" . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 92 | thus "y <=_r (a#l) ++_f y" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 93 | qed | 
| 
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 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 96 | lemma (in Semilat) pp_ub1: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 97 | shows "\<And>y. \<lbrakk>set ls \<subseteq> A; y \<in> A; x \<in> set ls\<rbrakk> \<Longrightarrow> x <=_r ls ++_f y" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 98 | proof (induct ls) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 99 | show "\<And>y. x \<in> set [] \<Longrightarrow> x <=_r [] ++_f y" by simp | 
| 
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 | fix y s ls | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 102 | assume "set (s#ls) \<subseteq> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 103 | then obtain s: "s \<in> A" and ls: "set ls \<subseteq> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 104 | assume y: "y \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 105 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 106 | assume | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 107 | "\<And>y. \<lbrakk>set ls \<subseteq> A; y \<in> A; x \<in> set ls\<rbrakk> \<Longrightarrow> x <=_r ls ++_f y" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 108 | hence IH: "\<And>y. x \<in> set ls \<Longrightarrow> y \<in> A \<Longrightarrow> x <=_r ls ++_f y" using ls . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 109 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 110 | assume "x \<in> set (s#ls)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 111 | then obtain xls: "x = s \<or> x \<in> set ls" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 112 |   moreover {
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 113 | assume xs: "x = s" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 114 | from s y have "s <=_r s +_f y" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 115 | also from s y have "s +_f y \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 116 | with ls have "(s +_f y) <=_r ls ++_f (s +_f y)" by (rule pp_ub2) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 117 | finally have "s <=_r ls ++_f (s +_f y)" . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 118 | with xs have "x <=_r ls ++_f (s +_f y)" by simp | 
| 
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 |   moreover {
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 121 | assume "x \<in> set ls" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 122 | hence "\<And>y. y \<in> A \<Longrightarrow> x <=_r ls ++_f y" by (rule IH) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 123 | moreover from s y have "s +_f y \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 124 | ultimately have "x <=_r ls ++_f (s +_f y)" . | 
| 
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 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 127 | have "x <=_r ls ++_f (s +_f y)" by blast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 128 | thus "x <=_r (s#ls) ++_f y" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 129 | qed | 
| 
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 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 132 | lemma (in Semilat) pp_lub: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 133 | assumes z: "z \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 134 | shows | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 135 | "\<And>y. y \<in> A \<Longrightarrow> set xs \<subseteq> A \<Longrightarrow> \<forall>x \<in> set xs. x <=_r z \<Longrightarrow> y <=_r z \<Longrightarrow> xs ++_f y <=_r z" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 136 | proof (induct xs) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 137 | fix y assume "y <=_r z" thus "[] ++_f y <=_r z" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 138 | next | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 139 | fix y l ls assume y: "y \<in> A" and "set (l#ls) \<subseteq> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 140 | then obtain l: "l \<in> A" and ls: "set ls \<subseteq> A" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 141 | assume "\<forall>x \<in> set (l#ls). x <=_r z" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 142 | then obtain lz: "l <=_r z" and lsz: "\<forall>x \<in> set ls. x <=_r z" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 143 | assume "y <=_r z" with lz have "l +_f y <=_r z" using l y z .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 144 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 145 | from l y have "l +_f y \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 146 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 147 | assume "\<And>y. y \<in> A \<Longrightarrow> set ls \<subseteq> A \<Longrightarrow> \<forall>x \<in> set ls. x <=_r z \<Longrightarrow> y <=_r z | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 148 | \<Longrightarrow> ls ++_f y <=_r z" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 149 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 150 | have "ls ++_f (l +_f y) <=_r z" using ls lsz by - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 151 | thus "(l#ls) ++_f y <=_r z" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 152 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 153 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 154 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 155 | lemma ub1': | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 156 | assumes "semilat (A, r, f)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 157 | shows "\<lbrakk>\<forall>(p,s) \<in> set S. s \<in> A; y \<in> A; (a,b) \<in> set S\<rbrakk> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 158 | \<Longrightarrow> b <=_r map snd [(p', t')\<leftarrow>S. p' = a] ++_f y" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 159 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 160 | interpret Semilat A r f using assms by (rule Semilat.intro) | 
| 
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 | let "b <=_r ?map ++_f y" = ?thesis | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 163 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 164 | assume "y \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 165 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 166 | assume "\<forall>(p,s) \<in> set S. s \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 167 | hence "set ?map \<subseteq> A" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 168 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 169 | assume "(a,b) \<in> set S" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 170 | hence "b \<in> set ?map" by (induct S, auto) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 171 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 172 | show ?thesis by - (rule pp_ub1) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 173 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 174 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 175 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 176 | lemma plusplus_empty: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 177 | "\<forall>s'. (q, s') \<in> set S \<longrightarrow> s' +_f ss ! q = ss ! q \<Longrightarrow> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 178 | (map snd [(p', t') \<leftarrow> S. p' = q] ++_f ss ! q) = ss ! q" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 179 | by (induct S) auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 180 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 181 | end |