--- a/src/HOL/Bali/AxCompl.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/AxCompl.thy Sat Jan 02 18:48:45 2016 +0100
@@ -2,18 +2,18 @@
Author: David von Oheimb and Norbert Schirmer
*)
-subsection {*
+subsection \<open>
Completeness proof for Axiomatic semantics of Java expressions and statements
-*}
+\<close>
theory AxCompl imports AxSem begin
-text {*
+text \<open>
design issues:
\begin{itemize}
\item proof structured by Most General Formulas (-> Thomas Kleymann)
\end{itemize}
-*}
+\<close>
@@ -170,10 +170,10 @@
done
lemmas MGFnD' = MGFnD [of _ _ _ _ "\<lambda>x. True"]
-text {* To derive the most general formula, we can always assume a normal
+text \<open>To derive the most general formula, we can always assume a normal
state in the precondition, since abrupt cases can be handled uniformally by
the abrupt rule.
-*}
+\<close>
lemma MGFNormalI: "G,A\<turnstile>{Normal \<doteq>} t\<succ> {G\<rightarrow>} \<Longrightarrow>
G,(A::state triple set)\<turnstile>{\<doteq>::state assn} t\<succ> {G\<rightarrow>}"
apply (unfold MGF_def)
@@ -191,8 +191,8 @@
apply clarsimp
done
-text {* Additionally to @{text MGFNormalI}, we also expand the definition of
-the most general formula here *}
+text \<open>Additionally to \<open>MGFNormalI\<close>, we also expand the definition of
+the most general formula here\<close>
lemma MGFn_NormalI:
"G,(A::state triple set)\<turnstile>{Normal((\<lambda>Y' s' s. s'=s \<and> normal s) \<and>. G\<turnstile>init\<le>n)}t\<succ>
{\<lambda>Y s' s. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y,s')} \<Longrightarrow> G,A\<turnstile>{=:n}t\<succ>{G\<rightarrow>}"
@@ -204,9 +204,9 @@
apply (clarsimp simp add: Let_def)
done
-text {* To derive the most general formula, we can restrict ourselves to
+text \<open>To derive the most general formula, we can restrict ourselves to
welltyped terms, since all others can be uniformally handled by the hazard
-rule. *}
+rule.\<close>
lemma MGFn_free_wt:
"(\<exists>T L C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T)
\<longrightarrow> G,(A::state triple set)\<turnstile>{=:n} t\<succ> {G\<rightarrow>}
@@ -216,10 +216,10 @@
apply (auto elim: conseq12 simp add: MGFn_def MGF_def)
done
-text {* To derive the most general formula, we can restrict ourselves to
+text \<open>To derive the most general formula, we can restrict ourselves to
welltyped terms and assume that the state in the precondition conforms to the
environment. All type violations can be uniformally handled by the hazard
-rule. *}
+rule.\<close>
lemma MGFn_free_wt_NormalConformI:
"(\<forall> T L C . \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T
\<longrightarrow> G,(A::state triple set)
@@ -240,10 +240,10 @@
apply blast
done
-text {* To derive the most general formula, we can restrict ourselves to
+text \<open>To derive the most general formula, we can restrict ourselves to
welltyped terms and assume that the state in the precondition conforms to the
environment and that the term is definetly assigned with respect to this state.
-All type violations can be uniformally handled by the hazard rule. *}
+All type violations can be uniformally handled by the hazard rule.\<close>
lemma MGFn_free_wt_da_NormalConformI:
"(\<forall> T L C B. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T
\<longrightarrow> G,(A::state triple set)
@@ -543,11 +543,11 @@
[where ?Env="\<lparr>prg=G, cls=C,lcl=L\<rparr>",simplified],auto)
-text {* To derive the most general formula for the loop statement, we need to
+text \<open>To derive the most general formula for the loop statement, we need to
come up with a proper loop invariant, which intuitively states that we are
currently inside the evaluation of the loop. To define such an invariant, we
unroll the loop in iterated evaluations of the expression and evaluations of
-the loop body. *}
+the loop body.\<close>
definition
unroll :: "prog \<Rightarrow> label \<Rightarrow> expr \<Rightarrow> stmt \<Rightarrow> (state \<times> state) set" where
@@ -719,7 +719,7 @@
show "Y' = \<diamondsuit> \<and> G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
proof -
from asm obtain v t where
- -- {* @{term "Z'"} gets instantiated with @{term "Norm s"} *}
+ \<comment> \<open>@{term "Z'"} gets instantiated with @{term "Norm s"}\<close>
unroll: "(Norm s, t) \<in> (unroll G l e c)\<^sup>*" and
eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'" and
normal_termination: "normal s' \<longrightarrow> \<not> the_Bool v" and
@@ -1037,13 +1037,13 @@
done
next
case (FVar accC statDeclC stat e fn)
- from MGFn_Init [OF hyp] and `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}` and wf
+ from MGFn_Init [OF hyp] and \<open>G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}\<close> and wf
show ?case
by (rule MGFn_FVar)
next
case (AVar e1 e2)
- note mgf_e1 = `G,A\<turnstile>{=:n} \<langle>e1\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
- note mgf_e2 = `G,A\<turnstile>{=:n} \<langle>e2\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
+ note mgf_e1 = \<open>G,A\<turnstile>{=:n} \<langle>e1\<rangle>\<^sub>e\<succ> {G\<rightarrow>}\<close>
+ note mgf_e2 = \<open>G,A\<turnstile>{=:n} \<langle>e2\<rangle>\<^sub>e\<succ> {G\<rightarrow>}\<close>
show "G,A\<turnstile>{=:n} \<langle>e1.[e2]\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
apply (rule MGFn_NormalI)
apply (rule ax_derivs.AVar)
@@ -1186,8 +1186,8 @@
done
next
case (Call accC statT mode e mn pTs' ps)
- note mgf_e = `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
- note mgf_ps = `G,A\<turnstile>{=:n} \<langle>ps\<rangle>\<^sub>l\<succ> {G\<rightarrow>}`
+ note mgf_e = \<open>G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}\<close>
+ note mgf_ps = \<open>G,A\<turnstile>{=:n} \<langle>ps\<rangle>\<^sub>l\<succ> {G\<rightarrow>}\<close>
from mgf_methds mgf_e mgf_ps wf
show "G,A\<turnstile>{=:n} \<langle>{accC,statT,mode}e\<cdot>mn({pTs'}ps)\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
by (rule MGFn_Call)
@@ -1198,7 +1198,7 @@
by simp
next
case (Body D c)
- note mgf_c = `G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
+ note mgf_c = \<open>G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}\<close>
from wf MGFn_Init [OF hyp] mgf_c
show "G,A\<turnstile>{=:n} \<langle>Body D c\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
by (rule MGFn_Body)
@@ -1266,8 +1266,8 @@
done
next
case (Loop l e c)
- note mgf_e = `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
- note mgf_c = `G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
+ note mgf_e = \<open>G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}\<close>
+ note mgf_c = \<open>G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}\<close>
from mgf_e mgf_c wf
show "G,A\<turnstile>{=:n} \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
by (rule MGFn_Loop)
@@ -1303,8 +1303,8 @@
done
next
case (Fin c1 c2)
- note mgf_c1 = `G,A\<turnstile>{=:n} \<langle>c1\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
- note mgf_c2 = `G,A\<turnstile>{=:n} \<langle>c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
+ note mgf_c1 = \<open>G,A\<turnstile>{=:n} \<langle>c1\<rangle>\<^sub>s\<succ> {G\<rightarrow>}\<close>
+ note mgf_c2 = \<open>G,A\<turnstile>{=:n} \<langle>c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}\<close>
from wf mgf_c1 mgf_c2
show "G,A\<turnstile>{=:n} \<langle>c1 Finally c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
by (rule MGFn_Fin)
@@ -1370,7 +1370,7 @@
apply -
apply (induct_tac "n")
apply (tactic "ALLGOALS (clarsimp_tac @{context})")
-apply (tactic {* dresolve_tac @{context} [Thm.permute_prems 0 1 @{thm card_seteq}] 1 *})
+apply (tactic \<open>dresolve_tac @{context} [Thm.permute_prems 0 1 @{thm card_seteq}] 1\<close>)
apply simp
apply (erule finite_imageI)
apply (simp add: MGF_asm ax_derivs_asm)
--- a/src/HOL/Bali/AxExample.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/AxExample.thy Sat Jan 02 18:48:45 2016 +0100
@@ -2,7 +2,7 @@
Author: David von Oheimb
*)
-subsection {* Example of a proof based on the Bali axiomatic semantics *}
+subsection \<open>Example of a proof based on the Bali axiomatic semantics\<close>
theory AxExample
imports AxSem Example
@@ -40,7 +40,7 @@
declare split_if_asm [split del]
declare lvar_def [simp]
-ML {*
+ML \<open>
fun inst1_tac ctxt s t xs st =
(case AList.lookup (op =) (rev (Term.add_var_names (Thm.prop_of st) [])) s of
SOME i => PRIMITIVE (Rule_Insts.read_instantiate ctxt [(((s, i), Position.none), t)] xs) st
@@ -50,7 +50,7 @@
REPEAT o resolve_tac ctxt [allI] THEN'
resolve_tac ctxt
@{thms ax_Skip ax_StatRef ax_MethdN ax_Alloc ax_Alloc_Arr ax_SXAlloc_Normal ax_derivs.intros(8-)};
-*}
+\<close>
theorem ax_test: "tprg,({}::'a triple set)\<turnstile>
@@ -64,8 +64,8 @@
precondition. *)
apply (tactic "ax_tac @{context} 1" (* Try *))
defer
-apply (tactic {* inst1_tac @{context} "Q"
- "\<lambda>Y s Z. arr_inv (snd s) \<and> tprg,s\<turnstile>catch SXcpt NullPointer" [] *})
+apply (tactic \<open>inst1_tac @{context} "Q"
+ "\<lambda>Y s Z. arr_inv (snd s) \<and> tprg,s\<turnstile>catch SXcpt NullPointer" []\<close>)
prefer 2
apply simp
apply (rule_tac P' = "Normal (\<lambda>Y s Z. arr_inv (snd s))" in conseq1)
@@ -84,7 +84,7 @@
apply (tactic "ax_tac @{context} 1" (* AVar *))
prefer 2
apply (rule ax_subst_Val_allI)
-apply (tactic {* inst1_tac @{context} "P'" "\<lambda>a. Normal (PP a\<leftarrow>x)" ["PP", "x"] *})
+apply (tactic \<open>inst1_tac @{context} "P'" "\<lambda>a. Normal (PP a\<leftarrow>x)" ["PP", "x"]\<close>)
apply (simp del: avar_def2 peek_and_def2)
apply (tactic "ax_tac @{context} 1")
apply (tactic "ax_tac @{context} 1")
@@ -125,25 +125,25 @@
apply (tactic "ax_tac @{context} 1") (* Ass *)
prefer 2
apply (rule ax_subst_Var_allI)
-apply (tactic {* inst1_tac @{context} "P'" "\<lambda>a vs l vf. PP a vs l vf\<leftarrow>x \<and>. p" ["PP", "x", "p"] *})
+apply (tactic \<open>inst1_tac @{context} "P'" "\<lambda>a vs l vf. PP a vs l vf\<leftarrow>x \<and>. p" ["PP", "x", "p"]\<close>)
apply (rule allI)
-apply (tactic {* simp_tac (@{context} delloop "split_all_tac" delsimps [@{thm peek_and_def2}, @{thm heap_def2}, @{thm subst_res_def2}, @{thm normal_def2}]) 1 *})
+apply (tactic \<open>simp_tac (@{context} delloop "split_all_tac" delsimps [@{thm peek_and_def2}, @{thm heap_def2}, @{thm subst_res_def2}, @{thm normal_def2}]) 1\<close>)
apply (rule ax_derivs.Abrupt)
apply (simp (no_asm))
apply (tactic "ax_tac @{context} 1" (* FVar *))
apply (tactic "ax_tac @{context} 2", tactic "ax_tac @{context} 2", tactic "ax_tac @{context} 2")
apply (tactic "ax_tac @{context} 1")
-apply (tactic {* inst1_tac @{context} "R" "\<lambda>a'. Normal ((\<lambda>Vals:vs (x, s) Z. arr_inv s \<and> inited Ext (globs s) \<and> a' \<noteq> Null \<and> vs = [Null]) \<and>. heap_free two)" [] *})
+apply (tactic \<open>inst1_tac @{context} "R" "\<lambda>a'. Normal ((\<lambda>Vals:vs (x, s) Z. arr_inv s \<and> inited Ext (globs s) \<and> a' \<noteq> Null \<and> vs = [Null]) \<and>. heap_free two)" []\<close>)
apply fastforce
prefer 4
apply (rule ax_derivs.Done [THEN conseq1],force)
apply (rule ax_subst_Val_allI)
-apply (tactic {* inst1_tac @{context} "P'" "\<lambda>a. Normal (PP a\<leftarrow>x)" ["PP", "x"] *})
+apply (tactic \<open>inst1_tac @{context} "P'" "\<lambda>a. Normal (PP a\<leftarrow>x)" ["PP", "x"]\<close>)
apply (simp (no_asm) del: peek_and_def2 heap_free_def2 normal_def2 o_apply)
apply (tactic "ax_tac @{context} 1")
prefer 2
apply (rule ax_subst_Val_allI)
-apply (tactic {* inst1_tac @{context} "P'" "\<lambda>aa v. Normal (QQ aa v\<leftarrow>y)" ["QQ", "y"] *})
+apply (tactic \<open>inst1_tac @{context} "P'" "\<lambda>aa v. Normal (QQ aa v\<leftarrow>y)" ["QQ", "y"]\<close>)
apply (simp del: peek_and_def2 heap_free_def2 normal_def2)
apply (tactic "ax_tac @{context} 1")
apply (tactic "ax_tac @{context} 1")
@@ -162,7 +162,7 @@
apply (tactic "ax_tac @{context} 1")
defer
apply (rule ax_subst_Var_allI)
-apply (tactic {* inst1_tac @{context} "P'" "\<lambda>vf. Normal (PP vf \<and>. p)" ["PP", "p"] *})
+apply (tactic \<open>inst1_tac @{context} "P'" "\<lambda>vf. Normal (PP vf \<and>. p)" ["PP", "p"]\<close>)
apply (simp (no_asm) del: split_paired_All peek_and_def2 initd_def2 heap_free_def2 normal_def2)
apply (tactic "ax_tac @{context} 1" (* NewC *))
apply (tactic "ax_tac @{context} 1" (* ax_Alloc *))
@@ -177,43 +177,43 @@
apply (rule ax_InitS)
apply force
apply (simp (no_asm))
-apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 1 *})
+apply (tactic \<open>simp_tac (@{context} delloop "split_all_tac") 1\<close>)
apply (rule ax_Init_Skip_lemma)
-apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 1 *})
+apply (tactic \<open>simp_tac (@{context} delloop "split_all_tac") 1\<close>)
apply (rule ax_InitS [THEN conseq1] (* init Base *))
apply force
apply (simp (no_asm))
apply (unfold arr_viewed_from_def)
apply (rule allI)
apply (rule_tac P' = "Normal P" and P = P for P in conseq1)
-apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 1 *})
+apply (tactic \<open>simp_tac (@{context} delloop "split_all_tac") 1\<close>)
apply (tactic "ax_tac @{context} 1")
apply (tactic "ax_tac @{context} 1")
apply (rule_tac [2] ax_subst_Var_allI)
-apply (tactic {* inst1_tac @{context} "P'" "\<lambda>vf l vfa. Normal (P vf l vfa)" ["P"] *})
-apply (tactic {* simp_tac (@{context} delloop "split_all_tac" delsimps [@{thm split_paired_All}, @{thm peek_and_def2}, @{thm heap_free_def2}, @{thm initd_def2}, @{thm normal_def2}, @{thm supd_lupd}]) 2 *})
+apply (tactic \<open>inst1_tac @{context} "P'" "\<lambda>vf l vfa. Normal (P vf l vfa)" ["P"]\<close>)
+apply (tactic \<open>simp_tac (@{context} delloop "split_all_tac" delsimps [@{thm split_paired_All}, @{thm peek_and_def2}, @{thm heap_free_def2}, @{thm initd_def2}, @{thm normal_def2}, @{thm supd_lupd}]) 2\<close>)
apply (tactic "ax_tac @{context} 2" (* NewA *))
apply (tactic "ax_tac @{context} 3" (* ax_Alloc_Arr *))
apply (tactic "ax_tac @{context} 3")
-apply (tactic {* inst1_tac @{context} "P" "\<lambda>vf l vfa. Normal (P vf l vfa\<leftarrow>\<diamondsuit>)" ["P"] *})
-apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 2 *})
+apply (tactic \<open>inst1_tac @{context} "P" "\<lambda>vf l vfa. Normal (P vf l vfa\<leftarrow>\<diamondsuit>)" ["P"]\<close>)
+apply (tactic \<open>simp_tac (@{context} delloop "split_all_tac") 2\<close>)
apply (tactic "ax_tac @{context} 2")
apply (tactic "ax_tac @{context} 1" (* FVar *))
apply (tactic "ax_tac @{context} 2" (* StatRef *))
apply (rule ax_derivs.Done [THEN conseq1])
-apply (tactic {* inst1_tac @{context} "Q" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf=lvar (VName e) (snd s)) \<and>. heap_free four \<and>. initd Base \<and>. initd Ext)" [] *})
+apply (tactic \<open>inst1_tac @{context} "Q" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf=lvar (VName e) (snd s)) \<and>. heap_free four \<and>. initd Base \<and>. initd Ext)" []\<close>)
apply (clarsimp split del: split_if)
apply (frule atleast_free_weaken [THEN atleast_free_weaken])
apply (drule initedD)
apply (clarsimp elim!: atleast_free_SucD simp add: arr_inv_def)
apply force
-apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 1 *})
+apply (tactic \<open>simp_tac (@{context} delloop "split_all_tac") 1\<close>)
apply (rule ax_triv_Init_Object [THEN peek_and_forget2, THEN conseq1])
apply (rule wf_tprg)
apply clarsimp
-apply (tactic {* inst1_tac @{context} "P" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf = lvar (VName e) (snd s)) \<and>. heap_free four \<and>. initd Ext)" [] *})
+apply (tactic \<open>inst1_tac @{context} "P" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf = lvar (VName e) (snd s)) \<and>. heap_free four \<and>. initd Ext)" []\<close>)
apply clarsimp
-apply (tactic {* inst1_tac @{context} "PP" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf = lvar (VName e) (snd s)) \<and>. heap_free four \<and>. Not \<circ> initd Base)" [] *})
+apply (tactic \<open>inst1_tac @{context} "PP" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf = lvar (VName e) (snd s)) \<and>. heap_free four \<and>. Not \<circ> initd Base)" []\<close>)
apply clarsimp
(* end init *)
apply (rule conseq1)
@@ -245,7 +245,7 @@
apply clarsimp
apply (tactic "ax_tac @{context} 1" (* If *))
apply (tactic
- {* inst1_tac @{context} "P'" "Normal (\<lambda>s.. (\<lambda>Y s Z. True)\<down>=Val (the (locals s i)))" [] *})
+ \<open>inst1_tac @{context} "P'" "Normal (\<lambda>s.. (\<lambda>Y s Z. True)\<down>=Val (the (locals s i)))" []\<close>)
apply (tactic "ax_tac @{context} 1")
apply (rule conseq1)
apply (tactic "ax_tac @{context} 1")
@@ -266,7 +266,7 @@
apply (tactic "ax_tac @{context} 1")
prefer 2
apply (rule ax_subst_Var_allI)
-apply (tactic {* inst1_tac @{context} "P'" "\<lambda>b Y ba Z vf. \<lambda>Y (x,s) Z. x=None \<and> snd vf = snd (lvar i s)" [] *})
+apply (tactic \<open>inst1_tac @{context} "P'" "\<lambda>b Y ba Z vf. \<lambda>Y (x,s) Z. x=None \<and> snd vf = snd (lvar i s)" []\<close>)
apply (rule allI)
apply (rule_tac P' = "Normal P" and P = P for P in conseq1)
prefer 2
--- a/src/HOL/Bali/AxSem.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/AxSem.thy Sat Jan 02 18:48:45 2016 +0100
@@ -2,12 +2,12 @@
Author: David von Oheimb
*)
-subsection {* Axiomatic semantics of Java expressions and statements
+subsection \<open>Axiomatic semantics of Java expressions and statements
(see also Eval.thy)
- *}
+\<close>
theory AxSem imports Evaln TypeSafe begin
-text {*
+text \<open>
design issues:
\begin{itemize}
\item a strong version of validity for triples with premises, namely one that
@@ -34,9 +34,9 @@
\item all triples in a derivation are of the same type (due to weak
polymorphism)
\end{itemize}
-*}
+\<close>
-type_synonym res = vals --{* result entry *}
+type_synonym res = vals \<comment>\<open>result entry\<close>
abbreviation (input)
Val where "Val x == In1 x"
@@ -57,7 +57,7 @@
"\<lambda>Var:v . b" == "(\<lambda>v. b) \<circ> CONST the_In2"
"\<lambda>Vals:v. b" == "(\<lambda>v. b) \<circ> CONST the_In3"
- --{* relation on result values, state and auxiliary variables *}
+ \<comment>\<open>relation on result values, state and auxiliary variables\<close>
type_synonym 'a assn = "res \<Rightarrow> state \<Rightarrow> 'a \<Rightarrow> bool"
translations
(type) "'a assn" <= (type) "vals \<Rightarrow> state \<Rightarrow> 'a \<Rightarrow> bool"
@@ -464,8 +464,8 @@
declare split_paired_All [simp del] split_paired_Ex [simp del]
declare split_if [split del] split_if_asm [split del]
option.split [split del] option.split_asm [split del]
-setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *}
-setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\<close>
+setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
inductive
ax_derivs :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triples \<Rightarrow> bool" ("_,_|\<turnstile>_" [61,58,58] 57)
@@ -496,7 +496,7 @@
| Abrupt: "G,A\<turnstile>{P\<leftarrow>(undefined3 t) \<and>. Not \<circ> normal} t\<succ> {P}"
- --{* variables *}
+ \<comment>\<open>variables\<close>
| LVar: " G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Var (lvar vn s))} LVar vn=\<succ> {P}"
| FVar: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Q};
@@ -506,7 +506,7 @@
| AVar: "\<lbrakk>G,A\<turnstile>{Normal P} e1-\<succ> {Q};
\<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} e2-\<succ> {\<lambda>Val:i:. avar G i a ..; R}\<rbrakk> \<Longrightarrow>
G,A\<turnstile>{Normal P} e1.[e2]=\<succ> {R}"
- --{* expressions *}
+ \<comment>\<open>expressions\<close>
| NewC: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Alloc G (CInst C) Q}\<rbrakk> \<Longrightarrow>
G,A\<turnstile>{Normal P} NewC C-\<succ> {Q}"
@@ -569,7 +569,7 @@
\<Longrightarrow>
G,A\<turnstile>{Normal P} Body D c-\<succ> {R}"
- --{* expression lists *}
+ \<comment>\<open>expression lists\<close>
| Nil: "G,A\<turnstile>{Normal (P\<leftarrow>Vals [])} []\<doteq>\<succ> {P}"
@@ -577,7 +577,7 @@
\<forall>v. G,A\<turnstile>{Q\<leftarrow>Val v} es\<doteq>\<succ> {\<lambda>Vals:vs:. R\<leftarrow>Vals (v#vs)}\<rbrakk> \<Longrightarrow>
G,A\<turnstile>{Normal P} e#es\<doteq>\<succ> {R}"
- --{* statements *}
+ \<comment>\<open>statements\<close>
| Skip: "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit>)} .Skip. {P}"
@@ -627,10 +627,10 @@
.init c. {set_lvars l .; R}\<rbrakk> \<Longrightarrow>
G,A\<turnstile>{Normal (P \<and>. Not \<circ> initd C)} .Init C. {R}"
--- {* Some dummy rules for the intermediate terms @{text Callee},
-@{text InsInitE}, @{text InsInitV}, @{text FinA} only used by the smallstep
+\<comment> \<open>Some dummy rules for the intermediate terms \<open>Callee\<close>,
+\<open>InsInitE\<close>, \<open>InsInitV\<close>, \<open>FinA\<close> only used by the smallstep
semantics.
-*}
+\<close>
| InsInitV: " G,A\<turnstile>{Normal P} InsInitV c v=\<succ> {Q}"
| InsInitE: " G,A\<turnstile>{Normal P} InsInitE c e-\<succ> {Q}"
| Callee: " G,A\<turnstile>{Normal P} Callee l e-\<succ> {Q}"
@@ -672,7 +672,7 @@
(* 37 subgoals *)
prefer 18 (* Methd *)
apply (rule ax_derivs.Methd, drule spec, erule mp, fast)
-apply (tactic {* TRYALL (resolve_tac @{context} ((funpow 5 tl) @{thms ax_derivs.intros})) *})
+apply (tactic \<open>TRYALL (resolve_tac @{context} ((funpow 5 tl) @{thms ax_derivs.intros}))\<close>)
apply auto
done
@@ -692,9 +692,9 @@
apply (erule ax_derivs.induct)
(*42 subgoals*)
apply (tactic "ALLGOALS (strip_tac @{context})")
-apply (tactic {* ALLGOALS(REPEAT o (EVERY'[dresolve_tac @{context} @{thms subset_singletonD},
+apply (tactic \<open>ALLGOALS(REPEAT o (EVERY'[dresolve_tac @{context} @{thms subset_singletonD},
eresolve_tac @{context} [disjE],
- fast_tac (@{context} addSIs @{thms ax_derivs.empty})]))*})
+ fast_tac (@{context} addSIs @{thms ax_derivs.empty})]))\<close>)
apply (tactic "TRYALL (hyp_subst_tac @{context})")
apply (simp, rule ax_derivs.empty)
apply (drule subset_insertD)
@@ -704,8 +704,8 @@
apply (fast intro: ax_derivs.weaken)
apply (rule ax_derivs.conseq, clarify, tactic "smp_tac @{context} 3 1", blast(* unused *))
(*37 subgoals*)
-apply (tactic {* TRYALL (resolve_tac @{context} ((funpow 5 tl) @{thms ax_derivs.intros})
- THEN_ALL_NEW fast_tac @{context}) *})
+apply (tactic \<open>TRYALL (resolve_tac @{context} ((funpow 5 tl) @{thms ax_derivs.intros})
+ THEN_ALL_NEW fast_tac @{context})\<close>)
(*1 subgoal*)
apply (clarsimp simp add: subset_mtriples_iff)
apply (rule ax_derivs.Methd)
@@ -719,16 +719,16 @@
subsubsection "rules derived from conseq"
-text {* In the following rules we often have to give some type annotations like:
+text \<open>In the following rules we often have to give some type annotations like:
@{term "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q}"}.
Given only the term above without annotations, Isabelle would infer a more
general type were we could have
different types of auxiliary variables in the assumption set (@{term A}) and
in the triple itself (@{term P} and @{term Q}). But
-@{text "ax_derivs.Methd"} enforces the same type in the inductive definition of
+\<open>ax_derivs.Methd\<close> enforces the same type in the inductive definition of
the derivation. So we have to restrict the types to be able to apply the
rules.
-*}
+\<close>
lemma conseq12: "\<lbrakk>G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'};
\<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>Y' s'. (\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow>
Q Y' s' Z)\<rbrakk>
@@ -738,7 +738,7 @@
apply blast
done
--- {* Nice variant, since it is so symmetric we might be able to memorise it. *}
+\<comment> \<open>Nice variant, since it is so symmetric we might be able to memorise it.\<close>
lemma conseq12': "\<lbrakk>G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'}; \<forall>s Y' s'.
(\<forall>Y Z. P' Y s Z \<longrightarrow> Q' Y' s' Z) \<longrightarrow>
(\<forall>Y Z. P Y s Z \<longrightarrow> Q Y' s' Z)\<rbrakk>
@@ -1008,7 +1008,7 @@
apply (auto simp add: type_ok_def)
done
-ML {* ML_Thms.bind_thms ("ax_Abrupts", sum3_instantiate @{context} @{thm ax_derivs.Abrupt}) *}
+ML \<open>ML_Thms.bind_thms ("ax_Abrupts", sum3_instantiate @{context} @{thm ax_derivs.Abrupt})\<close>
declare ax_Abrupts [intro!]
lemmas ax_Normal_cases = ax_cases [of _ _ _ normal]
--- a/src/HOL/Bali/AxSound.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/AxSound.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,9 +1,9 @@
(* Title: HOL/Bali/AxSound.thy
Author: David von Oheimb and Norbert Schirmer
*)
-subsection {* Soundness proof for Axiomatic semantics of Java expressions and
+subsection \<open>Soundness proof for Axiomatic semantics of Java expressions and
statements
- *}
+\<close>
theory AxSound imports AxSem begin
@@ -19,13 +19,13 @@
\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>A)) \<longrightarrow>
(\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s') \<longrightarrow> Q Y' s' Z \<and> s'\<Colon>\<preceq>(G,L)))))"
-text {* This definition differs from the ordinary @{text triple_valid_def}
+text \<open>This definition differs from the ordinary \<open>triple_valid_def\<close>
manly in the conclusion: We also ensures conformance of the result state. So
we don't have to apply the type soundness lemma all the time during
induction. This definition is only introduced for the soundness
proof of the axiomatic semantics, in the end we will conclude to
the ordinary definition.
-*}
+\<close>
definition
ax_valids2 :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triples \<Rightarrow> bool" ("_,_|\<Turnstile>\<Colon>_" [61,58,58] 57)
@@ -126,8 +126,8 @@
qed
next
case (Suc m)
- note hyp = `\<forall>t\<in>A. G\<Turnstile>m\<Colon>t \<Longrightarrow> \<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>m\<Colon>t`
- note prem = `\<forall>t\<in>A. G\<Turnstile>Suc m\<Colon>t`
+ note hyp = \<open>\<forall>t\<in>A. G\<Turnstile>m\<Colon>t \<Longrightarrow> \<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>m\<Colon>t\<close>
+ note prem = \<open>\<forall>t\<in>A. G\<Turnstile>Suc m\<Colon>t\<close>
show "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>Suc m\<Colon>t"
proof -
{
@@ -350,9 +350,9 @@
by (simp add: ax_valids2_def triple_valid2_def2)
next
case (insert A t ts)
- note valid_t = `G,A|\<Turnstile>\<Colon>{t}`
+ note valid_t = \<open>G,A|\<Turnstile>\<Colon>{t}\<close>
moreover
- note valid_ts = `G,A|\<Turnstile>\<Colon>ts`
+ note valid_ts = \<open>G,A|\<Turnstile>\<Colon>ts\<close>
{
fix n assume valid_A: "\<forall>t\<in>A. G\<Turnstile>n\<Colon>t"
have "G\<Turnstile>n\<Colon>t" and "\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t"
@@ -370,21 +370,21 @@
by (unfold ax_valids2_def) blast
next
case (asm ts A)
- from `ts \<subseteq> A`
+ from \<open>ts \<subseteq> A\<close>
show "G,A|\<Turnstile>\<Colon>ts"
by (auto simp add: ax_valids2_def triple_valid2_def)
next
case (weaken A ts' ts)
- note `G,A|\<Turnstile>\<Colon>ts'`
- moreover note `ts \<subseteq> ts'`
+ note \<open>G,A|\<Turnstile>\<Colon>ts'\<close>
+ moreover note \<open>ts \<subseteq> ts'\<close>
ultimately show "G,A|\<Turnstile>\<Colon>ts"
by (unfold ax_valids2_def triple_valid2_def) blast
next
case (conseq P A t Q)
- note con = `\<forall>Y s Z. P Y s Z \<longrightarrow>
+ note con = \<open>\<forall>Y s Z. P Y s Z \<longrightarrow>
(\<exists>P' Q'.
(G,A\<turnstile>{P'} t\<succ> {Q'} \<and> G,A|\<Turnstile>\<Colon>{ {P'} t\<succ> {Q'} }) \<and>
- (\<forall>Y' s'. (\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow> Q Y' s' Z))`
+ (\<forall>Y' s'. (\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow> Q Y' s' Z))\<close>
show "G,A|\<Turnstile>\<Colon>{ {P} t\<succ> {Q} }"
proof (rule validI)
fix n s0 L accC T C v s1 Y Z
@@ -469,8 +469,8 @@
qed
next
case (FVar A P statDeclC Q e stat fn R accC)
- note valid_init = `G,A|\<Turnstile>\<Colon>{ {Normal P} .Init statDeclC. {Q} }`
- note valid_e = `G,A|\<Turnstile>\<Colon>{ {Q} e-\<succ> {\<lambda>Val:a:. fvar statDeclC stat fn a ..; R} }`
+ note valid_init = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} .Init statDeclC. {Q} }\<close>
+ note valid_e = \<open>G,A|\<Turnstile>\<Colon>{ {Q} e-\<succ> {\<lambda>Val:a:. fvar statDeclC stat fn a ..; R} }\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} {accC,statDeclC,stat}e..fn=\<succ> {R} }"
proof (rule valid_var_NormalI)
fix n s0 L accC' T V vf s3 Y Z
@@ -564,7 +564,7 @@
qed
next
case (AVar A P e1 Q e2 R)
- note valid_e1 = `G,A|\<Turnstile>\<Colon>{ {Normal P} e1-\<succ> {Q} }`
+ note valid_e1 = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} e1-\<succ> {Q} }\<close>
have valid_e2: "\<And> a. G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>In1 a} e2-\<succ> {\<lambda>Val:i:. avar G i a ..; R} }"
using AVar.hyps by simp
show "G,A|\<Turnstile>\<Colon>{ {Normal P} e1.[e2]=\<succ> {R} }"
@@ -628,7 +628,7 @@
qed
next
case (NewC A P C Q)
- note valid_init = `G,A|\<Turnstile>\<Colon>{ {Normal P} .Init C. {Alloc G (CInst C) Q} }`
+ note valid_init = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} .Init C. {Alloc G (CInst C) Q} }\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} NewC C-\<succ> {Q} }"
proof (rule valid_expr_NormalI)
fix n s0 L accC T E v s2 Y Z
@@ -668,9 +668,9 @@
qed
next
case (NewA A P T Q e R)
- note valid_init = `G,A|\<Turnstile>\<Colon>{ {Normal P} .init_comp_ty T. {Q} }`
- note valid_e = `G,A|\<Turnstile>\<Colon>{ {Q} e-\<succ> {\<lambda>Val:i:. abupd (check_neg i) .;
- Alloc G (Arr T (the_Intg i)) R}}`
+ note valid_init = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} .init_comp_ty T. {Q} }\<close>
+ note valid_e = \<open>G,A|\<Turnstile>\<Colon>{ {Q} e-\<succ> {\<lambda>Val:i:. abupd (check_neg i) .;
+ Alloc G (Arr T (the_Intg i)) R}}\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} New T[e]-\<succ> {R} }"
proof (rule valid_expr_NormalI)
fix n s0 L accC arrT E v s3 Y Z
@@ -741,9 +741,9 @@
qed
next
case (Cast A P e T Q)
- note valid_e = `G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ>
+ note valid_e = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ>
{\<lambda>Val:v:. \<lambda>s.. abupd (raise_if (\<not> G,s\<turnstile>v fits T) ClassCast) .;
- Q\<leftarrow>In1 v} }`
+ Q\<leftarrow>In1 v} }\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} Cast T e-\<succ> {Q} }"
proof (rule valid_expr_NormalI)
fix n s0 L accC castT E v s2 Y Z
@@ -972,7 +972,7 @@
qed
next
case (Acc A P var Q)
- note valid_var = `G,A|\<Turnstile>\<Colon>{ {Normal P} var=\<succ> {\<lambda>Var:(v, f):. Q\<leftarrow>In1 v} }`
+ note valid_var = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} var=\<succ> {\<lambda>Var:(v, f):. Q\<leftarrow>In1 v} }\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} Acc var-\<succ> {Q} }"
proof (rule valid_expr_NormalI)
fix n s0 L accC T E v s1 Y Z
@@ -1008,7 +1008,7 @@
qed
next
case (Ass A P var Q e R)
- note valid_var = `G,A|\<Turnstile>\<Colon>{ {Normal P} var=\<succ> {Q} }`
+ note valid_var = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} var=\<succ> {Q} }\<close>
have valid_e: "\<And> vf.
G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>In2 vf} e-\<succ> {\<lambda>Val:v:. assign (snd vf) v .; R} }"
using Ass.hyps by simp
@@ -1120,7 +1120,7 @@
qed
next
case (Cond A P e0 P' e1 e2 Q)
- note valid_e0 = `G,A|\<Turnstile>\<Colon>{ {Normal P} e0-\<succ> {P'} }`
+ note valid_e0 = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} e0-\<succ> {P'} }\<close>
have valid_then_else:"\<And> b. G,A|\<Turnstile>\<Colon>{ {P'\<leftarrow>=b} (if b then e1 else e2)-\<succ> {Q} }"
using Cond.hyps by simp
show "G,A|\<Turnstile>\<Colon>{ {Normal P} e0 ? e1 : e2-\<succ> {Q} }"
@@ -1210,7 +1210,7 @@
qed
next
case (Call A P e Q args R mode statT mn pTs' S accC')
- note valid_e = `G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q} }`
+ note valid_e = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q} }\<close>
have valid_args: "\<And> a. G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>In1 a} args\<doteq>\<succ> {R a} }"
using Call.hyps by simp
have valid_methd: "\<And> a vs invC declC l.
@@ -1594,14 +1594,14 @@
qed
next
case (Methd A P Q ms)
- note valid_body = `G,A \<union> {{P} Methd-\<succ> {Q} | ms}|\<Turnstile>\<Colon>{{P} body G-\<succ> {Q} | ms}`
+ note valid_body = \<open>G,A \<union> {{P} Methd-\<succ> {Q} | ms}|\<Turnstile>\<Colon>{{P} body G-\<succ> {Q} | ms}\<close>
show "G,A|\<Turnstile>\<Colon>{{P} Methd-\<succ> {Q} | ms}"
by (rule Methd_sound) (rule Methd.hyps)
next
case (Body A P D Q c R)
- note valid_init = `G,A|\<Turnstile>\<Colon>{ {Normal P} .Init D. {Q} }`
- note valid_c = `G,A|\<Turnstile>\<Colon>{ {Q} .c.
- {\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>In1 (the (locals s Result))} }`
+ note valid_init = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} .Init D. {Q} }\<close>
+ note valid_c = \<open>G,A|\<Turnstile>\<Colon>{ {Q} .c.
+ {\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>In1 (the (locals s Result))} }\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} Body D c-\<succ> {R} }"
proof (rule valid_expr_NormalI)
fix n s0 L accC T E v s4 Y Z
@@ -1692,7 +1692,7 @@
qed
next
case (Cons A P e Q es R)
- note valid_e = `G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q} }`
+ note valid_e = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q} }\<close>
have valid_es: "\<And> v. G,A|\<Turnstile>\<Colon>{ {Q\<leftarrow>\<lfloor>v\<rfloor>\<^sub>e} es\<doteq>\<succ> {\<lambda>Vals:vs:. R\<leftarrow>\<lfloor>(v # vs)\<rfloor>\<^sub>l} }"
using Cons.hyps by simp
show "G,A|\<Turnstile>\<Colon>{ {Normal P} e # es\<doteq>\<succ> {R} }"
@@ -1774,7 +1774,7 @@
qed
next
case (Expr A P e Q)
- note valid_e = `G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q\<leftarrow>\<diamondsuit>} }`
+ note valid_e = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {Q\<leftarrow>\<diamondsuit>} }\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} .Expr e. {Q} }"
proof (rule valid_stmt_NormalI)
fix n s0 L accC C s1 Y Z
@@ -1804,7 +1804,7 @@
qed
next
case (Lab A P c l Q)
- note valid_c = `G,A|\<Turnstile>\<Colon>{ {Normal P} .c. {abupd (absorb l) .; Q} }`
+ note valid_c = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} .c. {abupd (absorb l) .; Q} }\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} .l\<bullet> c. {Q} }"
proof (rule valid_stmt_NormalI)
fix n s0 L accC C s2 Y Z
@@ -1841,8 +1841,8 @@
qed
next
case (Comp A P c1 Q c2 R)
- note valid_c1 = `G,A|\<Turnstile>\<Colon>{ {Normal P} .c1. {Q} }`
- note valid_c2 = `G,A|\<Turnstile>\<Colon>{ {Q} .c2. {R} }`
+ note valid_c1 = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} .c1. {Q} }\<close>
+ note valid_c2 = \<open>G,A|\<Turnstile>\<Colon>{ {Q} .c2. {R} }\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} .c1;; c2. {R} }"
proof (rule valid_stmt_NormalI)
fix n s0 L accC C s2 Y Z
@@ -1900,7 +1900,7 @@
qed
next
case (If A P e P' c1 c2 Q)
- note valid_e = `G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {P'} }`
+ note valid_e = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {P'} }\<close>
have valid_then_else: "\<And> b. G,A|\<Turnstile>\<Colon>{ {P'\<leftarrow>=b} .(if b then c1 else c2). {Q} }"
using If.hyps by simp
show "G,A|\<Turnstile>\<Colon>{ {Normal P} .If(e) c1 Else c2. {Q} }"
@@ -1977,10 +1977,10 @@
qed
next
case (Loop A P e P' c l)
- note valid_e = `G,A|\<Turnstile>\<Colon>{ {P} e-\<succ> {P'} }`
- note valid_c = `G,A|\<Turnstile>\<Colon>{ {Normal (P'\<leftarrow>=True)}
+ note valid_e = \<open>G,A|\<Turnstile>\<Colon>{ {P} e-\<succ> {P'} }\<close>
+ note valid_c = \<open>G,A|\<Turnstile>\<Colon>{ {Normal (P'\<leftarrow>=True)}
.c.
- {abupd (absorb (Cont l)) .; P} }`
+ {abupd (absorb (Cont l)) .; P} }\<close>
show "G,A|\<Turnstile>\<Colon>{ {P} .l\<bullet> While(e) c. {P'\<leftarrow>=False\<down>=\<diamondsuit>} }"
proof (rule valid_stmtI)
fix n s0 L accC C s3 Y Z
@@ -1993,7 +1993,7 @@
assume P: "P Y s0 Z"
show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
proof -
- --{* From the given hypothesises @{text valid_e} and @{text valid_c}
+ \<comment>\<open>From the given hypothesises \<open>valid_e\<close> and \<open>valid_c\<close>
we can only reach the state after unfolding the loop once, i.e.
@{term "P \<diamondsuit> s2 Z"}, where @{term s2} is the state after executing
@{term c}. To gain validity of the further execution of while, to
@@ -2001,9 +2001,9 @@
a hypothesis about the subsequent unfoldings (the whole loop again),
too. We can achieve this, by performing induction on the
evaluation relation, with all
- the necessary preconditions to apply @{text valid_e} and
- @{text valid_c} in the goal.
- *}
+ the necessary preconditions to apply \<open>valid_e\<close> and
+ \<open>valid_c\<close> in the goal.
+\<close>
{
fix t s s' v
assume "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s')"
@@ -2015,11 +2015,11 @@
(is "PROP ?Hyp n t s v s'")
proof (induct)
case (Loop s0' e' b n' s1' c' s2' l' s3' Y' T E)
- note while = `(\<langle>l'\<bullet> While(e') c'\<rangle>\<^sub>s::term) = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
+ note while = \<open>(\<langle>l'\<bullet> While(e') c'\<rangle>\<^sub>s::term) = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<close>
hence eqs: "l'=l" "e'=e" "c'=c" by simp_all
- note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
- note P = `P Y' (Norm s0') Z`
- note conf_s0' = `Norm s0'\<Colon>\<preceq>(G, L)`
+ note valid_A = \<open>\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t\<close>
+ note P = \<open>P Y' (Norm s0') Z\<close>
+ note conf_s0' = \<open>Norm s0'\<Colon>\<preceq>(G, L)\<close>
have wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<Colon>T"
using Loop.prems eqs by simp
have da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
@@ -2168,10 +2168,10 @@
qed
next
case (Abrupt abr s t' n' Y' T E)
- note t' = `t' = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
- note conf = `(Some abr, s)\<Colon>\<preceq>(G, L)`
- note P = `P Y' (Some abr, s) Z`
- note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
+ note t' = \<open>t' = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<close>
+ note conf = \<open>(Some abr, s)\<Colon>\<preceq>(G, L)\<close>
+ note P = \<open>P Y' (Some abr, s) Z\<close>
+ note valid_A = \<open>\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t\<close>
show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) (undefined3 t') (Some abr, s) Z"
proof -
have eval_e:
@@ -2234,7 +2234,7 @@
qed
next
case (Throw A P e Q)
- note valid_e = `G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>} }`
+ note valid_e = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} e-\<succ> {\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>} }\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} .Throw e. {Q} }"
proof (rule valid_stmt_NormalI)
fix n s0 L accC C s2 Y Z
@@ -2272,11 +2272,11 @@
qed
next
case (Try A P c1 Q C vn c2 R)
- note valid_c1 = `G,A|\<Turnstile>\<Colon>{ {Normal P} .c1. {SXAlloc G Q} }`
- note valid_c2 = `G,A|\<Turnstile>\<Colon>{ {Q \<and>. (\<lambda>s. G,s\<turnstile>catch C) ;. new_xcpt_var vn}
+ note valid_c1 = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} .c1. {SXAlloc G Q} }\<close>
+ note valid_c2 = \<open>G,A|\<Turnstile>\<Colon>{ {Q \<and>. (\<lambda>s. G,s\<turnstile>catch C) ;. new_xcpt_var vn}
.c2.
- {R} }`
- note Q_R = `(Q \<and>. (\<lambda>s. \<not> G,s\<turnstile>catch C)) \<Rightarrow> R`
+ {R} }\<close>
+ note Q_R = \<open>(Q \<and>. (\<lambda>s. \<not> G,s\<turnstile>catch C)) \<Rightarrow> R\<close>
show "G,A|\<Turnstile>\<Colon>{ {Normal P} .Try c1 Catch(C vn) c2. {R} }"
proof (rule valid_stmt_NormalI)
fix n s0 L accC E s3 Y Z
@@ -2404,7 +2404,7 @@
qed
next
case (Fin A P c1 Q c2 R)
- note valid_c1 = `G,A|\<Turnstile>\<Colon>{ {Normal P} .c1. {Q} }`
+ note valid_c1 = \<open>G,A|\<Turnstile>\<Colon>{ {Normal P} .c1. {Q} }\<close>
have valid_c2: "\<And> abr. G,A|\<Turnstile>\<Colon>{ {Q \<and>. (\<lambda>s. abr = fst s) ;. abupd (\<lambda>x. None)}
.c2.
{abupd (abrupt_if (abr \<noteq> None) abr) .; R} }"
@@ -2501,11 +2501,11 @@
qed
next
case (Init C c A P Q R)
- note c = `the (class G C) = c`
+ note c = \<open>the (class G C) = c\<close>
note valid_super =
- `G,A|\<Turnstile>\<Colon>{ {Normal (P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
+ \<open>G,A|\<Turnstile>\<Colon>{ {Normal (P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
.(if C = Object then Skip else Init (super c)).
- {Q} }`
+ {Q} }\<close>
have valid_init:
"\<And> l. G,A|\<Turnstile>\<Colon>{ {Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty}
.init c.
--- a/src/HOL/Bali/Basis.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Basis.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,7 +1,7 @@
(* Title: HOL/Bali/Basis.thy
Author: David von Oheimb
*)
-subsection {* Definitions extending HOL as logical basis of Bali *}
+subsection \<open>Definitions extending HOL as logical basis of Bali\<close>
theory Basis
imports Main "~~/src/HOL/Library/Old_Recdef"
@@ -9,10 +9,10 @@
subsubsection "misc"
-ML {* fun strip_tac ctxt i = REPEAT (resolve_tac ctxt [impI, allI] i) *}
+ML \<open>fun strip_tac ctxt i = REPEAT (resolve_tac ctxt [impI, allI] i)\<close>
declare split_if_asm [split] option.split [split] option.split_asm [split]
-setup {* map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac)) *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac))\<close>
declare if_weak_cong [cong del] option.case_cong_weak [cong del]
declare length_Suc_conv [iff]
@@ -176,13 +176,13 @@
abbreviation the_In1r :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'ar"
where "the_In1r \<equiv> the_Inr \<circ> the_In1"
-ML {*
+ML \<open>
fun sum3_instantiate ctxt thm =
map (fn s =>
simplify (ctxt delsimps @{thms not_None_eq})
(Rule_Insts.read_instantiate ctxt [((("t", 0), Position.none), "In" ^ s ^ " x")] ["x"] thm))
["1l","2","3","1r"]
-*}
+\<close>
(* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *)
@@ -203,7 +203,7 @@
subsubsection "Special map update"
-text{* Deemed too special for theory Map. *}
+text\<open>Deemed too special for theory Map.\<close>
definition chg_map :: "('b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> ('a \<rightharpoonup> 'b) \<Rightarrow> ('a \<rightharpoonup> 'b)"
where "chg_map f a m = (case m a of None \<Rightarrow> m | Some b \<Rightarrow> m(a\<mapsto>f b))"
@@ -255,7 +255,7 @@
definition lsplit :: "[['a, 'a list] \<Rightarrow> 'b, 'a list] \<Rightarrow> 'b"
where "lsplit = (\<lambda>f l. f (hd l) (tl l))"
-text {* list patterns -- extends pre-defined type "pttrn" used in abstractions *}
+text \<open>list patterns -- extends pre-defined type "pttrn" used in abstractions\<close>
syntax
"_lpttrn" :: "[pttrn, pttrn] \<Rightarrow> pttrn" ("_#/_" [901,900] 900)
translations
--- a/src/HOL/Bali/Conform.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Conform.thy Sat Jan 02 18:48:45 2016 +0100
@@ -2,11 +2,11 @@
Author: David von Oheimb
*)
-subsection {* Conformance notions for the type soundness proof for Java *}
+subsection \<open>Conformance notions for the type soundness proof for Java\<close>
theory Conform imports State begin
-text {*
+text \<open>
design issues:
\begin{itemize}
\item lconf allows for (arbitrary) inaccessible values
@@ -14,7 +14,7 @@
objects on the heap are indeed existing classes. Yet this can be
inferred for all referenced objs.
\end{itemize}
-*}
+\<close>
type_synonym env' = "prog \<times> (lname, ty) table" (* same as env of WellType.thy *)
@@ -25,12 +25,12 @@
definition gext :: "st \<Rightarrow> st \<Rightarrow> bool" ("_\<le>|_" [71,71] 70) where
"s\<le>|s' \<equiv> \<forall>r. \<forall>obj\<in>globs s r: \<exists>obj'\<in>globs s' r: tag obj'= tag obj"
-text {* For the the proof of type soundness we will need the
+text \<open>For the the proof of type soundness we will need the
property that during execution, objects are not lost and moreover retain the
values of their tags. So the object store grows conservatively. Note that if
we considered garbage collection, we would have to restrict this property to
accessible objects.
-*}
+\<close>
lemma gext_objD:
"\<lbrakk>s\<le>|s'; globs s r = Some obj\<rbrakk>
@@ -250,7 +250,7 @@
subsubsection "weak value list conformance"
-text {* Only if the value is defined it has to conform to its type.
+text \<open>Only if the value is defined it has to conform to its type.
This is the contribution of the definite assignment analysis to
the notion of conformance. The definite assignment analysis ensures
that the program only attempts to access local variables that
@@ -258,7 +258,7 @@
So conformance must only ensure that the
defined values are of the right type, and not also that the value
is defined.
-*}
+\<close>
definition
--- a/src/HOL/Bali/Decl.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Decl.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,15 +1,15 @@
(* Title: HOL/Bali/Decl.thy
Author: David von Oheimb and Norbert Schirmer
*)
-subsection {* Field, method, interface, and class declarations, whole Java programs
-*}
+subsection \<open>Field, method, interface, and class declarations, whole Java programs
+\<close>
theory Decl
imports Term Table
(** order is significant, because of clash for "var" **)
begin
-text {*
+text \<open>
improvements:
\begin{itemize}
\item clarification and correction of some aspects of the package/access concept
@@ -36,20 +36,20 @@
\item no main method
\end{itemize}
-*}
+\<close>
-subsection {* Modifier*}
+subsection \<open>Modifier\<close>
-subsubsection {* Access modifier *}
+subsubsection \<open>Access modifier\<close>
datatype acc_modi (* access modifier *)
= Private | Package | Protected | Public
-text {*
+text \<open>
We can define a linear order for the access modifiers. With Private yielding the
most restrictive access and public the most liberal access policy:
Private < Package < Protected < Public
-*}
+\<close>
instantiation acc_modi :: linorder
begin
@@ -70,14 +70,14 @@
fix x y z::acc_modi
show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split)
- show "x \<le> x" -- reflexivity
+ show "x \<le> x" \<comment> reflexivity
by (auto simp add: le_acc_def)
{
- assume "x \<le> y" "y \<le> z" -- transitivity
+ assume "x \<le> y" "y \<le> z" \<comment> transitivity
then show "x \<le> z"
by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split)
next
- assume "x \<le> y" "y \<le> x" -- antisymmetry
+ assume "x \<le> y" "y \<le> x" \<comment> antisymmetry
moreover have "\<forall> x y. x < (y::acc_modi) \<and> y < x \<longrightarrow> False"
by (auto simp add: less_acc_def split add: acc_modi.split)
ultimately show "x = y" by (unfold le_acc_def) iprover
@@ -137,11 +137,11 @@
using assms by (auto dest: acc_modi_Package_le)
-subsubsection {* Static Modifier *}
+subsubsection \<open>Static Modifier\<close>
type_synonym stat_modi = bool (* modifier: static *)
-subsection {* Declaration (base "class" for member,interface and class
- declarations *}
+subsection \<open>Declaration (base "class" for member,interface and class
+ declarations\<close>
record decl =
access :: acc_modi
@@ -150,7 +150,7 @@
(type) "decl" <= (type) "\<lparr>access::acc_modi\<rparr>"
(type) "decl" <= (type) "\<lparr>access::acc_modi,\<dots>::'a\<rparr>"
-subsection {* Member (field or method)*}
+subsection \<open>Member (field or method)\<close>
record member = decl +
static :: stat_modi
@@ -158,7 +158,7 @@
(type) "member" <= (type) "\<lparr>access::acc_modi,static::bool\<rparr>"
(type) "member" <= (type) "\<lparr>access::acc_modi,static::bool,\<dots>::'a\<rparr>"
-subsection {* Field *}
+subsection \<open>Field\<close>
record field = member +
type :: ty
@@ -173,7 +173,7 @@
translations
(type) "fdecl" <= (type) "vname \<times> field"
-subsection {* Method *}
+subsection \<open>Method\<close>
record mhead = member + (* method head (excluding signature) *)
pars ::"vname list" (* parameter names *)
@@ -219,9 +219,9 @@
lemma resT_mhead [simp]:"resT (mhead m) = resT m"
by (simp add: mhead_def)
-text {* To be able to talk uniformaly about field and method declarations we
+text \<open>To be able to talk uniformaly about field and method declarations we
introduce the notion of a member declaration (e.g. useful to define
-accessiblity ) *}
+accessiblity )\<close>
datatype memberdecl = fdecl fdecl | mdecl mdecl
@@ -293,16 +293,16 @@
by (simp add: is_method_def)
-subsection {* Interface *}
+subsection \<open>Interface\<close>
-record ibody = decl + --{* interface body *}
- imethods :: "(sig \<times> mhead) list" --{* method heads *}
+record ibody = decl + \<comment>\<open>interface body\<close>
+ imethods :: "(sig \<times> mhead) list" \<comment>\<open>method heads\<close>
-record iface = ibody + --{* interface *}
- isuperIfs:: "qtname list" --{* superinterface list *}
+record iface = ibody + \<comment>\<open>interface\<close>
+ isuperIfs:: "qtname list" \<comment>\<open>superinterface list\<close>
type_synonym
- idecl --{* interface declaration, cf. 9.1 *}
+ idecl \<comment>\<open>interface declaration, cf. 9.1\<close>
= "qtname \<times> iface"
translations
@@ -324,17 +324,17 @@
lemma imethods_ibody [simp]: "(imethods (ibody i)) = imethods i"
by (simp add: ibody_def)
-subsection {* Class *}
-record cbody = decl + --{* class body *}
+subsection \<open>Class\<close>
+record cbody = decl + \<comment>\<open>class body\<close>
cfields:: "fdecl list"
methods:: "mdecl list"
- init :: "stmt" --{* initializer *}
+ init :: "stmt" \<comment>\<open>initializer\<close>
-record "class" = cbody + --{* class *}
- super :: "qtname" --{* superclass *}
- superIfs:: "qtname list" --{* implemented interfaces *}
+record "class" = cbody + \<comment>\<open>class\<close>
+ super :: "qtname" \<comment>\<open>superclass\<close>
+ superIfs:: "qtname list" \<comment>\<open>implemented interfaces\<close>
type_synonym
- cdecl --{* class declaration, cf. 8.1 *}
+ cdecl \<comment>\<open>class declaration, cf. 8.1\<close>
= "qtname \<times> class"
translations
@@ -370,16 +370,16 @@
subsubsection "standard classes"
consts
- Object_mdecls :: "mdecl list" --{* methods of Object *}
- SXcpt_mdecls :: "mdecl list" --{* methods of SXcpts *}
+ Object_mdecls :: "mdecl list" \<comment>\<open>methods of Object\<close>
+ SXcpt_mdecls :: "mdecl list" \<comment>\<open>methods of SXcpts\<close>
definition
- ObjectC :: "cdecl" --{* declaration of root class *} where
+ ObjectC :: "cdecl" \<comment>\<open>declaration of root class\<close> where
"ObjectC = (Object,\<lparr>access=Public,cfields=[],methods=Object_mdecls,
init=Skip,super=undefined,superIfs=[]\<rparr>)"
definition
- SXcptC ::"xname \<Rightarrow> cdecl" --{* declarations of throwable classes *} where
+ SXcptC ::"xname \<Rightarrow> cdecl" \<comment>\<open>declarations of throwable classes\<close> where
"SXcptC xn = (SXcpt xn,\<lparr>access=Public,cfields=[],methods=SXcpt_mdecls,
init=Skip,
super=if xn = Throwable then Object
@@ -448,11 +448,11 @@
subsubsection "subinterface and subclass relation, in anticipation of TypeRel.thy"
definition
- subint1 :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{* direct subinterface *}
+ subint1 :: "prog \<Rightarrow> (qtname \<times> qtname) set" \<comment>\<open>direct subinterface\<close>
where "subint1 G = {(I,J). \<exists>i\<in>iface G I: J\<in>set (isuperIfs i)}"
definition
- subcls1 :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{* direct subclass *}
+ subcls1 :: "prog \<Rightarrow> (qtname \<times> qtname) set" \<comment>\<open>direct subclass\<close>
where "subcls1 G = {(C,D). C\<noteq>Object \<and> (\<exists>c\<in>class G C: super c = D)}"
abbreviation
@@ -815,7 +815,7 @@
definition
imethds :: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables" where
- --{* methods of an interface, with overriding and inheritance, cf. 9.2 *}
+ \<comment>\<open>methods of an interface, with overriding and inheritance, cf. 9.2\<close>
"imethds G I = iface_rec G I
(\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus>
(set_option \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
--- a/src/HOL/Bali/DeclConcepts.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/DeclConcepts.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,8 +1,8 @@
(* Title: HOL/Bali/DeclConcepts.thy
Author: Norbert Schirmer
*)
-subsection {* Advanced concepts on Java declarations like overriding, inheritance,
-dynamic method lookup*}
+subsection \<open>Advanced concepts on Java declarations like overriding, inheritance,
+dynamic method lookup\<close>
theory DeclConcepts imports TypeRel begin
@@ -16,10 +16,10 @@
| Some c \<Rightarrow> access c = Public)"
subsection "accessibility of types (cf. 6.6.1)"
-text {*
+text \<open>
Primitive types are always accessible, interfaces and classes are accessible
in their package or if they are defined public, an array type is accessible if
-its element type is accessible *}
+its element type is accessible\<close>
primrec
accessible_in :: "prog \<Rightarrow> ty \<Rightarrow> pname \<Rightarrow> bool" ("_ \<turnstile> _ accessible'_in _" [61,61,61] 60) and
@@ -71,16 +71,16 @@
by (simp add: is_acc_reftype_def)
subsection "accessibility of members"
-text {*
+text \<open>
The accessibility of members is more involved as the accessibility of types.
We have to distinguish several cases to model the different effects of
accessibility during inheritance, overriding and ordinary member access
-*}
+\<close>
-subsubsection {* Various technical conversion and selection functions *}
+subsubsection \<open>Various technical conversion and selection functions\<close>
-text {* overloaded selector @{text accmodi} to select the access modifier
-out of various HOL types *}
+text \<open>overloaded selector \<open>accmodi\<close> to select the access modifier
+out of various HOL types\<close>
class has_accmodi =
fixes accmodi:: "'a \<Rightarrow> acc_modi"
@@ -144,8 +144,8 @@
"accmodi (mdecl m) = accmodi m"
by (simp add: memberdecl_acc_modi_def)
-text {* overloaded selector @{text declclass} to select the declaring class
-out of various HOL types *}
+text \<open>overloaded selector \<open>declclass\<close> to select the declaring class
+out of various HOL types\<close>
class has_declclass =
fixes declclass:: "'a \<Rightarrow> qtname"
@@ -180,8 +180,8 @@
lemma pair_declclass_simp[simp]: "declclass (c,x) = declclass c"
by (simp add: pair_declclass_def)
-text {* overloaded selector @{text is_static} to select the static modifier
-out of various HOL types *}
+text \<open>overloaded selector \<open>is_static\<close> to select the static modifier
+out of various HOL types\<close>
class has_static =
fixes is_static :: "'a \<Rightarrow> bool"
@@ -246,32 +246,32 @@
lemma mhead_static_simp [simp]: "is_static (mhead m) = is_static m"
by (cases m) (simp add: mhead_def member_is_static_simp)
--- {* some mnemotic selectors for various pairs *}
+\<comment> \<open>some mnemotic selectors for various pairs\<close>
definition
decliface :: "qtname \<times> 'a decl_scheme \<Rightarrow> qtname" where
- "decliface = fst" --{* get the interface component *}
+ "decliface = fst" \<comment>\<open>get the interface component\<close>
definition
mbr :: "qtname \<times> memberdecl \<Rightarrow> memberdecl" where
- "mbr = snd" --{* get the memberdecl component *}
+ "mbr = snd" \<comment>\<open>get the memberdecl component\<close>
definition
mthd :: "'b \<times> 'a \<Rightarrow> 'a" where
- "mthd = snd" --{* get the method component *}
- --{* also used for mdecl, mhead *}
+ "mthd = snd" \<comment>\<open>get the method component\<close>
+ \<comment>\<open>also used for mdecl, mhead\<close>
definition
fld :: "'b \<times> 'a decl_scheme \<Rightarrow> 'a decl_scheme" where
- "fld = snd" --{* get the field component *}
- --{* also used for @{text "((vname \<times> qtname)\<times> field)"} *}
+ "fld = snd" \<comment>\<open>get the field component\<close>
+ \<comment>\<open>also used for \<open>((vname \<times> qtname)\<times> field)\<close>\<close>
--- {* some mnemotic selectors for @{text "(vname \<times> qtname)"} *}
+\<comment> \<open>some mnemotic selectors for \<open>(vname \<times> qtname)\<close>\<close>
definition
fname:: "vname \<times> 'a \<Rightarrow> vname"
where "fname = fst"
- --{* also used for fdecl *}
+ \<comment>\<open>also used for fdecl\<close>
definition
declclassf:: "(vname \<times> qtname) \<Rightarrow> qtname"
@@ -326,7 +326,7 @@
lemma declclassf_simp[simp]:"declclassf (n,c) = c"
by (simp add: declclassf_def)
- --{* some mnemotic selectors for @{text "(vname \<times> qtname)"} *}
+ \<comment>\<open>some mnemotic selectors for \<open>(vname \<times> qtname)\<close>\<close>
definition
fldname :: "vname \<times> qtname \<Rightarrow> vname"
@@ -345,8 +345,8 @@
lemma ext_fieldname_simp[simp]: "(fldname f,fldclass f) = f"
by (simp add: fldname_def fldclass_def)
-text {* Convert a qualified method declaration (qualified with its declaring
-class) to a qualified member declaration: @{text methdMembr} *}
+text \<open>Convert a qualified method declaration (qualified with its declaring
+class) to a qualified member declaration: \<open>methdMembr\<close>\<close>
definition
methdMembr :: "qtname \<times> mdecl \<Rightarrow> qtname \<times> memberdecl"
@@ -364,8 +364,8 @@
lemma declclass_methdMembr_simp[simp]: "declclass (methdMembr m) = declclass m"
by (cases m) (simp add: methdMembr_def)
-text {* Convert a qualified method (qualified with its declaring
-class) to a qualified member declaration: @{text method} *}
+text \<open>Convert a qualified method (qualified with its declaring
+class) to a qualified member declaration: \<open>method\<close>\<close>
definition
"method" :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> memberdecl)"
@@ -411,8 +411,8 @@
lemma memberid_fieldm_simp[simp]: "memberid (fieldm n f) = fid n"
by (simp add: fieldm_def)
-text {* Select the signature out of a qualified method declaration:
- @{text msig} *}
+text \<open>Select the signature out of a qualified method declaration:
+ \<open>msig\<close>\<close>
definition
msig :: "(qtname \<times> mdecl) \<Rightarrow> sig"
@@ -421,8 +421,8 @@
lemma msig_simp[simp]: "msig (c,(s,m)) = s"
by (simp add: msig_def)
-text {* Convert a qualified method (qualified with its declaring
-class) to a qualified method declaration: @{text qmdecl} *}
+text \<open>Convert a qualified method (qualified with its declaring
+class) to a qualified method declaration: \<open>qmdecl\<close>\<close>
definition
qmdecl :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> mdecl)"
@@ -451,8 +451,8 @@
"methdMembr (qmdecl sig old) = method sig old"
by (simp add: methdMembr_def qmdecl_def method_def)
-text {* overloaded selector @{text resTy} to select the result type
-out of various HOL types *}
+text \<open>overloaded selector \<open>resTy\<close> to select the result type
+out of various HOL types\<close>
class has_resTy =
fixes resTy:: "'a \<Rightarrow> ty"
@@ -507,8 +507,8 @@
by (cases m) (simp add: mthd_def )
subsubsection "inheritable-in"
-text {*
-@{text "G\<turnstile>m inheritable_in P"}: m can be inherited by
+text \<open>
+\<open>G\<turnstile>m inheritable_in P\<close>: m can be inherited by
classes in package P if:
\begin{itemize}
\item the declaration class of m is accessible in P and
@@ -517,7 +517,7 @@
class of m is also P. If the member m is declared with private access
it is not accessible for inheritance at all.
\end{itemize}
-*}
+\<close>
definition
inheritable_in :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> pname \<Rightarrow> bool" ("_ \<turnstile> _ inheritable'_in _" [61,61,61] 60)
where
@@ -603,14 +603,14 @@
| Inherited: "\<lbrakk>G\<turnstile>m inheritable_in (pid C); G\<turnstile>memberid m undeclared_in C;
G\<turnstile>C \<prec>\<^sub>C1 S; G\<turnstile>(Class S) accessible_in (pid C);G\<turnstile>m member_of S
\<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
-text {* Note that in the case of an inherited member only the members of the
+text \<open>Note that in the case of an inherited member only the members of the
direct superclass are concerned. If a member of a superclass of the direct
superclass isn't inherited in the direct superclass (not member of the
direct superclass) than it can't be a member of the class. E.g. If a
member of a class A is defined with package access it isn't member of a
subclass S if S isn't in the same package as A. Any further subclasses
of S will not inherit the member, regardless if they are in the same
-package as A or not.*}
+package as A or not.\<close>
abbreviation
method_member_of:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
@@ -641,10 +641,10 @@
definition
member_in :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool" ("_ \<turnstile> _ member'_in _" [61,61,61] 60)
where "G\<turnstile>m member_in C = (\<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> m member_of provC)"
-text {* A member is in a class if it is member of the class or a superclass.
+text \<open>A member is in a class if it is member of the class or a superclass.
If a member is in a class we can select this member. This additional notion
is necessary since not all members are inherited to subclasses. So such
-members are not member-of the subclass but member-in the subclass.*}
+members are not member-of the subclass but member-in the subclass.\<close>
abbreviation
method_member_in:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
@@ -669,12 +669,12 @@
subsubsection "overriding"
-text {* Unfortunately the static notion of overriding (used during the
+text \<open>Unfortunately the static notion of overriding (used during the
typecheck of the compiler) and the dynamic notion of overriding (used during
execution in the JVM) are not exactly the same.
-*}
+\<close>
-text {* Static overriding (used during the typecheck of the compiler) *}
+text \<open>Static overriding (used during the typecheck of the compiler)\<close>
inductive
stat_overridesR :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
@@ -693,7 +693,7 @@
| Indirect: "\<lbrakk>G\<turnstile>new overrides\<^sub>S intr; G\<turnstile>intr overrides\<^sub>S old\<rbrakk>
\<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
-text {* Dynamic overriding (used during the typecheck of the compiler) *}
+text \<open>Dynamic overriding (used during the typecheck of the compiler)\<close>
inductive
overridesR :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
@@ -799,15 +799,15 @@
(G\<turnstile>accclass \<prec>\<^sub>C declclass membr
\<and> (G\<turnstile>cls \<preceq>\<^sub>C accclass \<or> is_static membr))
| Public \<Rightarrow> True)"
-text {*
+text \<open>
The subcondition of the @{term "Protected"} case:
@{term "G\<turnstile>accclass \<prec>\<^sub>C declclass membr"} could also be relaxed to:
@{term "G\<turnstile>accclass \<preceq>\<^sub>C declclass membr"} since in case both classes are the
same the other condition @{term "(pid (declclass membr) = pid accclass)"}
holds anyway.
-*}
+\<close>
-text {* Like in case of overriding, the static and dynamic accessibility
+text \<open>Like in case of overriding, the static and dynamic accessibility
of members is not uniform.
\begin{itemize}
\item Statically the class/interface of the member must be accessible for the
@@ -819,7 +819,7 @@
\item Statically the member we want to access must be "member of" the class.
Dynamically it must only be "member in" the class.
\end{itemize}
-*}
+\<close>
inductive
accessible_fromR :: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
@@ -1403,7 +1403,7 @@
"imethds G I =
iface_rec G I (\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus>
(set_option \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
-text {* methods of an interface, with overriding and inheritance, cf. 9.2 *}
+text \<open>methods of an interface, with overriding and inheritance, cf. 9.2\<close>
definition
accimethds :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables" where
@@ -1411,7 +1411,7 @@
(if G\<turnstile>Iface I accessible_in pack
then imethds G I
else (\<lambda> k. {}))"
-text {* only returns imethds if the interface is accessible *}
+text \<open>only returns imethds if the interface is accessible\<close>
definition
methd :: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
@@ -1422,25 +1422,25 @@
subcls_mthds
++
table_of (map (\<lambda>(s,m). (s,C,m)) (methods c)))"
-text {* @{term "methd G C"}: methods of a class C (statically visible from C),
+text \<open>@{term "methd G C"}: methods of a class C (statically visible from C),
with inheritance and hiding cf. 8.4.6;
- Overriding is captured by @{text dynmethd}.
+ Overriding is captured by \<open>dynmethd\<close>.
Every new method with the same signature coalesces the
- method of a superclass. *}
+ method of a superclass.\<close>
definition
accmethd :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
"accmethd G S C =
filter_tab (\<lambda>sig m. G\<turnstile>method sig m of C accessible_from S) (methd G C)"
-text {* @{term "accmethd G S C"}: only those methods of @{term "methd G C"},
- accessible from S *}
+text \<open>@{term "accmethd G S C"}: only those methods of @{term "methd G C"},
+ accessible from S\<close>
-text {* Note the class component in the accessibility filter. The class where
+text \<open>Note the class component in the accessibility filter. The class where
method @{term m} is declared (@{term declC}) isn't necessarily accessible
from the current scope @{term S}. The method can be made accessible
through inheritance, too.
So we must test accessibility of method @{term m} of class @{term C}
- (not @{term "declclass m"}) *}
+ (not @{term "declclass m"})\<close>
definition
dynmethd :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
@@ -1461,13 +1461,13 @@
)
else None))"
-text {* @{term "dynmethd G statC dynC"}: dynamic method lookup of a reference
- with dynamic class @{term dynC} and static class @{term statC} *}
-text {* Note some kind of duality between @{term methd} and @{term dynmethd}
+text \<open>@{term "dynmethd G statC dynC"}: dynamic method lookup of a reference
+ with dynamic class @{term dynC} and static class @{term statC}\<close>
+text \<open>Note some kind of duality between @{term methd} and @{term dynmethd}
in the @{term class_rec} arguments. Whereas @{term methd} filters the
subclass methods (to get only the inherited ones), @{term dynmethd}
filters the new methods (to get only those methods which actually
- override the methods of the static class) *}
+ override the methods of the static class)\<close>
definition
dynimethd :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
@@ -1475,9 +1475,9 @@
(\<lambda>sig. if imethds G I sig \<noteq> {}
then methd G dynC sig
else dynmethd G Object dynC sig)"
-text {* @{term "dynimethd G I dynC"}: dynamic method lookup of a reference with
- dynamic class dynC and static interface type I *}
-text {*
+text \<open>@{term "dynimethd G I dynC"}: dynamic method lookup of a reference with
+ dynamic class dynC and static interface type I\<close>
+text \<open>
When calling an interface method, we must distinguish if the method signature
was defined in the interface or if it must be an Object method in the other
case. If it was an interface method we search the class hierarchy
@@ -1487,7 +1487,7 @@
effects like in case of dynmethd. The method will be inherited or
overridden in all classes from the first class implementing the interface
down to the actual dynamic class.
- *}
+\<close>
definition
dynlookup :: "prog \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
@@ -1497,19 +1497,19 @@
| IfaceT I \<Rightarrow> dynimethd G I dynC
| ClassT statC \<Rightarrow> dynmethd G statC dynC
| ArrayT ty \<Rightarrow> dynmethd G Object dynC)"
-text {* @{term "dynlookup G statT dynC"}: dynamic lookup of a method within the
+text \<open>@{term "dynlookup G statT dynC"}: dynamic lookup of a method within the
static reference type statT and the dynamic class dynC.
In a wellformd context statT will not be NullT and in case
- statT is an array type, dynC=Object *}
+ statT is an array type, dynC=Object\<close>
definition
fields :: "prog \<Rightarrow> qtname \<Rightarrow> ((vname \<times> qtname) \<times> field) list" where
"fields G C =
class_rec G C [] (\<lambda>C c ts. map (\<lambda>(n,t). ((n,C),t)) (cfields c) @ ts)"
-text {* @{term "fields G C"}
+text \<open>@{term "fields G C"}
list of fields of a class, including all the fields of the superclasses
(private, inherited and hidden ones) not only the accessible ones
- (an instance of a object allocates all these fields *}
+ (an instance of a object allocates all these fields\<close>
definition
accfield :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (vname, qtname \<times> field) table" where
@@ -1517,15 +1517,15 @@
(let field_tab = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (fields G C))
in filter_tab (\<lambda>n (declC,f). G\<turnstile> (declC,fdecl (n,f)) of C accessible_from S)
field_tab)"
-text {* @{term "accfield G C S"}: fields of a class @{term C} which are
+text \<open>@{term "accfield G C S"}: fields of a class @{term C} which are
accessible from scope of class
- @{term S} with inheritance and hiding, cf. 8.3 *}
-text {* note the class component in the accessibility filter (see also
+ @{term S} with inheritance and hiding, cf. 8.3\<close>
+text \<open>note the class component in the accessibility filter (see also
@{term methd}).
The class declaring field @{term f} (@{term declC}) isn't necessarily
accessible from scope @{term S}. The field can be made visible through
inheritance, too. So we must test accessibility of field @{term f} of class
- @{term C} (not @{term "declclass f"}) *}
+ @{term C} (not @{term "declclass f"})\<close>
definition
is_methd :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> bool"
--- a/src/HOL/Bali/DefiniteAssignment.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/DefiniteAssignment.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,8 +1,8 @@
-subsection {* Definite Assignment *}
+subsection \<open>Definite Assignment\<close>
theory DefiniteAssignment imports WellType begin
-text {* Definite Assignment Analysis (cf. 16)
+text \<open>Definite Assignment Analysis (cf. 16)
The definite assignment analysis approximates the sets of local
variables that will be assigned at a certain point of evaluation, and ensures
@@ -37,17 +37,17 @@
\item analysis of definite unassigned
\item special treatment of final fields
\end{itemize}
-*}
+\<close>
-subsubsection {* Correct nesting of jump statements *}
+subsubsection \<open>Correct nesting of jump statements\<close>
-text {* For definite assignment it becomes crucial, that jumps (break,
+text \<open>For definite assignment it becomes crucial, that jumps (break,
continue, return) are nested correctly i.e. a continue jump is nested in a
matching while statement, a break jump is nested in a proper label statement,
a class initialiser does not terminate abruptly with a return. With this we
can for example ensure that evaluation of an expression will never end up
with a jump, since no breaks, continues or returns are allowed in an
-expression. *}
+expression.\<close>
primrec jumpNestingOkS :: "jump set \<Rightarrow> stmt \<Rightarrow> bool"
where
@@ -59,8 +59,8 @@
| "jumpNestingOkS jmps (If(e) c1 Else c2) = (jumpNestingOkS jmps c1 \<and>
jumpNestingOkS jmps c2)"
| "jumpNestingOkS jmps (l\<bullet> While(e) c) = jumpNestingOkS ({Cont l} \<union> jmps) c"
---{* The label of the while loop only handles continue jumps. Breaks are only
- handled by @{term Lab} *}
+\<comment>\<open>The label of the while loop only handles continue jumps. Breaks are only
+ handled by @{term Lab}\<close>
| "jumpNestingOkS jmps (Jmp j) = (j \<in> jmps)"
| "jumpNestingOkS jmps (Throw e) = True"
| "jumpNestingOkS jmps (Try c1 Catch(C vn) c2) = (jumpNestingOkS jmps c1 \<and>
@@ -68,9 +68,9 @@
| "jumpNestingOkS jmps (c1 Finally c2) = (jumpNestingOkS jmps c1 \<and>
jumpNestingOkS jmps c2)"
| "jumpNestingOkS jmps (Init C) = True"
- --{* wellformedness of the program must enshure that for all initializers
- jumpNestingOkS {} holds *}
---{* Dummy analysis for intermediate smallstep term @{term FinA} *}
+ \<comment>\<open>wellformedness of the program must enshure that for all initializers
+ jumpNestingOkS {} holds\<close>
+\<comment>\<open>Dummy analysis for intermediate smallstep term @{term FinA}\<close>
| "jumpNestingOkS jmps (FinA a c) = False"
@@ -111,10 +111,10 @@
-subsubsection {* Calculation of assigned variables for boolean expressions*}
+subsubsection \<open>Calculation of assigned variables for boolean expressions\<close>
-subsection {* Very restricted calculation fallback calculation *}
+subsection \<open>Very restricted calculation fallback calculation\<close>
primrec the_LVar_name :: "var \<Rightarrow> lname"
where "the_LVar_name (LVar n) = n"
@@ -140,8 +140,8 @@
| "assignsE (b? e1 : e2) = (assignsE b) \<union> ((assignsE e1) \<inter> (assignsE e2))"
| "assignsE ({accC,statT,mode}objRef\<cdot>mn({pTs}args))
= (assignsE objRef) \<union> (assignsEs args)"
--- {* Only dummy analysis for intermediate expressions
- @{term Methd}, @{term Body}, @{term InsInitE} and @{term Callee} *}
+\<comment> \<open>Only dummy analysis for intermediate expressions
+ @{term Methd}, @{term Body}, @{term InsInitE} and @{term Callee}\<close>
| "assignsE (Methd C sig) = {}"
| "assignsE (Body C s) = {}"
| "assignsE (InsInitE s e) = {}"
@@ -216,9 +216,9 @@
| False\<Rightarrow> (case (constVal e1) of
None \<Rightarrow> None
| Some v \<Rightarrow> constVal e2)))"
---{* Note that @{text "constVal (Cond b e1 e2)"} is stricter as it could be.
+\<comment>\<open>Note that \<open>constVal (Cond b e1 e2)\<close> is stricter as it could be.
It requires that all tree expressions are constant even if we can decide
- which branch to choose, provided the constant value of @{term b} *}
+ which branch to choose, provided the constant value of @{term b}\<close>
| "constVal (Call accC statT mode objRef mn pTs args) = None"
| "constVal (Methd C sig) = None"
| "constVal (Body C s) = None"
@@ -274,19 +274,19 @@
by (induct rule: constVal_Some_induct) simp_all
-subsection {* Main analysis for boolean expressions *}
+subsection \<open>Main analysis for boolean expressions\<close>
-text {* Assigned local variables after evaluating the expression if it evaluates
+text \<open>Assigned local variables after evaluating the expression if it evaluates
to a specific boolean value. If the expression cannot evaluate to a
@{term Boolean} value UNIV is returned. If we expect true/false the opposite
-constant false/true will also lead to UNIV. *}
+constant false/true will also lead to UNIV.\<close>
primrec assigns_if :: "bool \<Rightarrow> expr \<Rightarrow> lname set"
where
- "assigns_if b (NewC c) = UNIV" --{*can never evaluate to Boolean*}
-| "assigns_if b (NewA t e) = UNIV" --{*can never evaluate to Boolean*}
+ "assigns_if b (NewC c) = UNIV" \<comment>\<open>can never evaluate to Boolean\<close>
+| "assigns_if b (NewA t e) = UNIV" \<comment>\<open>can never evaluate to Boolean\<close>
| "assigns_if b (Cast t e) = assigns_if b e"
-| "assigns_if b (Inst e r) = assignsE e" --{*Inst has type Boolean but
- e is a reference type*}
+| "assigns_if b (Inst e r) = assignsE e" \<comment>\<open>Inst has type Boolean but
+ e is a reference type\<close>
| "assigns_if b (Lit val) = (if val=Bool b then {} else UNIV)"
| "assigns_if b (UnOp unop e) = (case constVal (UnOp unop e) of
None \<Rightarrow> (if unop = UNot
@@ -311,7 +311,7 @@
else assignsE e1 \<union> assignsE e2))
| Some v \<Rightarrow> (if v=Bool b then {} else UNIV))"
-| "assigns_if b (Super) = UNIV" --{*can never evaluate to Boolean*}
+| "assigns_if b (Super) = UNIV" \<comment>\<open>can never evaluate to Boolean\<close>
| "assigns_if b (Acc v) = (assignsV v)"
| "assigns_if b (v := e) = (assignsE (Ass v e))"
| "assigns_if b (c? e1 : e2) = (assignsE c) \<union>
@@ -323,8 +323,8 @@
| False \<Rightarrow> assigns_if b e2))"
| "assigns_if b ({accC,statT,mode}objRef\<cdot>mn({pTs}args))
= assignsE ({accC,statT,mode}objRef\<cdot>mn({pTs}args)) "
--- {* Only dummy analysis for intermediate expressions
- @{term Methd}, @{term Body}, @{term InsInitE} and @{term Callee} *}
+\<comment> \<open>Only dummy analysis for intermediate expressions
+ @{term Methd}, @{term Body}, @{term InsInitE} and @{term Callee}\<close>
| "assigns_if b (Methd C sig) = {}"
| "assigns_if b (Body C s) = {}"
| "assigns_if b (InsInitE s e) = {}"
@@ -347,10 +347,10 @@
by (cases binop) (simp_all)
next
case (Cond c e1 e2 b)
- note hyp_c = `\<And> b. ?Const b c \<Longrightarrow> ?Ass b c`
- note hyp_e1 = `\<And> b. ?Const b e1 \<Longrightarrow> ?Ass b e1`
- note hyp_e2 = `\<And> b. ?Const b e2 \<Longrightarrow> ?Ass b e2`
- note const = `constVal (c ? e1 : e2) = Some (Bool b)`
+ note hyp_c = \<open>\<And> b. ?Const b c \<Longrightarrow> ?Ass b c\<close>
+ note hyp_e1 = \<open>\<And> b. ?Const b e1 \<Longrightarrow> ?Ass b e1\<close>
+ note hyp_e2 = \<open>\<And> b. ?Const b e2 \<Longrightarrow> ?Ass b e2\<close>
+ note const = \<open>constVal (c ? e1 : e2) = Some (Bool b)\<close>
then obtain bv where bv: "constVal c = Some bv"
by simp
hence emptyC: "assignsE c = {}" by (rule assignsE_const_simp)
@@ -395,10 +395,10 @@
by (cases binop) (simp_all)
next
case (Cond c e1 e2 b)
- note hyp_c = `\<And> b. ?Const b c \<Longrightarrow> ?Ass b c`
- note hyp_e1 = `\<And> b. ?Const b e1 \<Longrightarrow> ?Ass b e1`
- note hyp_e2 = `\<And> b. ?Const b e2 \<Longrightarrow> ?Ass b e2`
- note const = `constVal (c ? e1 : e2) = Some (Bool b)`
+ note hyp_c = \<open>\<And> b. ?Const b c \<Longrightarrow> ?Ass b c\<close>
+ note hyp_e1 = \<open>\<And> b. ?Const b e1 \<Longrightarrow> ?Ass b e1\<close>
+ note hyp_e2 = \<open>\<And> b. ?Const b e2 \<Longrightarrow> ?Ass b e2\<close>
+ note const = \<open>constVal (c ? e1 : e2) = Some (Bool b)\<close>
then obtain bv where bv: "constVal c = Some bv"
by simp
show ?case
@@ -425,7 +425,7 @@
by blast
qed
-subsection {* Lifting set operations to range of tables (map to a set) *}
+subsection \<open>Lifting set operations to range of tables (map to a set)\<close>
definition
union_ts :: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" ("_ \<Rightarrow>\<union> _" [67,67] 65)
@@ -439,7 +439,7 @@
all_union_ts :: "('a,'b) tables \<Rightarrow> 'b set \<Rightarrow> ('a,'b) tables" (infixl "\<Rightarrow>\<union>\<^sub>\<forall>" 40)
where "(A \<Rightarrow>\<union>\<^sub>\<forall> B) = (\<lambda> k. A k \<union> B)"
-subsubsection {* Binary union of tables *}
+subsubsection \<open>Binary union of tables\<close>
lemma union_ts_iff [simp]: "(c \<in> (A \<Rightarrow>\<union> B) k) = (c \<in> A k \<or> c \<in> B k)"
by (unfold union_ts_def) blast
@@ -457,7 +457,7 @@
"\<lbrakk>c \<in> (A \<Rightarrow>\<union> B) k; (c \<in> A k \<Longrightarrow> P); (c \<in> B k \<Longrightarrow> P)\<rbrakk> \<Longrightarrow> P"
by (unfold union_ts_def) blast
-subsubsection {* Binary intersection of tables *}
+subsubsection \<open>Binary intersection of tables\<close>
lemma intersect_ts_iff [simp]: "c \<in> (A \<Rightarrow>\<inter> B) k = (c \<in> A k \<and> c \<in> B k)"
by (unfold intersect_ts_def) blast
@@ -476,7 +476,7 @@
by simp
-subsubsection {* All-Union of tables and set *}
+subsubsection \<open>All-Union of tables and set\<close>
lemma all_union_ts_iff [simp]: "(c \<in> (A \<Rightarrow>\<union>\<^sub>\<forall> B) k) = (c \<in> A k \<or> c \<in> B)"
by (unfold all_union_ts_def) blast
@@ -499,14 +499,14 @@
type_synonym breakass = "(label, lname) tables"
---{* Mapping from a break label, to the set of variables that will be assigned
- if the evaluation terminates with this break *}
+\<comment>\<open>Mapping from a break label, to the set of variables that will be assigned
+ if the evaluation terminates with this break\<close>
record assigned =
- nrm :: "lname set" --{* Definetly assigned variables
- for normal completion*}
- brk :: "breakass" --{* Definetly assigned variables for
- abrupt completion with a break *}
+ nrm :: "lname set" \<comment>\<open>Definetly assigned variables
+ for normal completion\<close>
+ brk :: "breakass" \<comment>\<open>Definetly assigned variables for
+ abrupt completion with a break\<close>
definition
rmlab :: "'a \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables"
@@ -522,14 +522,14 @@
range_inter_ts :: "('a,'b) tables \<Rightarrow> 'b set" ("\<Rightarrow>\<Inter>_" 80)
where "\<Rightarrow>\<Inter>A = {x |x. \<forall> k. x \<in> A k}"
-text {*
-In @{text "E\<turnstile> B \<guillemotright>t\<guillemotright> A"},
-@{text B} denotes the ''assigned'' variables before evaluating term @{text t},
-whereas @{text A} denotes the ''assigned'' variables after evaluating term @{text t}.
-The environment @{term E} is only needed for the conditional @{text "_ ? _ : _"}.
+text \<open>
+In \<open>E\<turnstile> B \<guillemotright>t\<guillemotright> A\<close>,
+\<open>B\<close> denotes the ''assigned'' variables before evaluating term \<open>t\<close>,
+whereas \<open>A\<close> denotes the ''assigned'' variables after evaluating term \<open>t\<close>.
+The environment @{term E} is only needed for the conditional \<open>_ ? _ : _\<close>.
The definite assignment rules refer to the typing rules here to
distinguish boolean and other expressions.
-*}
+\<close>
inductive
da :: "env \<Rightarrow> lname set \<Rightarrow> term \<Rightarrow> assigned \<Rightarrow> bool" ("_\<turnstile> _ \<guillemotright>_\<guillemotright> _" [65,65,65,65] 71)
@@ -556,7 +556,7 @@
\<Longrightarrow>
Env\<turnstile> B \<guillemotright>\<langle>If(e) c1 Else c2\<rangle>\<guillemotright> A"
---{* Note that @{term E} is not further used, because we take the specialized
+\<comment>\<open>Note that @{term E} is not further used, because we take the specialized
sets that also consider if the expression evaluates to true or false.
Inside of @{term e} there is no {\tt break} or {\tt finally}, so the break
map of @{term E} will be the trivial one. So
@@ -572,7 +572,7 @@
to @{term UNIV} too, because @{term "assigns_if False e = UNIV"}. So
in the intersection of the break maps the path @{term c2} will have no
contribution.
- *}
+\<close>
| Loop: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E;
Env\<turnstile> (B \<union> assigns_if True e) \<guillemotright>\<langle>c\<rangle>\<guillemotright> C;
@@ -580,7 +580,7 @@
brk A = brk C\<rbrakk>
\<Longrightarrow>
Env\<turnstile> B \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A"
---{* The @{text Loop} rule resembles some of the ideas of the @{text If} rule.
+\<comment>\<open>The \<open>Loop\<close> rule resembles some of the ideas of the \<open>If\<close> rule.
For the @{term "nrm A"} the set @{term "B \<union> assigns_if False e"}
will be @{term UNIV} if the condition is constantly true. To normally exit
the while loop, we must consider the body @{term c} to be completed
@@ -589,7 +589,7 @@
only handles continue labels, not break labels. The break label will be
handled by an enclosing @{term Lab} statement. So we don't have to
handle the breaks specially.
- *}
+\<close>
| Jmp: "\<lbrakk>jump=Ret \<longrightarrow> Result \<in> B;
nrm A = UNIV;
@@ -599,13 +599,13 @@
| Ret \<Rightarrow> \<lambda> k. UNIV)\<rbrakk>
\<Longrightarrow>
Env\<turnstile> B \<guillemotright>\<langle>Jmp jump\<rangle>\<guillemotright> A"
---{* In case of a break to label @{term l} the corresponding break set is all
+\<comment>\<open>In case of a break to label @{term l} the corresponding break set is all
variables assigned before the break. The assigned variables for normal
completion of the @{term Jmp} is @{term UNIV}, because the statement will
never complete normally. For continue and return the break map is the
trivial one. In case of a return we enshure that the result value is
assigned.
- *}
+\<close>
| Throw: "\<lbrakk>Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> E; nrm A = UNIV; brk A = (\<lambda> l. UNIV)\<rbrakk>
\<Longrightarrow> Env\<turnstile> B \<guillemotright>\<langle>Throw e\<rangle>\<guillemotright> A"
@@ -622,23 +622,23 @@
brk A = ((brk C1) \<Rightarrow>\<union>\<^sub>\<forall> (nrm C2)) \<Rightarrow>\<inter> (brk C2)\<rbrakk>
\<Longrightarrow>
Env\<turnstile> B \<guillemotright>\<langle>c1 Finally c2\<rangle>\<guillemotright> A"
---{* The set of assigned variables before execution @{term c2} are the same
+\<comment>\<open>The set of assigned variables before execution @{term c2} are the same
as before execution @{term c1}, because @{term c1} could throw an exception
and so we can't guarantee that any variable will be assigned in @{term c1}.
- The @{text Finally} statement completes
+ The \<open>Finally\<close> statement completes
normally if both @{term c1} and @{term c2} complete normally. If @{term c1}
completes abruptly with a break, then @{term c2} also will be executed
and may terminate normally or with a break. The overall break map then is
the intersection of the maps of both paths. If @{term c2} terminates
normally we have to extend all break sets in @{term "brk C1"} with
- @{term "nrm C2"} (@{text "\<Rightarrow>\<union>\<^sub>\<forall>"}). If @{term c2} exits with a break this
+ @{term "nrm C2"} (\<open>\<Rightarrow>\<union>\<^sub>\<forall>\<close>). If @{term c2} exits with a break this
break will appear in the overall result state. We don't know if
@{term c1} completed normally or abruptly (maybe with an exception not only
a break) so @{term c1} has no contribution to the break map following this
path.
- *}
+\<close>
---{* Evaluation of expressions and the break sets of definite assignment:
+\<comment>\<open>Evaluation of expressions and the break sets of definite assignment:
Thinking of a Java expression we assume that we can never have
a break statement inside of a expression. So for all expressions the
break sets could be set to the trivial one: @{term "\<lambda> l. UNIV"}.
@@ -656,18 +656,18 @@
the analysis of the correct nesting of breaks in the typing judgments
right now. So we have decided to adjust the rules of definite assignment
to fit to these circumstances. If an initialization is involved during
- evaluation of the expression (evaluation rules @{text FVar}, @{text NewC}
- and @{text NewA}
-*}
+ evaluation of the expression (evaluation rules \<open>FVar\<close>, \<open>NewC\<close>
+ and \<open>NewA\<close>
+\<close>
| Init: "Env\<turnstile> B \<guillemotright>\<langle>Init C\<rangle>\<guillemotright> \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>"
---{* Wellformedness of a program will ensure, that every static initialiser
+\<comment>\<open>Wellformedness of a program will ensure, that every static initialiser
is definetly assigned and the jumps are nested correctly. The case here
for @{term Init} is just for convenience, to get a proper precondition
for the induction hypothesis in various proofs, so that we don't have to
expand the initialisation on every point where it is triggerred by the
evaluation rules.
- *}
+\<close>
| NewC: "Env\<turnstile> B \<guillemotright>\<langle>NewC C\<rangle>\<guillemotright> \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>"
| NewA: "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> A
@@ -715,9 +715,9 @@
nrm A = B; brk A = (\<lambda> k. UNIV)\<rbrakk>
\<Longrightarrow>
Env\<turnstile> B \<guillemotright>\<langle>Acc (LVar vn)\<rangle>\<guillemotright> A"
---{* To properly access a local variable we have to test the definite
+\<comment>\<open>To properly access a local variable we have to test the definite
assignment here. The variable must occur in the set @{term B}
- *}
+\<close>
| Acc: "\<lbrakk>\<forall> vn. v \<noteq> LVar vn;
Env\<turnstile> B \<guillemotright>\<langle>v\<rangle>\<guillemotright> A\<rbrakk>
@@ -754,7 +754,7 @@
\<Longrightarrow>
Env\<turnstile> B \<guillemotright>\<langle>{accC,statT,mode}e\<cdot>mn({pTs}args)\<rangle>\<guillemotright> A"
--- {* The interplay of @{term Call}, @{term Methd} and @{term Body}:
+\<comment> \<open>The interplay of @{term Call}, @{term Methd} and @{term Body}:
Why rules for @{term Methd} and @{term Body} at all? Note that a
Java source program will not include bare @{term Methd} or @{term Body}
terms. These terms are just introduced during evaluation. So definite
@@ -774,7 +774,7 @@
sub-evaluation during the type-safety proof. Note that well-typedness is
also a precondition for type-safety and so we can omit some assertion
that are already ensured by well-typedness.
- *}
+\<close>
| Methd: "\<lbrakk>methd (prg Env) D sig = Some m;
Env\<turnstile> B \<guillemotright>\<langle>Body (declclass m) (stmt (mbody (mthd m)))\<rangle>\<guillemotright> A
\<rbrakk>
@@ -785,7 +785,7 @@
nrm A = B; brk A = (\<lambda> l. UNIV)\<rbrakk>
\<Longrightarrow>
Env\<turnstile> B \<guillemotright>\<langle>Body D c\<rangle>\<guillemotright> A"
--- {* Note that @{term A} is not correlated to @{term C}. If the body
+\<comment> \<open>Note that @{term A} is not correlated to @{term C}. If the body
statement returns abruptly with return, evaluation of @{term Body}
will absorb this return and complete normally. So we cannot trivially
get the assigned variables of the body statement since it has not
@@ -797,7 +797,7 @@
for a return the @{term Jump} rule ensures that the result variable is
set and then this information must be carried over to the @{term Body}
rule by the conformance predicate of the state.
- *}
+\<close>
| LVar: "Env\<turnstile> B \<guillemotright>\<langle>LVar vn\<rangle>\<guillemotright> \<lparr>nrm=B, brk=\<lambda> l. UNIV\<rparr>"
| FVar: "Env\<turnstile> B \<guillemotright>\<langle>e\<rangle>\<guillemotright> A
@@ -818,7 +818,7 @@
declare inj_term_sym_simps [simp]
declare assigns_if.simps [simp del]
declare split_paired_All [simp del] split_paired_Ex [simp del]
-setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\<close>
inductive_cases da_elim_cases [cases set]:
"Env\<turnstile> B \<guillemotright>\<langle>Skip\<rangle>\<guillemotright> A"
@@ -884,7 +884,7 @@
declare inj_term_sym_simps [simp del]
declare assigns_if.simps [simp]
declare split_paired_All [simp] split_paired_Ex [simp]
-setup {* map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac)) *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac))\<close>
(* To be able to eliminate both the versions with the overloaded brackets:
(B \<guillemotright>\<langle>Skip\<rangle>\<guillemotright> A) and with the explicit constructor (B \<guillemotright>In1r Skip\<guillemotright> A),
@@ -956,7 +956,7 @@
case (Cast T e)
have "E\<turnstile>e\<Colon>- (PrimT Boolean)"
proof -
- from `E\<turnstile>(Cast T e)\<Colon>- (PrimT Boolean)`
+ from \<open>E\<turnstile>(Cast T e)\<Colon>- (PrimT Boolean)\<close>
obtain Te where "E\<turnstile>e\<Colon>-Te"
"prg E\<turnstile>Te\<preceq>? PrimT Boolean"
by cases simp
@@ -986,10 +986,10 @@
by - (cases binop, auto simp add: assignsE_const_simp)
next
case (Cond c e1 e2)
- note hyp_c = `?Boolean c \<Longrightarrow> ?Incl c`
- note hyp_e1 = `?Boolean e1 \<Longrightarrow> ?Incl e1`
- note hyp_e2 = `?Boolean e2 \<Longrightarrow> ?Incl e2`
- note wt = `E\<turnstile>(c ? e1 : e2)\<Colon>-PrimT Boolean`
+ note hyp_c = \<open>?Boolean c \<Longrightarrow> ?Incl c\<close>
+ note hyp_e1 = \<open>?Boolean e1 \<Longrightarrow> ?Incl e1\<close>
+ note hyp_e2 = \<open>?Boolean e2 \<Longrightarrow> ?Incl e2\<close>
+ note wt = \<open>E\<turnstile>(c ? e1 : e2)\<Colon>-PrimT Boolean\<close>
then obtain
boolean_c: "E\<turnstile>c\<Colon>-PrimT Boolean" and
boolean_e1: "E\<turnstile>e1\<Colon>-PrimT Boolean" and
@@ -1067,10 +1067,10 @@
show ?case by cases simp
next
case (Lab Env B c C A l B' A')
- note A = `nrm A = nrm C \<inter> brk C l` `brk A = rmlab l (brk C)`
- note `PROP ?Hyp Env B \<langle>c\<rangle> C`
+ note A = \<open>nrm A = nrm C \<inter> brk C l\<close> \<open>brk A = rmlab l (brk C)\<close>
+ note \<open>PROP ?Hyp Env B \<langle>c\<rangle> C\<close>
moreover
- note `B \<subseteq> B'`
+ note \<open>B \<subseteq> B'\<close>
moreover
obtain C'
where "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
@@ -1093,19 +1093,19 @@
by simp
next
case (Comp Env B c1 C1 c2 C2 A B' A')
- note A = `nrm A = nrm C2` `brk A = brk C1 \<Rightarrow>\<inter> brk C2`
- from `Env\<turnstile> B' \<guillemotright>\<langle>c1;; c2\<rangle>\<guillemotright> A'`
+ note A = \<open>nrm A = nrm C2\<close> \<open>brk A = brk C1 \<Rightarrow>\<inter> brk C2\<close>
+ from \<open>Env\<turnstile> B' \<guillemotright>\<langle>c1;; c2\<rangle>\<guillemotright> A'\<close>
obtain C1' C2'
where da_c1: "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
da_c2: "Env\<turnstile> nrm C1' \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
A': "nrm A' = nrm C2'" "brk A' = brk C1' \<Rightarrow>\<inter> brk C2'"
by cases auto
- note `PROP ?Hyp Env B \<langle>c1\<rangle> C1`
- moreover note `B \<subseteq> B'`
+ note \<open>PROP ?Hyp Env B \<langle>c1\<rangle> C1\<close>
+ moreover note \<open>B \<subseteq> B'\<close>
moreover note da_c1
ultimately have C1': "nrm C1 \<subseteq> nrm C1'" "(\<forall>l. brk C1 l \<subseteq> brk C1' l)"
by auto
- note `PROP ?Hyp Env (nrm C1) \<langle>c2\<rangle> C2`
+ note \<open>PROP ?Hyp Env (nrm C1) \<langle>c2\<rangle> C2\<close>
with da_c2 C1'
have C2': "nrm C2 \<subseteq> nrm C2'" "(\<forall>l. brk C2 l \<subseteq> brk C2' l)"
by auto
@@ -1114,19 +1114,19 @@
by auto
next
case (If Env B e E c1 C1 c2 C2 A B' A')
- note A = `nrm A = nrm C1 \<inter> nrm C2` `brk A = brk C1 \<Rightarrow>\<inter> brk C2`
- from `Env\<turnstile> B' \<guillemotright>\<langle>If(e) c1 Else c2\<rangle>\<guillemotright> A'`
+ note A = \<open>nrm A = nrm C1 \<inter> nrm C2\<close> \<open>brk A = brk C1 \<Rightarrow>\<inter> brk C2\<close>
+ from \<open>Env\<turnstile> B' \<guillemotright>\<langle>If(e) c1 Else c2\<rangle>\<guillemotright> A'\<close>
obtain C1' C2'
where da_c1: "Env\<turnstile> B' \<union> assigns_if True e \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
da_c2: "Env\<turnstile> B' \<union> assigns_if False e \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
A': "nrm A' = nrm C1' \<inter> nrm C2'" "brk A' = brk C1' \<Rightarrow>\<inter> brk C2'"
by cases auto
- note `PROP ?Hyp Env (B \<union> assigns_if True e) \<langle>c1\<rangle> C1`
- moreover note B' = `B \<subseteq> B'`
+ note \<open>PROP ?Hyp Env (B \<union> assigns_if True e) \<langle>c1\<rangle> C1\<close>
+ moreover note B' = \<open>B \<subseteq> B'\<close>
moreover note da_c1
ultimately obtain C1': "nrm C1 \<subseteq> nrm C1'" "(\<forall>l. brk C1 l \<subseteq> brk C1' l)"
by blast
- note `PROP ?Hyp Env (B \<union> assigns_if False e) \<langle>c2\<rangle> C2`
+ note \<open>PROP ?Hyp Env (B \<union> assigns_if False e) \<langle>c2\<rangle> C2\<close>
with da_c2 B'
obtain C2': "nrm C2 \<subseteq> nrm C2'" "(\<forall>l. brk C2 l \<subseteq> brk C2' l)"
by blast
@@ -1135,16 +1135,16 @@
by auto
next
case (Loop Env B e E c C A l B' A')
- note A = `nrm A = nrm C \<inter> (B \<union> assigns_if False e)` `brk A = brk C`
- from `Env\<turnstile> B' \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A'`
+ note A = \<open>nrm A = nrm C \<inter> (B \<union> assigns_if False e)\<close> \<open>brk A = brk C\<close>
+ from \<open>Env\<turnstile> B' \<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<guillemotright> A'\<close>
obtain C'
where
da_c': "Env\<turnstile> B' \<union> assigns_if True e \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'" and
A': "nrm A' = nrm C' \<inter> (B' \<union> assigns_if False e)"
"brk A' = brk C'"
by cases auto
- note `PROP ?Hyp Env (B \<union> assigns_if True e) \<langle>c\<rangle> C`
- moreover note B' = `B \<subseteq> B'`
+ note \<open>PROP ?Hyp Env (B \<union> assigns_if True e) \<langle>c\<rangle> C\<close>
+ moreover note B' = \<open>B \<subseteq> B'\<close>
moreover note da_c'
ultimately obtain C': "nrm C \<subseteq> nrm C'" "(\<forall>l. brk C l \<subseteq> brk C' l)"
by blast
@@ -1175,8 +1175,8 @@
case Throw thus ?case by (elim da_elim_cases) auto
next
case (Try Env B c1 C1 vn C c2 C2 A B' A')
- note A = `nrm A = nrm C1 \<inter> nrm C2` `brk A = brk C1 \<Rightarrow>\<inter> brk C2`
- from `Env\<turnstile> B' \<guillemotright>\<langle>Try c1 Catch(C vn) c2\<rangle>\<guillemotright> A'`
+ note A = \<open>nrm A = nrm C1 \<inter> nrm C2\<close> \<open>brk A = brk C1 \<Rightarrow>\<inter> brk C2\<close>
+ from \<open>Env\<turnstile> B' \<guillemotright>\<langle>Try c1 Catch(C vn) c2\<rangle>\<guillemotright> A'\<close>
obtain C1' C2'
where da_c1': "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
da_c2': "Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>\<turnstile> B' \<union> {VName vn}
@@ -1184,13 +1184,13 @@
A': "nrm A' = nrm C1' \<inter> nrm C2'"
"brk A' = brk C1' \<Rightarrow>\<inter> brk C2'"
by cases auto
- note `PROP ?Hyp Env B \<langle>c1\<rangle> C1`
- moreover note B' = `B \<subseteq> B'`
+ note \<open>PROP ?Hyp Env B \<langle>c1\<rangle> C1\<close>
+ moreover note B' = \<open>B \<subseteq> B'\<close>
moreover note da_c1'
ultimately obtain C1': "nrm C1 \<subseteq> nrm C1'" "(\<forall>l. brk C1 l \<subseteq> brk C1' l)"
by blast
- note `PROP ?Hyp (Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>)
- (B \<union> {VName vn}) \<langle>c2\<rangle> C2`
+ note \<open>PROP ?Hyp (Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>)
+ (B \<union> {VName vn}) \<langle>c2\<rangle> C2\<close>
with B' da_c2'
obtain "nrm C2 \<subseteq> nrm C2'" "(\<forall>l. brk C2 l \<subseteq> brk C2' l)"
by blast
@@ -1199,21 +1199,21 @@
by auto
next
case (Fin Env B c1 C1 c2 C2 A B' A')
- note A = `nrm A = nrm C1 \<union> nrm C2`
- `brk A = (brk C1 \<Rightarrow>\<union>\<^sub>\<forall> nrm C2) \<Rightarrow>\<inter> (brk C2)`
- from `Env\<turnstile> B' \<guillemotright>\<langle>c1 Finally c2\<rangle>\<guillemotright> A'`
+ note A = \<open>nrm A = nrm C1 \<union> nrm C2\<close>
+ \<open>brk A = (brk C1 \<Rightarrow>\<union>\<^sub>\<forall> nrm C2) \<Rightarrow>\<inter> (brk C2)\<close>
+ from \<open>Env\<turnstile> B' \<guillemotright>\<langle>c1 Finally c2\<rangle>\<guillemotright> A'\<close>
obtain C1' C2'
where da_c1': "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'" and
da_c2': "Env\<turnstile> B' \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'" and
A': "nrm A' = nrm C1' \<union> nrm C2'"
"brk A' = (brk C1' \<Rightarrow>\<union>\<^sub>\<forall> nrm C2') \<Rightarrow>\<inter> (brk C2')"
by cases auto
- note `PROP ?Hyp Env B \<langle>c1\<rangle> C1`
- moreover note B' = `B \<subseteq> B'`
+ note \<open>PROP ?Hyp Env B \<langle>c1\<rangle> C1\<close>
+ moreover note B' = \<open>B \<subseteq> B'\<close>
moreover note da_c1'
ultimately obtain C1': "nrm C1 \<subseteq> nrm C1'" "(\<forall>l. brk C1 l \<subseteq> brk C1' l)"
by blast
- note hyp_c2 = `PROP ?Hyp Env B \<langle>c2\<rangle> C2`
+ note hyp_c2 = \<open>PROP ?Hyp Env B \<langle>c2\<rangle> C2\<close>
from da_c2' B'
obtain "nrm C2 \<subseteq> nrm C2'" "(\<forall>l. brk C2 l \<subseteq> brk C2' l)"
by - (drule hyp_c2,auto)
@@ -1236,17 +1236,17 @@
case UnOp thus ?case by (elim da_elim_cases) auto
next
case (CondAnd Env B e1 E1 e2 E2 A B' A')
- note A = `nrm A = B \<union>
+ note A = \<open>nrm A = B \<union>
assigns_if True (BinOp CondAnd e1 e2) \<inter>
- assigns_if False (BinOp CondAnd e1 e2)`
- `brk A = (\<lambda>l. UNIV)`
- from `Env\<turnstile> B' \<guillemotright>\<langle>BinOp CondAnd e1 e2\<rangle>\<guillemotright> A'`
+ assigns_if False (BinOp CondAnd e1 e2)\<close>
+ \<open>brk A = (\<lambda>l. UNIV)\<close>
+ from \<open>Env\<turnstile> B' \<guillemotright>\<langle>BinOp CondAnd e1 e2\<rangle>\<guillemotright> A'\<close>
obtain A': "nrm A' = B' \<union>
assigns_if True (BinOp CondAnd e1 e2) \<inter>
assigns_if False (BinOp CondAnd e1 e2)"
"brk A' = (\<lambda>l. UNIV)"
by cases auto
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
with A A' show ?case
by auto
next
@@ -1265,13 +1265,13 @@
case Ass thus ?case by (elim da_elim_cases) auto
next
case (CondBool Env c e1 e2 B C E1 E2 A B' A')
- note A = `nrm A = B \<union>
+ note A = \<open>nrm A = B \<union>
assigns_if True (c ? e1 : e2) \<inter>
- assigns_if False (c ? e1 : e2)`
- `brk A = (\<lambda>l. UNIV)`
- note `Env\<turnstile> (c ? e1 : e2)\<Colon>- (PrimT Boolean)`
+ assigns_if False (c ? e1 : e2)\<close>
+ \<open>brk A = (\<lambda>l. UNIV)\<close>
+ note \<open>Env\<turnstile> (c ? e1 : e2)\<Colon>- (PrimT Boolean)\<close>
moreover
- note `Env\<turnstile> B' \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A'`
+ note \<open>Env\<turnstile> B' \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A'\<close>
ultimately
obtain A': "nrm A' = B' \<union>
assigns_if True (c ? e1 : e2) \<inter>
@@ -1279,14 +1279,14 @@
"brk A' = (\<lambda>l. UNIV)"
by (elim da_elim_cases) (auto simp add: inj_term_simps)
(* inj_term_simps needed to handle wt (defined without \<langle>\<rangle>) *)
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
with A A' show ?case
by auto
next
case (Cond Env c e1 e2 B C E1 E2 A B' A')
- note A = `nrm A = nrm E1 \<inter> nrm E2` `brk A = (\<lambda>l. UNIV)`
- note not_bool = `\<not> Env\<turnstile> (c ? e1 : e2)\<Colon>- (PrimT Boolean)`
- from `Env\<turnstile> B' \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A'`
+ note A = \<open>nrm A = nrm E1 \<inter> nrm E2\<close> \<open>brk A = (\<lambda>l. UNIV)\<close>
+ note not_bool = \<open>\<not> Env\<turnstile> (c ? e1 : e2)\<Colon>- (PrimT Boolean)\<close>
+ from \<open>Env\<turnstile> B' \<guillemotright>\<langle>c ? e1 : e2\<rangle>\<guillemotright> A'\<close>
obtain E1' E2'
where da_e1': "Env\<turnstile> B' \<union> assigns_if True c \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'" and
da_e2': "Env\<turnstile> B' \<union> assigns_if False c \<guillemotright>\<langle>e2\<rangle>\<guillemotright> E2'" and
@@ -1295,12 +1295,12 @@
using not_bool
by (elim da_elim_cases) (auto simp add: inj_term_simps)
(* inj_term_simps needed to handle wt (defined without \<langle>\<rangle>) *)
- note `PROP ?Hyp Env (B \<union> assigns_if True c) \<langle>e1\<rangle> E1`
- moreover note B' = `B \<subseteq> B'`
+ note \<open>PROP ?Hyp Env (B \<union> assigns_if True c) \<langle>e1\<rangle> E1\<close>
+ moreover note B' = \<open>B \<subseteq> B'\<close>
moreover note da_e1'
ultimately obtain E1': "nrm E1 \<subseteq> nrm E1'" "(\<forall>l. brk E1 l \<subseteq> brk E1' l)"
by blast
- note `PROP ?Hyp Env (B \<union> assigns_if False c) \<langle>e2\<rangle> E2`
+ note \<open>PROP ?Hyp Env (B \<union> assigns_if False c) \<langle>e2\<rangle> E2\<close>
with B' da_e2'
obtain "nrm E2 \<subseteq> nrm E2'" "(\<forall>l. brk E2 l \<subseteq> brk E2' l)"
by blast
@@ -1326,7 +1326,7 @@
next
case Cons thus ?case by (elim da_elim_cases) auto
qed
- from this [OF da' `B \<subseteq> B'`] show ?thesis .
+ from this [OF da' \<open>B \<subseteq> B'\<close>] show ?thesis .
qed
lemma da_weaken:
@@ -1342,9 +1342,9 @@
case Expr thus ?case by (iprover intro: da.Expr)
next
case (Lab Env B c C A l B')
- note `PROP ?Hyp Env B \<langle>c\<rangle>`
+ note \<open>PROP ?Hyp Env B \<langle>c\<rangle>\<close>
moreover
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
ultimately obtain C' where "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
by iprover
then obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>Break l\<bullet> c\<rangle>\<guillemotright> A'"
@@ -1352,10 +1352,10 @@
thus ?case ..
next
case (Comp Env B c1 C1 c2 C2 A B')
- note da_c1 = `Env\<turnstile> B \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1`
- note `PROP ?Hyp Env B \<langle>c1\<rangle>`
+ note da_c1 = \<open>Env\<turnstile> B \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1\<close>
+ note \<open>PROP ?Hyp Env B \<langle>c1\<rangle>\<close>
moreover
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
ultimately obtain C1' where da_c1': "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'"
by iprover
with da_c1 B'
@@ -1363,7 +1363,7 @@
"nrm C1 \<subseteq> nrm C1'"
by (rule da_monotone [elim_format]) simp
moreover
- note `PROP ?Hyp Env (nrm C1) \<langle>c2\<rangle>`
+ note \<open>PROP ?Hyp Env (nrm C1) \<langle>c2\<rangle>\<close>
ultimately obtain C2' where "Env\<turnstile> nrm C1' \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2'"
by iprover
with da_c1' obtain A' where "Env\<turnstile> B' \<guillemotright>\<langle>c1;; c2\<rangle>\<guillemotright> A'"
@@ -1371,7 +1371,7 @@
thus ?case ..
next
case (If Env B e E c1 C1 c2 C2 A B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain E' where "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
proof -
have "PROP ?Hyp Env B \<langle>e\<rangle>" by (rule If.hyps)
@@ -1405,7 +1405,7 @@
thus ?case ..
next
case (Loop Env B e E c C A l B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain E' where "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
proof -
have "PROP ?Hyp Env B \<langle>e\<rangle>" by (rule Loop.hyps)
@@ -1429,7 +1429,7 @@
thus ?case ..
next
case (Jmp jump B A Env B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
with Jmp.hyps have "jump = Ret \<longrightarrow> Result \<in> B' "
by auto
moreover
@@ -1448,7 +1448,7 @@
case Throw thus ?case by (iprover intro: da.Throw )
next
case (Try Env B c1 C1 vn C c2 C2 A B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain C1' where "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'"
proof -
have "PROP ?Hyp Env B \<langle>c1\<rangle>" by (rule Try.hyps)
@@ -1473,7 +1473,7 @@
thus ?case ..
next
case (Fin Env B c1 C1 c2 C2 A B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain C1' where C1': "Env\<turnstile> B' \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1'"
proof -
have "PROP ?Hyp Env B \<langle>c1\<rangle>" by (rule Fin.hyps)
@@ -1507,7 +1507,7 @@
case UnOp thus ?case by (iprover intro: da.UnOp)
next
case (CondAnd Env B e1 E1 e2 E2 A B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain E1' where "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
proof -
have "PROP ?Hyp Env B \<langle>e1\<rangle>" by (rule CondAnd.hyps)
@@ -1529,7 +1529,7 @@
thus ?case ..
next
case (CondOr Env B e1 E1 e2 E2 A B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain E1' where "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
proof -
have "PROP ?Hyp Env B \<langle>e1\<rangle>" by (rule CondOr.hyps)
@@ -1551,7 +1551,7 @@
thus ?case ..
next
case (BinOp Env B e1 E1 e2 A binop B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain E1' where E1': "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
proof -
have "PROP ?Hyp Env B \<langle>e1\<rangle>" by (rule BinOp.hyps)
@@ -1575,22 +1575,22 @@
thus ?case ..
next
case (Super B Env B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
with Super.hyps have "This \<in> B'"
by auto
thus ?case by (iprover intro: da.Super)
next
case (AccLVar vn B A Env B')
- note `vn \<in> B`
+ note \<open>vn \<in> B\<close>
moreover
- note `B \<subseteq> B'`
+ note \<open>B \<subseteq> B'\<close>
ultimately have "vn \<in> B'" by auto
thus ?case by (iprover intro: da.AccLVar)
next
case Acc thus ?case by (iprover intro: da.Acc)
next
case (AssLVar Env B e E A vn B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
then obtain E' where "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
by (rule AssLVar.hyps [elim_format]) iprover
then obtain A' where
@@ -1599,8 +1599,8 @@
thus ?case ..
next
case (Ass v Env B V e A B')
- note B' = `B \<subseteq> B'`
- note `\<forall>vn. v \<noteq> LVar vn`
+ note B' = \<open>B \<subseteq> B'\<close>
+ note \<open>\<forall>vn. v \<noteq> LVar vn\<close>
moreover
obtain V' where V': "Env\<turnstile> B' \<guillemotright>\<langle>v\<rangle>\<guillemotright> V'"
proof -
@@ -1625,8 +1625,8 @@
thus ?case ..
next
case (CondBool Env c e1 e2 B C E1 E2 A B')
- note B' = `B \<subseteq> B'`
- note `Env\<turnstile>(c ? e1 : e2)\<Colon>-(PrimT Boolean)`
+ note B' = \<open>B \<subseteq> B'\<close>
+ note \<open>Env\<turnstile>(c ? e1 : e2)\<Colon>-(PrimT Boolean)\<close>
moreover obtain C' where C': "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
proof -
have "PROP ?Hyp Env B \<langle>c\<rangle>" by (rule CondBool.hyps)
@@ -1661,8 +1661,8 @@
thus ?case ..
next
case (Cond Env c e1 e2 B C E1 E2 A B')
- note B' = `B \<subseteq> B'`
- note `\<not> Env\<turnstile>(c ? e1 : e2)\<Colon>-(PrimT Boolean)`
+ note B' = \<open>B \<subseteq> B'\<close>
+ note \<open>\<not> Env\<turnstile>(c ? e1 : e2)\<Colon>-(PrimT Boolean)\<close>
moreover obtain C' where C': "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'"
proof -
have "PROP ?Hyp Env B \<langle>c\<rangle>" by (rule Cond.hyps)
@@ -1697,7 +1697,7 @@
thus ?case ..
next
case (Call Env B e E args A accC statT mode mn pTs B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain E' where E': "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
proof -
have "PROP ?Hyp Env B \<langle>e\<rangle>" by (rule Call.hyps)
@@ -1723,7 +1723,7 @@
case Methd thus ?case by (iprover intro: da.Methd)
next
case (Body Env B c C A D B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain C' where C': "Env\<turnstile> B' \<guillemotright>\<langle>c\<rangle>\<guillemotright> C'" and nrm_C': "nrm C \<subseteq> nrm C'"
proof -
have "Env\<turnstile> B \<guillemotright>\<langle>c\<rangle>\<guillemotright> C" by (rule Body.hyps)
@@ -1737,10 +1737,10 @@
with da_c that show ?thesis by iprover
qed
moreover
- note `Result \<in> nrm C`
+ note \<open>Result \<in> nrm C\<close>
with nrm_C' have "Result \<in> nrm C'"
by blast
- moreover note `jumpNestingOkS {Ret} c`
+ moreover note \<open>jumpNestingOkS {Ret} c\<close>
ultimately obtain A' where
"Env\<turnstile> B' \<guillemotright>\<langle>Body D c\<rangle>\<guillemotright> A'"
by (iprover intro: da.Body)
@@ -1751,7 +1751,7 @@
case FVar thus ?case by (iprover intro: da.FVar)
next
case (AVar Env B e1 E1 e2 A B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain E1' where E1': "Env\<turnstile> B' \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1'"
proof -
have "PROP ?Hyp Env B \<langle>e1\<rangle>" by (rule AVar.hyps)
@@ -1777,7 +1777,7 @@
case Nil thus ?case by (iprover intro: da.Nil)
next
case (Cons Env B e E es A B')
- note B' = `B \<subseteq> B'`
+ note B' = \<open>B \<subseteq> B'\<close>
obtain E' where E': "Env\<turnstile> B' \<guillemotright>\<langle>e\<rangle>\<guillemotright> E'"
proof -
have "PROP ?Hyp Env B \<langle>e\<rangle>" by (rule Cons.hyps)
@@ -1800,7 +1800,7 @@
by (iprover intro: da.Cons)
thus ?case ..
qed
- from this [OF `B \<subseteq> B'`] show ?thesis .
+ from this [OF \<open>B \<subseteq> B'\<close>] show ?thesis .
qed
(* Remarks about the proof style:
--- a/src/HOL/Bali/DefiniteAssignmentCorrect.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/DefiniteAssignmentCorrect.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,4 +1,4 @@
-subsection {* Correctness of Definite Assignment *}
+subsection \<open>Correctness of Definite Assignment\<close>
theory DefiniteAssignmentCorrect imports WellForm Eval begin
@@ -104,8 +104,8 @@
"\<And> jmps' jmps. \<lbrakk>jumpNestingOkS jmps' c; jmps' \<subseteq> jmps\<rbrakk> \<Longrightarrow> jumpNestingOkS jmps c"
proof (induct rule: var.induct expr.induct stmt.induct)
case (Lab j c jmps' jmps)
- note jmpOk = `jumpNestingOkS jmps' (j\<bullet> c)`
- note jmps = `jmps' \<subseteq> jmps`
+ note jmpOk = \<open>jumpNestingOkS jmps' (j\<bullet> c)\<close>
+ note jmps = \<open>jmps' \<subseteq> jmps\<close>
with jmpOk have "jumpNestingOkS ({j} \<union> jmps') c" by simp
moreover from jmps have "({j} \<union> jmps') \<subseteq> ({j} \<union> jmps)" by auto
ultimately
@@ -135,10 +135,10 @@
by simp
next
case (Loop l e c jmps' jmps)
- from `jumpNestingOkS jmps' (l\<bullet> While(e) c)`
+ from \<open>jumpNestingOkS jmps' (l\<bullet> While(e) c)\<close>
have "jumpNestingOkS ({Cont l} \<union> jmps') c" by simp
moreover
- from `jmps' \<subseteq> jmps`
+ from \<open>jmps' \<subseteq> jmps\<close>
have "{Cont l} \<union> jmps' \<subseteq> {Cont l} \<union> jmps" by auto
ultimately
have "jumpNestingOkS ({Cont l} \<union> jmps) c"
@@ -240,7 +240,7 @@
by (cases s) (simp add: avar_def2 abrupt_if_def)
-text {*
+text \<open>
The next theorem expresses: If jumps (breaks, continues, returns) are nested
correctly, we won't find an unexpected jump in the result state of the
evaluation. For exeample, a break can't leave its enclosing loop, an return
@@ -266,7 +266,7 @@
The wellformedness of the program is used to enshure that for all
classinitialisations and methods the nesting of jumps is wellformed, too.
-*}
+\<close>
theorem jumpNestingOk_eval:
assumes eval: "G\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)"
and jmpOk: "jumpNestingOk jmps t"
@@ -287,19 +287,19 @@
(\<forall> jmps T Env.
?Jmp jmps s0 \<longrightarrow> jumpNestingOk jmps t \<longrightarrow> Env\<turnstile>t\<Colon>T \<longrightarrow> prg Env=G\<longrightarrow>
?Jmp jmps s1 \<and> ?Upd v s1)"
- -- {* Variable @{text ?HypObj} is the following goal spelled in terms of
+ \<comment> \<open>Variable \<open>?HypObj\<close> is the following goal spelled in terms of
the object logic, instead of the meta logic. It is needed in some
cases of the induction were, the atomize-rulify process of induct
does not work fine, because the eval rules mix up object and meta
- logic. See for example the case for the loop. *}
+ logic. See for example the case for the loop.\<close>
from eval
have "\<And> jmps T Env. \<lbrakk>?Jmp jmps s0; jumpNestingOk jmps t; Env\<turnstile>t\<Colon>T;prg Env=G\<rbrakk>
\<Longrightarrow> ?Jmp jmps s1 \<and> ?Upd v s1"
(is "PROP ?Hyp t s0 s1 v")
- -- {* We need to abstract over @{term jmps} since @{term jmps} are extended
+ \<comment> \<open>We need to abstract over @{term jmps} since @{term jmps} are extended
during analysis of @{term Lab}. Also we need to abstract over
@{term T} and @{term Env} since they are altered in various
- typing judgements. *}
+ typing judgements.\<close>
proof (induct)
case Abrupt thus ?case by simp
next
@@ -308,8 +308,8 @@
case Expr thus ?case by (elim wt_elim_cases) simp
next
case (Lab s0 c s1 jmp jmps T Env)
- note jmpOK = `jumpNestingOk jmps (In1r (jmp\<bullet> c))`
- note G = `prg Env = G`
+ note jmpOK = \<open>jumpNestingOk jmps (In1r (jmp\<bullet> c))\<close>
+ note G = \<open>prg Env = G\<close>
have wt_c: "Env\<turnstile>c\<Colon>\<surd>"
using Lab.prems by (elim wt_elim_cases)
{
@@ -319,7 +319,7 @@
proof -
from ab_s1 have jmp_s1: "abrupt s1 = Some (Jump j)"
by (cases s1) (simp add: absorb_def)
- note hyp_c = `PROP ?Hyp (In1r c) (Norm s0) s1 \<diamondsuit>`
+ note hyp_c = \<open>PROP ?Hyp (In1r c) (Norm s0) s1 \<diamondsuit>\<close>
from ab_s1 have "j \<noteq> jmp"
by (cases s1) (simp add: absorb_def)
moreover have "j \<in> {jmp} \<union> jmps"
@@ -337,8 +337,8 @@
thus ?case by simp
next
case (Comp s0 c1 s1 c2 s2 jmps T Env)
- note jmpOk = `jumpNestingOk jmps (In1r (c1;; c2))`
- note G = `prg Env = G`
+ note jmpOk = \<open>jumpNestingOk jmps (In1r (c1;; c2))\<close>
+ note G = \<open>prg Env = G\<close>
from Comp.prems obtain
wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and wt_c2: "Env\<turnstile>c2\<Colon>\<surd>"
by (elim wt_elim_cases)
@@ -349,11 +349,11 @@
proof -
have jmp: "?Jmp jmps s1"
proof -
- note hyp_c1 = `PROP ?Hyp (In1r c1) (Norm s0) s1 \<diamondsuit>`
+ note hyp_c1 = \<open>PROP ?Hyp (In1r c1) (Norm s0) s1 \<diamondsuit>\<close>
with wt_c1 jmpOk G
show ?thesis by simp
qed
- moreover note hyp_c2 = `PROP ?Hyp (In1r c2) s1 s2 (\<diamondsuit>::vals)`
+ moreover note hyp_c2 = \<open>PROP ?Hyp (In1r c2) s1 s2 (\<diamondsuit>::vals)\<close>
have jmpOk': "jumpNestingOk jmps (In1r c2)" using jmpOk by simp
moreover note wt_c2 G abr_s2
ultimately show "j \<in> jmps"
@@ -362,8 +362,8 @@
} thus ?case by simp
next
case (If s0 e b s1 c1 c2 s2 jmps T Env)
- note jmpOk = `jumpNestingOk jmps (In1r (If(e) c1 Else c2))`
- note G = `prg Env = G`
+ note jmpOk = \<open>jumpNestingOk jmps (In1r (If(e) c1 Else c2))\<close>
+ note G = \<open>prg Env = G\<close>
from If.prems obtain
wt_e: "Env\<turnstile>e\<Colon>-PrimT Boolean" and
wt_then_else: "Env\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
@@ -373,11 +373,11 @@
assume jmp: "abrupt s2 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note `PROP ?Hyp (In1l e) (Norm s0) s1 (In1 b)`
+ note \<open>PROP ?Hyp (In1l e) (Norm s0) s1 (In1 b)\<close>
with wt_e G have "?Jmp jmps s1"
by simp
moreover note hyp_then_else =
- `PROP ?Hyp (In1r (if the_Bool b then c1 else c2)) s1 s2 \<diamondsuit>`
+ \<open>PROP ?Hyp (In1r (if the_Bool b then c1 else c2)) s1 s2 \<diamondsuit>\<close>
have "jumpNestingOk jmps (In1r (if the_Bool b then c1 else c2))"
using jmpOk by (cases "the_Bool b") simp_all
moreover note wt_then_else G jmp
@@ -388,9 +388,9 @@
thus ?case by simp
next
case (Loop s0 e b s1 c s2 l s3 jmps T Env)
- note jmpOk = `jumpNestingOk jmps (In1r (l\<bullet> While(e) c))`
- note G = `prg Env = G`
- note wt = `Env\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T`
+ note jmpOk = \<open>jumpNestingOk jmps (In1r (l\<bullet> While(e) c))\<close>
+ note G = \<open>prg Env = G\<close>
+ note wt = \<open>Env\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T\<close>
then obtain
wt_e: "Env\<turnstile>e\<Colon>-PrimT Boolean" and
wt_c: "Env\<turnstile>c\<Colon>\<surd>"
@@ -400,7 +400,7 @@
assume jmp: "abrupt s3 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note `PROP ?Hyp (In1l e) (Norm s0) s1 (In1 b)`
+ note \<open>PROP ?Hyp (In1l e) (Norm s0) s1 (In1 b)\<close>
with wt_e G have jmp_s1: "?Jmp jmps s1"
by simp
show ?thesis
@@ -468,8 +468,8 @@
case (Jmp s j jmps T Env) thus ?case by simp
next
case (Throw s0 e a s1 jmps T Env)
- note jmpOk = `jumpNestingOk jmps (In1r (Throw e))`
- note G = `prg Env = G`
+ note jmpOk = \<open>jumpNestingOk jmps (In1r (Throw e))\<close>
+ note G = \<open>prg Env = G\<close>
from Throw.prems obtain Te where
wt_e: "Env\<turnstile>e\<Colon>-Te"
by (elim wt_elim_cases)
@@ -478,7 +478,7 @@
assume jmp: "abrupt (abupd (throw a) s1) = Some (Jump j)"
have "j\<in>jmps"
proof -
- from `PROP ?Hyp (In1l e) (Norm s0) s1 (In1 a)`
+ from \<open>PROP ?Hyp (In1l e) (Norm s0) s1 (In1 a)\<close>
have "?Jmp jmps s1" using wt_e G by simp
moreover
from jmp
@@ -490,8 +490,8 @@
thus ?case by simp
next
case (Try s0 c1 s1 s2 C vn c2 s3 jmps T Env)
- note jmpOk = `jumpNestingOk jmps (In1r (Try c1 Catch(C vn) c2))`
- note G = `prg Env = G`
+ note jmpOk = \<open>jumpNestingOk jmps (In1r (Try c1 Catch(C vn) c2))\<close>
+ note G = \<open>prg Env = G\<close>
from Try.prems obtain
wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and
wt_c2: "Env\<lparr>lcl := lcl Env(VName vn\<mapsto>Class C)\<rparr>\<turnstile>c2\<Colon>\<surd>"
@@ -501,10 +501,10 @@
assume jmp: "abrupt s3 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note `PROP ?Hyp (In1r c1) (Norm s0) s1 (\<diamondsuit>::vals)`
+ note \<open>PROP ?Hyp (In1r c1) (Norm s0) s1 (\<diamondsuit>::vals)\<close>
with jmpOk wt_c1 G
have jmp_s1: "?Jmp jmps s1" by simp
- note s2 = `G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2`
+ note s2 = \<open>G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2\<close>
show "j \<in> jmps"
proof (cases "G,s2\<turnstile>catch C")
case False
@@ -542,8 +542,8 @@
thus ?case by simp
next
case (Fin s0 c1 x1 s1 c2 s2 s3 jmps T Env)
- note jmpOk = `jumpNestingOk jmps (In1r (c1 Finally c2))`
- note G = `prg Env = G`
+ note jmpOk = \<open>jumpNestingOk jmps (In1r (c1 Finally c2))\<close>
+ note G = \<open>prg Env = G\<close>
from Fin.prems obtain
wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and wt_c2: "Env\<turnstile>c2\<Colon>\<surd>"
by (elim wt_elim_cases)
@@ -553,14 +553,14 @@
have "j \<in> jmps"
proof (cases "x1=Some (Jump j)")
case True
- note hyp_c1 = `PROP ?Hyp (In1r c1) (Norm s0) (x1,s1) \<diamondsuit>`
+ note hyp_c1 = \<open>PROP ?Hyp (In1r c1) (Norm s0) (x1,s1) \<diamondsuit>\<close>
with True jmpOk wt_c1 G show ?thesis
by - (rule hyp_c1 [THEN conjunct1,rule_format (no_asm)],simp_all)
next
case False
- note hyp_c2 = `PROP ?Hyp (In1r c2) (Norm s1) s2 \<diamondsuit>`
- note `s3 = (if \<exists>err. x1 = Some (Error err) then (x1, s1)
- else abupd (abrupt_if (x1 \<noteq> None) x1) s2)`
+ note hyp_c2 = \<open>PROP ?Hyp (In1r c2) (Norm s1) s2 \<diamondsuit>\<close>
+ note \<open>s3 = (if \<exists>err. x1 = Some (Error err) then (x1, s1)
+ else abupd (abrupt_if (x1 \<noteq> None) x1) s2)\<close>
with False jmp have "abrupt s2 = Some (Jump j)"
by (cases s2) (simp add: abrupt_if_def)
with jmpOk wt_c2 G show ?thesis
@@ -570,9 +570,9 @@
thus ?case by simp
next
case (Init C c s0 s3 s1 s2 jmps T Env)
- note `jumpNestingOk jmps (In1r (Init C))`
- note G = `prg Env = G`
- note `the (class G C) = c`
+ note \<open>jumpNestingOk jmps (In1r (Init C))\<close>
+ note G = \<open>prg Env = G\<close>
+ note \<open>the (class G C) = c\<close>
with Init.prems have c: "class G C = Some c"
by (elim wt_elim_cases) auto
{
@@ -640,15 +640,15 @@
assume jmp: "abrupt s2 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note `prg Env = G`
- moreover note hyp_init = `PROP ?Hyp (In1r (Init C)) (Norm s0) s1 \<diamondsuit>`
+ note \<open>prg Env = G\<close>
+ moreover note hyp_init = \<open>PROP ?Hyp (In1r (Init C)) (Norm s0) s1 \<diamondsuit>\<close>
moreover from wf NewC.prems
have "Env\<turnstile>(Init C)\<Colon>\<surd>"
by (elim wt_elim_cases) (drule is_acc_classD,simp)
moreover
have "abrupt s1 = Some (Jump j)"
proof -
- from `G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2` and jmp show ?thesis
+ from \<open>G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2\<close> and jmp show ?thesis
by (rule halloc_no_jump')
qed
ultimately show "j \<in> jmps"
@@ -663,20 +663,20 @@
assume jmp: "abrupt s3 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from NewA.prems
obtain wt_init: "Env\<turnstile>init_comp_ty elT\<Colon>\<surd>" and
wt_size: "Env\<turnstile>e\<Colon>-PrimT Integer"
by (elim wt_elim_cases) (auto dest: wt_init_comp_ty')
- note `PROP ?Hyp (In1r (init_comp_ty elT)) (Norm s0) s1 \<diamondsuit>`
+ note \<open>PROP ?Hyp (In1r (init_comp_ty elT)) (Norm s0) s1 \<diamondsuit>\<close>
with wt_init G
have "?Jmp jmps s1"
by (simp add: init_comp_ty_def)
moreover
- note hyp_e = `PROP ?Hyp (In1l e) s1 s2 (In1 i)`
+ note hyp_e = \<open>PROP ?Hyp (In1l e) s1 s2 (In1 i)\<close>
have "abrupt s2 = Some (Jump j)"
proof -
- note `G\<turnstile>abupd (check_neg i) s2\<midarrow>halloc Arr elT (the_Intg i)\<succ>a\<rightarrow> s3`
+ note \<open>G\<turnstile>abupd (check_neg i) s2\<midarrow>halloc Arr elT (the_Intg i)\<succ>a\<rightarrow> s3\<close>
moreover note jmp
ultimately
have "abrupt (abupd (check_neg i) s2) = Some (Jump j)"
@@ -695,14 +695,14 @@
assume jmp: "abrupt s2 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note hyp_e = `PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)`
- note `prg Env = G`
+ note hyp_e = \<open>PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)\<close>
+ note \<open>prg Env = G\<close>
moreover from Cast.prems
obtain eT where "Env\<turnstile>e\<Colon>-eT" by (elim wt_elim_cases)
moreover
have "abrupt s1 = Some (Jump j)"
proof -
- note `s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits cT) ClassCast) s1`
+ note \<open>s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits cT) ClassCast) s1\<close>
moreover note jmp
ultimately show ?thesis by (cases s1) (simp add: abrupt_if_def)
qed
@@ -718,8 +718,8 @@
assume jmp: "abrupt s1 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note hyp_e = `PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)`
- note `prg Env = G`
+ note hyp_e = \<open>PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)\<close>
+ note \<open>prg Env = G\<close>
moreover from Inst.prems
obtain eT where "Env\<turnstile>e\<Colon>-eT" by (elim wt_elim_cases)
moreover note jmp
@@ -737,8 +737,8 @@
assume jmp: "abrupt s1 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note hyp_e = `PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)`
- note `prg Env = G`
+ note hyp_e = \<open>PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)\<close>
+ note \<open>prg Env = G\<close>
moreover from UnOp.prems
obtain eT where "Env\<turnstile>e\<Colon>-eT" by (elim wt_elim_cases)
moreover note jmp
@@ -754,17 +754,17 @@
assume jmp: "abrupt s2 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from BinOp.prems
obtain e1T e2T where
wt_e1: "Env\<turnstile>e1\<Colon>-e1T" and
wt_e2: "Env\<turnstile>e2\<Colon>-e2T"
by (elim wt_elim_cases)
- note `PROP ?Hyp (In1l e1) (Norm s0) s1 (In1 v1)`
+ note \<open>PROP ?Hyp (In1l e1) (Norm s0) s1 (In1 v1)\<close>
with G wt_e1 have jmp_s1: "?Jmp jmps s1" by simp
note hyp_e2 =
- `PROP ?Hyp (if need_second_arg binop v1 then In1l e2 else In1r Skip)
- s1 s2 (In1 v2)`
+ \<open>PROP ?Hyp (if need_second_arg binop v1 then In1l e2 else In1r Skip)
+ s1 s2 (In1 v2)\<close>
show "j\<in>jmps"
proof (cases "need_second_arg binop v1")
case True with jmp_s1 wt_e2 jmp G
@@ -787,8 +787,8 @@
assume jmp: "abrupt s1 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note hyp_va = `PROP ?Hyp (In2 va) (Norm s0) s1 (In2 (v,f))`
- note `prg Env = G`
+ note hyp_va = \<open>PROP ?Hyp (In2 va) (Norm s0) s1 (In2 (v,f))\<close>
+ note \<open>prg Env = G\<close>
moreover from Acc.prems
obtain vT where "Env\<turnstile>va\<Colon>=vT" by (elim wt_elim_cases)
moreover note jmp
@@ -799,14 +799,14 @@
thus ?case by simp
next
case (Ass s0 va w f s1 e v s2 jmps T Env)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Ass.prems
obtain vT eT where
wt_va: "Env\<turnstile>va\<Colon>=vT" and
wt_e: "Env\<turnstile>e\<Colon>-eT"
by (elim wt_elim_cases)
- note hyp_v = `PROP ?Hyp (In2 va) (Norm s0) s1 (In2 (w,f))`
- note hyp_e = `PROP ?Hyp (In1l e) s1 s2 (In1 v)`
+ note hyp_v = \<open>PROP ?Hyp (In2 va) (Norm s0) s1 (In2 (w,f))\<close>
+ note hyp_e = \<open>PROP ?Hyp (In1l e) s1 s2 (In1 v)\<close>
{
fix j
assume jmp: "abrupt (assign f v s2) = Some (Jump j)"
@@ -815,7 +815,7 @@
have "abrupt s2 = Some (Jump j)"
proof (cases "normal s2")
case True
- from `G\<turnstile>s1 \<midarrow>e-\<succ>v\<rightarrow> s2` and True have nrm_s1: "normal s1"
+ from \<open>G\<turnstile>s1 \<midarrow>e-\<succ>v\<rightarrow> s2\<close> and True have nrm_s1: "normal s1"
by (rule eval_no_abrupt_lemma [rule_format])
with nrm_s1 wt_va G True
have "abrupt (f v s2) \<noteq> Some (Jump j)"
@@ -838,9 +838,9 @@
thus ?case by simp
next
case (Cond s0 e0 b s1 e1 e2 v s2 jmps T Env)
- note G = `prg Env = G`
- note hyp_e0 = `PROP ?Hyp (In1l e0) (Norm s0) s1 (In1 b)`
- note hyp_e1_e2 = `PROP ?Hyp (In1l (if the_Bool b then e1 else e2)) s1 s2 (In1 v)`
+ note G = \<open>prg Env = G\<close>
+ note hyp_e0 = \<open>PROP ?Hyp (In1l e0) (Norm s0) s1 (In1 b)\<close>
+ note hyp_e1_e2 = \<open>PROP ?Hyp (In1l (if the_Bool b then e1 else e2)) s1 s2 (In1 v)\<close>
from Cond.prems
obtain e1T e2T
where wt_e0: "Env\<turnstile>e0\<Colon>-PrimT Boolean"
@@ -873,7 +873,7 @@
next
case (Call s0 e a s1 args vs s2 D mode statT mn pTs s3 s3' accC v s4
jmps T Env)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Call.prems
obtain eT argsT
where wt_e: "Env\<turnstile>e\<Colon>-eT" and wt_args: "Env\<turnstile>args\<Colon>\<doteq>argsT"
@@ -884,26 +884,26 @@
= Some (Jump j)"
have "j\<in>jmps"
proof -
- note hyp_e = `PROP ?Hyp (In1l e) (Norm s0) s1 (In1 a)`
+ note hyp_e = \<open>PROP ?Hyp (In1l e) (Norm s0) s1 (In1 a)\<close>
from wt_e G
have jmp_s1: "?Jmp jmps s1"
by - (rule hyp_e [THEN conjunct1],simp_all)
- note hyp_args = `PROP ?Hyp (In3 args) s1 s2 (In3 vs)`
+ note hyp_args = \<open>PROP ?Hyp (In3 args) s1 s2 (In3 vs)\<close>
have "abrupt s2 = Some (Jump j)"
proof -
- note `G\<turnstile>s3' \<midarrow>Methd D \<lparr>name = mn, parTs = pTs\<rparr>-\<succ>v\<rightarrow> s4`
+ note \<open>G\<turnstile>s3' \<midarrow>Methd D \<lparr>name = mn, parTs = pTs\<rparr>-\<succ>v\<rightarrow> s4\<close>
moreover
from jmp have "abrupt s4 = Some (Jump j)"
by (cases s4) simp
ultimately have "abrupt s3' = Some (Jump j)"
by - (rule ccontr,drule (1) Methd_no_jump,simp)
- moreover note `s3' = check_method_access G accC statT mode
- \<lparr>name = mn, parTs = pTs\<rparr> a s3`
+ moreover note \<open>s3' = check_method_access G accC statT mode
+ \<lparr>name = mn, parTs = pTs\<rparr> a s3\<close>
ultimately have "abrupt s3 = Some (Jump j)"
by (cases s3)
(simp add: check_method_access_def abrupt_if_def Let_def)
moreover
- note `s3 = init_lvars G D \<lparr>name=mn, parTs=pTs\<rparr> mode a vs s2`
+ note \<open>s3 = init_lvars G D \<lparr>name=mn, parTs=pTs\<rparr> mode a vs s2\<close>
ultimately show ?thesis
by (cases s2) (auto simp add: init_lvars_def2)
qed
@@ -915,7 +915,7 @@
thus ?case by simp
next
case (Methd s0 D sig v s1 jmps T Env)
- from `G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<rightarrow> s1`
+ from \<open>G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<rightarrow> s1\<close>
have "G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<rightarrow> s1"
by (rule eval.Methd)
hence "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
@@ -934,7 +934,7 @@
thus ?case by (simp add: lvar_def Let_def)
next
case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC jmps T Env)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from wf FVar.prems
obtain statC f where
wt_e: "Env\<turnstile>e\<Colon>-Class statC" and
@@ -951,21 +951,21 @@
thus ?thesis
by simp
qed
- note fvar = `(v, s2') = fvar statDeclC stat fn a s2`
+ note fvar = \<open>(v, s2') = fvar statDeclC stat fn a s2\<close>
{
fix j
assume jmp: "abrupt s3 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note hyp_init = `PROP ?Hyp (In1r (Init statDeclC)) (Norm s0) s1 \<diamondsuit>`
+ note hyp_init = \<open>PROP ?Hyp (In1r (Init statDeclC)) (Norm s0) s1 \<diamondsuit>\<close>
from G wt_init
have "?Jmp jmps s1"
by - (rule hyp_init [THEN conjunct1],auto)
moreover
- note hyp_e = `PROP ?Hyp (In1l e) s1 s2 (In1 a)`
+ note hyp_e = \<open>PROP ?Hyp (In1l e) s1 s2 (In1 a)\<close>
have "abrupt s2 = Some (Jump j)"
proof -
- note `s3 = check_field_access G accC statDeclC fn stat a s2'`
+ note \<open>s3 = check_field_access G accC statDeclC fn stat a s2'\<close>
with jmp have "abrupt s2' = Some (Jump j)"
by (cases s2')
(simp add: check_field_access_def abrupt_if_def Let_def)
@@ -993,23 +993,23 @@
ultimately show ?case using v by simp
next
case (AVar s0 e1 a s1 e2 i s2 v s2' jmps T Env)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from AVar.prems
obtain e1T e2T where
wt_e1: "Env\<turnstile>e1\<Colon>-e1T" and wt_e2: "Env\<turnstile>e2\<Colon>-e2T"
by (elim wt_elim_cases) simp
- note avar = `(v, s2') = avar G i a s2`
+ note avar = \<open>(v, s2') = avar G i a s2\<close>
{
fix j
assume jmp: "abrupt s2' = Some (Jump j)"
have "j\<in>jmps"
proof -
- note hyp_e1 = `PROP ?Hyp (In1l e1) (Norm s0) s1 (In1 a)`
+ note hyp_e1 = \<open>PROP ?Hyp (In1l e1) (Norm s0) s1 (In1 a)\<close>
from G wt_e1
have "?Jmp jmps s1"
by - (rule hyp_e1 [THEN conjunct1], auto)
moreover
- note hyp_e2 = `PROP ?Hyp (In1l e2) s1 s2 (In1 i)`
+ note hyp_e2 = \<open>PROP ?Hyp (In1l e2) s1 s2 (In1 i)\<close>
have "abrupt s2 = Some (Jump j)"
proof -
from avar have "s2' = snd (avar G i a s2)"
@@ -1039,7 +1039,7 @@
case Nil thus ?case by simp
next
case (Cons s0 e v s1 es vs s2 jmps T Env)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Cons.prems obtain eT esT
where wt_e: "Env\<turnstile>e\<Colon>-eT" and wt_e2: "Env\<turnstile>es\<Colon>\<doteq>esT"
by (elim wt_elim_cases) simp
@@ -1048,12 +1048,12 @@
assume jmp: "abrupt s2 = Some (Jump j)"
have "j\<in>jmps"
proof -
- note hyp_e = `PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)`
+ note hyp_e = \<open>PROP ?Hyp (In1l e) (Norm s0) s1 (In1 v)\<close>
from G wt_e
have "?Jmp jmps s1"
by - (rule hyp_e [THEN conjunct1],simp_all)
moreover
- note hyp_es = `PROP ?Hyp (In3 es) s1 s2 (In3 vs)`
+ note hyp_es = \<open>PROP ?Hyp (In3 es) s1 s2 (In3 vs)\<close>
ultimately show ?thesis
using wt_e2 G jmp
by - (rule hyp_es [THEN conjunct1, rule_format (no_asm)],
@@ -1255,7 +1255,7 @@
\<subseteq> dom (locals (store (snd (avar G i a s))))"
by (cases s, simp add: avar_def2)
- text {*
+ text \<open>
Since assignments are modelled as functions from states to states, we
must take into account these functions. They appear only in the assignment
rule and as result from evaluating a variable. Thats why we need the
@@ -1268,7 +1268,7 @@
could also think of a pair of a value and a reference in the store, instead of
the generic update function. But as only array updates can cause a special
exception (if the types mismatch) and not array reads we then have to introduce
-two different rules to handle array reads and updates *}
+two different rules to handle array reads and updates\<close>
lemma dom_locals_eval_mono:
assumes eval: "G\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)"
shows "dom (locals (store s0)) \<subseteq> dom (locals (store s1)) \<and>
@@ -1334,7 +1334,7 @@
then
have s0_s1: "dom (locals (store ((Norm s0)::state)))
\<subseteq> dom (locals (store s1))" by simp
- from `G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2`
+ from \<open>G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2\<close>
have s1_s2: "dom (locals (store s1)) \<subseteq> dom (locals (store s2))"
by (rule dom_locals_sxalloc_mono)
thus ?case
@@ -1402,7 +1402,7 @@
qed
next
case (NewC s0 C s1 a s2)
- note halloc = `G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2`
+ note halloc = \<open>G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2\<close>
from NewC.hyps
have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
by simp
@@ -1412,7 +1412,7 @@
finally show ?case by simp
next
case (NewA s0 T s1 e i s2 a s3)
- note halloc = `G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3`
+ note halloc = \<open>G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3\<close>
from NewA.hyps
have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
by simp
@@ -1470,7 +1470,7 @@
finally show ?thesis by simp
next
case False
- with `G\<turnstile>s1 \<midarrow>e-\<succ>v\<rightarrow> s2`
+ with \<open>G\<turnstile>s1 \<midarrow>e-\<succ>v\<rightarrow> s2\<close>
have "s2=s1"
by auto
with s0_s1 False
@@ -1492,7 +1492,7 @@
finally show ?case by simp
next
case (Call s0 e a' s1 args vs s2 D mode statT mn pTs s3 s3' accC v s4)
- note s3 = `s3 = init_lvars G D \<lparr>name = mn, parTs = pTs\<rparr> mode a' vs s2`
+ note s3 = \<open>s3 = init_lvars G D \<lparr>name = mn, parTs = pTs\<rparr> mode a' vs s2\<close>
from Call.hyps
have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
by simp
@@ -1521,10 +1521,10 @@
also
have "\<dots> \<subseteq> dom (locals (store (abupd (absorb Ret) s3)))"
proof -
- from `s3 =
+ from \<open>s3 =
(if \<exists>l. abrupt s2 = Some (Jump (Break l)) \<or>
abrupt s2 = Some (Jump (Cont l))
- then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)`
+ then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)\<close>
show ?thesis
by simp
qed
@@ -1546,7 +1546,7 @@
by (simp add: dom_locals_fvar_vvar_mono)
hence v_ok: "(\<forall>vv. In2 v = In2 vv \<and> normal s3 \<longrightarrow> ?V_ok)"
by - (intro strip, simp)
- note s3 = `s3 = check_field_access G accC statDeclC fn stat a s2'`
+ note s3 = \<open>s3 = check_field_access G accC statDeclC fn stat a s2'\<close>
from FVar.hyps
have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
by simp
@@ -1659,7 +1659,7 @@
from eval normal show ?thesis
proof (induct)
case Abrupt thus ?case by simp
- next -- {* For statements its trivial, since then @{term "assigns t = {}"} *}
+ next \<comment> \<open>For statements its trivial, since then @{term "assigns t = {}"}\<close>
case Skip show ?case by simp
next
case Expr show ?case by simp
@@ -1685,7 +1685,7 @@
case NewC show ?case by simp
next
case (NewA s0 T s1 e i s2 a s3)
- note halloc = `G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3`
+ note halloc = \<open>G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3\<close>
have "assigns (In1l e) \<subseteq> dom (locals (store s2))"
proof -
from NewA
@@ -1728,8 +1728,8 @@
also
have "\<dots> \<subseteq> dom (locals (store s2))"
proof -
- note `G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
- else In1r Skip)\<succ>\<rightarrow> (In1 v2, s2)`
+ note \<open>G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
+ else In1r Skip)\<succ>\<rightarrow> (In1 v2, s2)\<close>
thus ?thesis
by (rule dom_locals_eval_mono_elim)
qed
@@ -1752,7 +1752,7 @@
case Acc thus ?case by simp
next
case (Ass s0 va w f s1 e v s2)
- note nrm_ass_s2 = `normal (assign f v s2)`
+ note nrm_ass_s2 = \<open>normal (assign f v s2)\<close>
hence nrm_s2: "normal s2"
by (cases s2, simp add: assign_def Let_def)
with Ass.hyps
@@ -1843,16 +1843,16 @@
case (Call s0 e a' s1 args vs s2 D mode statT mn pTs s3 s3' accC v s4)
have nrm_s2: "normal s2"
proof -
- from `normal ((set_lvars (locals (snd s2))) s4)`
+ from \<open>normal ((set_lvars (locals (snd s2))) s4)\<close>
have normal_s4: "normal s4" by simp
hence "normal s3'" using Call.hyps
by - (erule eval_no_abrupt_lemma [rule_format])
moreover note
- `s3' = check_method_access G accC statT mode \<lparr>name=mn, parTs=pTs\<rparr> a' s3`
+ \<open>s3' = check_method_access G accC statT mode \<lparr>name=mn, parTs=pTs\<rparr> a' s3\<close>
ultimately have "normal s3"
by (cases s3) (simp add: check_method_access_def Let_def)
moreover
- note s3 = `s3 = init_lvars G D \<lparr>name = mn, parTs = pTs\<rparr> mode a' vs s2`
+ note s3 = \<open>s3 = init_lvars G D \<lparr>name = mn, parTs = pTs\<rparr> mode a' vs s2\<close>
ultimately show "normal s2"
by (cases s2) (simp add: init_lvars_def2)
qed
@@ -1883,11 +1883,11 @@
case LVar thus ?case by simp
next
case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC)
- note s3 = `s3 = check_field_access G accC statDeclC fn stat a s2'`
- note avar = `(v, s2') = fvar statDeclC stat fn a s2`
+ note s3 = \<open>s3 = check_field_access G accC statDeclC fn stat a s2'\<close>
+ note avar = \<open>(v, s2') = fvar statDeclC stat fn a s2\<close>
have nrm_s2: "normal s2"
proof -
- note `normal s3`
+ note \<open>normal s3\<close>
with s3 have "normal s2'"
by (cases s2') (simp add: check_field_access_def Let_def)
with avar show "normal s2"
@@ -1912,10 +1912,10 @@
by simp
next
case (AVar s0 e1 a s1 e2 i s2 v s2')
- note avar = `(v, s2') = avar G i a s2`
+ note avar = \<open>(v, s2') = avar G i a s2\<close>
have nrm_s2: "normal s2"
proof -
- from avar and `normal s2'`
+ from avar and \<open>normal s2'\<close>
show ?thesis by (cases s2) (simp add: avar_def2)
qed
with AVar.hyps
@@ -2018,17 +2018,17 @@
case Inst hence False by simp thus ?case ..
next
case (Lit val c v s0 s)
- note `constVal (Lit val) = Some c`
+ note \<open>constVal (Lit val) = Some c\<close>
moreover
- from `G\<turnstile>Norm s0 \<midarrow>Lit val-\<succ>v\<rightarrow> s`
+ from \<open>G\<turnstile>Norm s0 \<midarrow>Lit val-\<succ>v\<rightarrow> s\<close>
obtain "v=val" and "normal s"
by cases simp
ultimately show "v=c \<and> normal s" by simp
next
case (UnOp unop e c v s0 s)
- note const = `constVal (UnOp unop e) = Some c`
+ note const = \<open>constVal (UnOp unop e) = Some c\<close>
then obtain ce where ce: "constVal e = Some ce" by simp
- from `G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>v\<rightarrow> s`
+ from \<open>G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>v\<rightarrow> s\<close>
obtain ve where ve: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>ve\<rightarrow> s" and
v: "v = eval_unop unop ve"
by cases simp
@@ -2042,12 +2042,12 @@
show ?case ..
next
case (BinOp binop e1 e2 c v s0 s)
- note const = `constVal (BinOp binop e1 e2) = Some c`
+ note const = \<open>constVal (BinOp binop e1 e2) = Some c\<close>
then obtain c1 c2 where c1: "constVal e1 = Some c1" and
c2: "constVal e2 = Some c2" and
c: "c = eval_binop binop c1 c2"
by simp
- from `G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>v\<rightarrow> s`
+ from \<open>G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>v\<rightarrow> s\<close>
obtain v1 s1 v2
where v1: "G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1" and
v2: "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
@@ -2089,13 +2089,13 @@
case Ass hence False by simp thus ?case ..
next
case (Cond b e1 e2 c v s0 s)
- note c = `constVal (b ? e1 : e2) = Some c`
+ note c = \<open>constVal (b ? e1 : e2) = Some c\<close>
then obtain cb c1 c2 where
cb: "constVal b = Some cb" and
c1: "constVal e1 = Some c1" and
c2: "constVal e2 = Some c2"
by (auto split: bool.splits)
- from `G\<turnstile>Norm s0 \<midarrow>b ? e1 : e2-\<succ>v\<rightarrow> s`
+ from \<open>G\<turnstile>Norm s0 \<midarrow>b ? e1 : e2-\<succ>v\<rightarrow> s\<close>
obtain vb s1
where vb: "G\<turnstile>Norm s0 \<midarrow>b-\<succ>vb\<rightarrow> s1" and
eval_v: "G\<turnstile>s1 \<midarrow>(if the_Bool vb then e1 else e2)-\<succ>v\<rightarrow> s"
@@ -2166,27 +2166,27 @@
case Inst hence False by simp thus ?case ..
next
case (Lit v c)
- from `constVal (Lit v) = Some c`
+ from \<open>constVal (Lit v) = Some c\<close>
have "c=v" by simp
moreover
- from `Env\<turnstile>Lit v\<Colon>-PrimT Boolean`
+ from \<open>Env\<turnstile>Lit v\<Colon>-PrimT Boolean\<close>
have "typeof empty_dt v = Some (PrimT Boolean)"
by cases simp
ultimately show ?case by simp
next
case (UnOp unop e c)
- from `Env\<turnstile>UnOp unop e\<Colon>-PrimT Boolean`
+ from \<open>Env\<turnstile>UnOp unop e\<Colon>-PrimT Boolean\<close>
have "Boolean = unop_type unop" by cases simp
moreover
- from `constVal (UnOp unop e) = Some c`
+ from \<open>constVal (UnOp unop e) = Some c\<close>
obtain ce where "c = eval_unop unop ce" by auto
ultimately show ?case by (simp add: eval_unop_type)
next
case (BinOp binop e1 e2 c)
- from `Env\<turnstile>BinOp binop e1 e2\<Colon>-PrimT Boolean`
+ from \<open>Env\<turnstile>BinOp binop e1 e2\<Colon>-PrimT Boolean\<close>
have "Boolean = binop_type binop" by cases simp
moreover
- from `constVal (BinOp binop e1 e2) = Some c`
+ from \<open>constVal (BinOp binop e1 e2) = Some c\<close>
obtain c1 c2 where "c = eval_binop binop c1 c2" by auto
ultimately show ?case by (simp add: eval_binop_type)
next
@@ -2197,13 +2197,13 @@
case Ass hence False by simp thus ?case ..
next
case (Cond b e1 e2 c)
- note c = `constVal (b ? e1 : e2) = Some c`
+ note c = \<open>constVal (b ? e1 : e2) = Some c\<close>
then obtain cb c1 c2 where
cb: "constVal b = Some cb" and
c1: "constVal e1 = Some c1" and
c2: "constVal e2 = Some c2"
by (auto split: bool.splits)
- note wt = `Env\<turnstile>b ? e1 : e2\<Colon>-PrimT Boolean`
+ note wt = \<open>Env\<turnstile>b ? e1 : e2\<Colon>-PrimT Boolean\<close>
then
obtain T1 T2
where "Env\<turnstile>b\<Colon>-PrimT Boolean" and
@@ -2239,8 +2239,8 @@
bool: "Env\<turnstile> e\<Colon>-PrimT Boolean"
shows "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
proof -
- -- {* To properly perform induction on the evaluation relation we have to
- generalize the lemma to terms not only expressions. *}
+ \<comment> \<open>To properly perform induction on the evaluation relation we have to
+ generalize the lemma to terms not only expressions.\<close>
{ fix t val
assume eval': "prg Env\<turnstile> s0 \<midarrow>t\<succ>\<rightarrow> (val,s1)"
assume bool': "Env\<turnstile> t\<Colon>Inl (PrimT Boolean)"
@@ -2252,26 +2252,26 @@
case Abrupt thus ?case by simp
next
case (NewC s0 C s1 a s2)
- from `Env\<turnstile>NewC C\<Colon>-PrimT Boolean`
+ from \<open>Env\<turnstile>NewC C\<Colon>-PrimT Boolean\<close>
have False
by cases simp
thus ?case ..
next
case (NewA s0 T s1 e i s2 a s3)
- from `Env\<turnstile>New T[e]\<Colon>-PrimT Boolean`
+ from \<open>Env\<turnstile>New T[e]\<Colon>-PrimT Boolean\<close>
have False
by cases simp
thus ?case ..
next
case (Cast s0 e b s1 s2 T)
- note s2 = `s2 = abupd (raise_if (\<not> prg Env,snd s1\<turnstile>b fits T) ClassCast) s1`
+ note s2 = \<open>s2 = abupd (raise_if (\<not> prg Env,snd s1\<turnstile>b fits T) ClassCast) s1\<close>
have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
proof -
- from s2 and `normal s2`
+ from s2 and \<open>normal s2\<close>
have "normal s1"
by (cases s1) simp
moreover
- from `Env\<turnstile>Cast T e\<Colon>-PrimT Boolean`
+ from \<open>Env\<turnstile>Cast T e\<Colon>-PrimT Boolean\<close>
have "Env\<turnstile>e\<Colon>- PrimT Boolean"
by cases (auto dest: cast_Boolean2)
ultimately show ?thesis
@@ -2283,14 +2283,14 @@
finally show ?case by simp
next
case (Inst s0 e v s1 b T)
- from `prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1` and `normal s1`
+ from \<open>prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<close> and \<open>normal s1\<close>
have "assignsE e \<subseteq> dom (locals (store s1))"
by (rule assignsE_good_approx)
thus ?case
by simp
next
case (Lit s v)
- from `Env\<turnstile>Lit v\<Colon>-PrimT Boolean`
+ from \<open>Env\<turnstile>Lit v\<Colon>-PrimT Boolean\<close>
have "typeof empty_dt v = Some (PrimT Boolean)"
by cases simp
then obtain b where "v=Bool b"
@@ -2299,13 +2299,13 @@
by simp
next
case (UnOp s0 e v s1 unop)
- note bool = `Env\<turnstile>UnOp unop e\<Colon>-PrimT Boolean`
+ note bool = \<open>Env\<turnstile>UnOp unop e\<Colon>-PrimT Boolean\<close>
hence bool_e: "Env\<turnstile>e\<Colon>-PrimT Boolean"
by cases (cases unop,simp_all)
show ?case
proof (cases "constVal (UnOp unop e)")
case None
- note `normal s1`
+ note \<open>normal s1\<close>
moreover note bool_e
ultimately have "assigns_if (the_Bool v) e \<subseteq> dom (locals (store s1))"
by (rule UnOp.hyps [elim_format]) auto
@@ -2321,7 +2321,7 @@
next
case (Some c)
moreover
- from `prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1`
+ from \<open>prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<close>
have "prg Env\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>eval_unop unop v\<rightarrow> s1"
by (rule eval.UnOp)
with Some
@@ -2339,7 +2339,7 @@
qed
next
case (BinOp s0 e1 v1 s1 binop e2 v2 s2)
- note bool = `Env\<turnstile>BinOp binop e1 e2\<Colon>-PrimT Boolean`
+ note bool = \<open>Env\<turnstile>BinOp binop e1 e2\<Colon>-PrimT Boolean\<close>
show ?case
proof (cases "constVal (BinOp binop e1 e2)")
case (Some c)
@@ -2392,7 +2392,7 @@
with condAnd
have need_second: "need_second_arg binop v1"
by (simp add: need_second_arg_def)
- from `normal s2`
+ from \<open>normal s2\<close>
have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
by (rule BinOp.hyps [elim_format])
(simp add: need_second bool_e2)+
@@ -2412,7 +2412,7 @@
obtain "the_Bool v1=True" and "the_Bool v2 = False"
by (simp add: need_second_arg_def)
moreover
- from `normal s2`
+ from \<open>normal s2\<close>
have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
by (rule BinOp.hyps [elim_format]) (simp add: True bool_e2)+
with e1_s2
@@ -2440,7 +2440,7 @@
with condOr
have need_second: "need_second_arg binop v1"
by (simp add: need_second_arg_def)
- from `normal s2`
+ from \<open>normal s2\<close>
have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
by (rule BinOp.hyps [elim_format])
(simp add: need_second bool_e2)+
@@ -2460,7 +2460,7 @@
obtain "the_Bool v1=False" and "the_Bool v2 = True"
by (simp add: need_second_arg_def)
moreover
- from `normal s2`
+ from \<open>normal s2\<close>
have "assigns_if (the_Bool v2) e2 \<subseteq> dom (locals (store s2))"
by (rule BinOp.hyps [elim_format]) (simp add: True bool_e2)+
with e1_s2
@@ -2483,12 +2483,12 @@
qed
next
case False
- note `\<not> (binop = CondAnd \<or> binop = CondOr)`
+ note \<open>\<not> (binop = CondAnd \<or> binop = CondOr)\<close>
from BinOp.hyps
have
"prg Env\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>eval_binop binop v1 v2\<rightarrow> s2"
by - (rule eval.BinOp)
- moreover note `normal s2`
+ moreover note \<open>normal s2\<close>
ultimately
have "assignsE (BinOp binop e1 e2) \<subseteq> dom (locals (store s2))"
by (rule assignsE_good_approx)
@@ -2499,13 +2499,13 @@
qed
next
case Super
- note `Env\<turnstile>Super\<Colon>-PrimT Boolean`
+ note \<open>Env\<turnstile>Super\<Colon>-PrimT Boolean\<close>
hence False
by cases simp
thus ?case ..
next
case (Acc s0 va v f s1)
- from `prg Env\<turnstile>Norm s0 \<midarrow>va=\<succ>(v, f)\<rightarrow> s1` and `normal s1`
+ from \<open>prg Env\<turnstile>Norm s0 \<midarrow>va=\<succ>(v, f)\<rightarrow> s1\<close> and \<open>normal s1\<close>
have "assignsV va \<subseteq> dom (locals (store s1))"
by (rule assignsV_good_approx)
thus ?case by simp
@@ -2513,23 +2513,23 @@
case (Ass s0 va w f s1 e v s2)
hence "prg Env\<turnstile>Norm s0 \<midarrow>va := e-\<succ>v\<rightarrow> assign f v s2"
by - (rule eval.Ass)
- moreover note `normal (assign f v s2)`
+ moreover note \<open>normal (assign f v s2)\<close>
ultimately
have "assignsE (va := e) \<subseteq> dom (locals (store (assign f v s2)))"
by (rule assignsE_good_approx)
thus ?case by simp
next
case (Cond s0 e0 b s1 e1 e2 v s2)
- from `Env\<turnstile>e0 ? e1 : e2\<Colon>-PrimT Boolean`
+ from \<open>Env\<turnstile>e0 ? e1 : e2\<Colon>-PrimT Boolean\<close>
obtain wt_e1: "Env\<turnstile>e1\<Colon>-PrimT Boolean" and
wt_e2: "Env\<turnstile>e2\<Colon>-PrimT Boolean"
by cases (auto dest: widen_Boolean2)
- note eval_e0 = `prg Env\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1`
+ note eval_e0 = \<open>prg Env\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1\<close>
have e0_s2: "assignsE e0 \<subseteq> dom (locals (store s2))"
proof -
note eval_e0
moreover
- from Cond.hyps and `normal s2` have "normal s1"
+ from Cond.hyps and \<open>normal s2\<close> have "normal s1"
by - (erule eval_no_abrupt_lemma [rule_format],simp)
ultimately
have "assignsE e0 \<subseteq> dom (locals (store s1))"
@@ -2547,14 +2547,14 @@
\<subseteq> dom (locals (store s2))"
proof (cases "the_Bool b")
case True
- from `normal s2`
+ from \<open>normal s2\<close>
have "assigns_if (the_Bool v) e1 \<subseteq> dom (locals (store s2))"
by (rule Cond.hyps [elim_format]) (simp_all add: wt_e1 True)
thus ?thesis
by blast
next
case False
- from `normal s2`
+ from \<open>normal s2\<close>
have "assigns_if (the_Bool v) e2 \<subseteq> dom (locals (store s2))"
by (rule Cond.hyps [elim_format]) (simp_all add: wt_e2 False)
thus ?thesis
@@ -2574,7 +2574,7 @@
show ?thesis
proof (cases "the_Bool c")
case True
- from `normal s2`
+ from \<open>normal s2\<close>
have "assigns_if (the_Bool v) e1 \<subseteq> dom (locals (store s2))"
by (rule Cond.hyps [elim_format]) (simp_all add: eq_b_c True wt_e1)
with e0_s2
@@ -2584,7 +2584,7 @@
by simp
next
case False
- from `normal s2`
+ from \<open>normal s2\<close>
have "assigns_if (the_Bool v) e2 \<subseteq> dom (locals (store s2))"
by (rule Cond.hyps [elim_format]) (simp_all add: eq_b_c False wt_e2)
with e0_s2
@@ -2602,14 +2602,14 @@
by - (rule eval.Call)
hence "assignsE ({accC,statT,mode}e\<cdot>mn( {pTs}args))
\<subseteq> dom (locals (store ((set_lvars (locals (store s2))) s4)))"
- using `normal ((set_lvars (locals (snd s2))) s4)`
+ using \<open>normal ((set_lvars (locals (snd s2))) s4)\<close>
by (rule assignsE_good_approx)
thus ?case by simp
next
case Methd show ?case by simp
next
case Body show ?case by simp
- qed simp+ -- {* all the statements and variables *}
+ qed simp+ \<comment> \<open>all the statements and variables\<close>
}
note generalized = this
from eval bool show ?thesis
@@ -2653,7 +2653,7 @@
let ?HypObj = "\<lambda> t s0 s1.
\<forall> Env T A. ?Wt Env t T \<longrightarrow> ?Da Env s0 t A \<longrightarrow> prg Env = G
\<longrightarrow> ?NormalAssigned s1 A \<and> ?BreakAssigned s0 s1 A \<and> ?ResAssigned s0 s1"
- -- {* Goal in object logic variant *}
+ \<comment> \<open>Goal in object logic variant\<close>
let ?Hyp = "\<lambda>t s0 s1. (\<And> Env T A. \<lbrakk>?Wt Env t T; ?Da Env s0 t A; prg Env = G\<rbrakk>
\<Longrightarrow> ?NormalAssigned s1 A \<and> ?BreakAssigned s0 s1 A \<and> ?ResAssigned s0 s1)"
from eval and wt da G
@@ -2692,7 +2692,7 @@
(rule Expr.hyps, auto)
next
case (Lab s0 c s1 j Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Lab.prems
obtain C l where
da_c: "Env\<turnstile> dom (locals (snd (Norm s0))) \<guillemotright>\<langle>c\<rangle>\<guillemotright> C" and
@@ -2751,7 +2751,7 @@
ultimately show ?case by (intro conjI)
next
case (Comp s0 c1 s1 c2 s2 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Comp.prems
obtain C1 C2
where da_c1: "Env\<turnstile> dom (locals (snd (Norm s0))) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1" and
@@ -2762,7 +2762,7 @@
obtain wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and
wt_c2: "Env\<turnstile>c2\<Colon>\<surd>"
by (elim wt_elim_cases) simp
- note `PROP ?Hyp (In1r c1) (Norm s0) s1`
+ note \<open>PROP ?Hyp (In1r c1) (Norm s0) s1\<close>
with wt_c1 da_c1 G
obtain nrm_c1: "?NormalAssigned s1 C1" and
brk_c1: "?BreakAssigned (Norm s0) s1 C1" and
@@ -2777,7 +2777,7 @@
nrm_c2: "nrm C2 \<subseteq> nrm C2'" and
brk_c2: "\<forall> l. brk C2 l \<subseteq> brk C2' l"
by (rule da_weakenE) iprover
- note `PROP ?Hyp (In1r c2) s1 s2`
+ note \<open>PROP ?Hyp (In1r c2) s1 s2\<close>
with wt_c2 da_c2' G
obtain nrm_c2': "?NormalAssigned s2 C2'" and
brk_c2': "?BreakAssigned s1 s2 C2'" and
@@ -2797,7 +2797,7 @@
ultimately show ?thesis by (intro conjI)
next
case False
- with `G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2`
+ with \<open>G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2\<close>
have eq_s1_s2: "s2=s1" by auto
with False have "?NormalAssigned s2 A" by blast
moreover
@@ -2824,7 +2824,7 @@
qed
next
case (If s0 e b s1 c1 c2 s2 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
with If.hyps have eval_e: "prg Env \<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1" by simp
from If.prems
obtain E C1 C2 where
@@ -2920,7 +2920,7 @@
moreover
have "s2 = s1"
proof -
- from abr and `G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2`
+ from abr and \<open>G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2\<close>
show ?thesis
by (cases s1) simp
qed
@@ -2928,7 +2928,7 @@
qed
next
case (Loop s0 e b s1 c s2 l s3 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
with Loop.hyps have eval_e: "prg Env\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1"
by (simp (no_asm_simp))
from Loop.prems
@@ -3134,7 +3134,7 @@
ultimately show ?case by (intro conjI)
next
case (Throw s0 e a s1 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Throw.prems obtain E where
da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> E"
by (elim da_elim_cases)
@@ -3164,7 +3164,7 @@
ultimately show ?case by (intro conjI)
next
case (Try s0 c1 s1 s2 C vn c2 s3 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Try.prems obtain C1 C2 where
da_c1: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1" and
da_c2:
@@ -3178,7 +3178,7 @@
by (elim wt_elim_cases)
have sxalloc: "prg Env\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" using Try.hyps G
by (simp (no_asm_simp))
- note `PROP ?Hyp (In1r c1) (Norm s0) s1`
+ note \<open>PROP ?Hyp (In1r c1) (Norm s0) s1\<close>
with wt_c1 da_c1 G
obtain nrm_C1: "?NormalAssigned s1 C1" and
brk_C1: "?BreakAssigned (Norm s0) s1 C1" and
@@ -3236,7 +3236,7 @@
have "(dom (locals (store ((Norm s0)::state))) \<union> {VName vn})
\<subseteq> dom (locals (store (new_xcpt_var vn s2)))"
proof -
- from `G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1`
+ from \<open>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1\<close>
have "dom (locals (store ((Norm s0)::state)))
\<subseteq> dom (locals (store s1))"
by (rule dom_locals_eval_mono_elim)
@@ -3311,7 +3311,7 @@
qed
next
case (Fin s0 c1 x1 s1 c2 s2 s3 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Fin.prems obtain C1 C2 where
da_C1: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>c1\<rangle>\<guillemotright> C1" and
da_C2: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>c2\<rangle>\<guillemotright> C2" and
@@ -3322,7 +3322,7 @@
wt_c1: "Env\<turnstile>c1\<Colon>\<surd>" and
wt_c2: "Env\<turnstile>c2\<Colon>\<surd>"
by (elim wt_elim_cases)
- note `PROP ?Hyp (In1r c1) (Norm s0) (x1,s1)`
+ note \<open>PROP ?Hyp (In1r c1) (Norm s0) (x1,s1)\<close>
with wt_c1 da_C1 G
obtain nrmAss_C1: "?NormalAssigned (x1,s1) C1" and
brkAss_C1: "?BreakAssigned (Norm s0) (x1,s1) C1" and
@@ -3342,7 +3342,7 @@
nrm_C2': "nrm C2 \<subseteq> nrm C2'" and
brk_C2': "\<forall> l. brk C2 l \<subseteq> brk C2' l"
by (rule da_weakenE) simp
- note `PROP ?Hyp (In1r c2) (Norm s1) s2`
+ note \<open>PROP ?Hyp (In1r c2) (Norm s1) s2\<close>
with wt_c2 da_C2' G
obtain nrmAss_C2': "?NormalAssigned s2 C2'" and
brkAss_C2': "?BreakAssigned (Norm s1) s2 C2'" and
@@ -3357,11 +3357,11 @@
show ?thesis
using that resAss_s2' by simp
qed
- note s3 = `s3 = (if \<exists>err. x1 = Some (Error err) then (x1, s1)
- else abupd (abrupt_if (x1 \<noteq> None) x1) s2)`
+ note s3 = \<open>s3 = (if \<exists>err. x1 = Some (Error err) then (x1, s1)
+ else abupd (abrupt_if (x1 \<noteq> None) x1) s2)\<close>
have s1_s2: "dom (locals s1) \<subseteq> dom (locals (store s2))"
proof -
- from `G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2`
+ from \<open>G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2\<close>
show ?thesis
by (rule dom_locals_eval_mono_elim) simp
qed
@@ -3470,7 +3470,7 @@
ultimately show ?case by (intro conjI)
next
case (Init C c s0 s3 s1 s2 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Init.hyps
have eval: "prg Env\<turnstile> Norm s0 \<midarrow>Init C\<rightarrow> s3"
apply (simp only: G)
@@ -3480,7 +3480,7 @@
apply (simp only: if_False )
apply (elim conjE,intro conjI,assumption+,simp)
done (* auto or simp alone always do too much *)
- from Init.prems and `the (class G C) = c`
+ from Init.prems and \<open>the (class G C) = c\<close>
have c: "class G C = Some c"
by (elim wt_elim_cases) auto
from Init.prems obtain
@@ -3525,7 +3525,7 @@
qed
next
case (NewC s0 C s1 a s2 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from NewC.prems
obtain A: "nrm A = dom (locals (store ((Norm s0)::state)))"
"brk A = (\<lambda> l. UNIV)"
@@ -3565,7 +3565,7 @@
ultimately show ?case by (intro conjI)
next
case (NewA s0 elT s1 e i s2 a s3 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from NewA.prems obtain
da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
by (elim da_elim_cases)
@@ -3573,7 +3573,7 @@
wt_init: "Env\<turnstile>init_comp_ty elT\<Colon>\<surd>" and
wt_size: "Env\<turnstile>e\<Colon>-PrimT Integer"
by (elim wt_elim_cases) (auto dest: wt_init_comp_ty')
- note halloc = `G\<turnstile>abupd (check_neg i) s2\<midarrow>halloc Arr elT (the_Intg i)\<succ>a\<rightarrow>s3`
+ note halloc = \<open>G\<turnstile>abupd (check_neg i) s2\<midarrow>halloc Arr elT (the_Intg i)\<succ>a\<rightarrow>s3\<close>
have "dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
by (rule dom_locals_eval_mono_elim) (rule NewA.hyps)
with da_e obtain A' where
@@ -3581,7 +3581,7 @@
and nrm_A_A': "nrm A \<subseteq> nrm A'"
and brk_A_A': "\<forall> l. brk A l \<subseteq> brk A' l"
by (rule da_weakenE) simp
- note `PROP ?Hyp (In1l e) s1 s2`
+ note \<open>PROP ?Hyp (In1l e) s1 s2\<close>
with wt_size da_e' G obtain
nrmAss_A': "?NormalAssigned s2 A'" and
brkAss_A': "?BreakAssigned s1 s2 A'"
@@ -3630,19 +3630,19 @@
ultimately show ?case by (intro conjI)
next
case (Cast s0 e v s1 s2 cT Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Cast.prems obtain
da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
by (elim da_elim_cases)
from Cast.prems obtain eT where
wt_e: "Env\<turnstile>e\<Colon>-eT"
by (elim wt_elim_cases)
- note `PROP ?Hyp (In1l e) (Norm s0) s1`
+ note \<open>PROP ?Hyp (In1l e) (Norm s0) s1\<close>
with wt_e da_e G obtain
nrmAss_A: "?NormalAssigned s1 A" and
brkAss_A: "?BreakAssigned (Norm s0) s1 A"
by simp
- note s2 = `s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits cT) ClassCast) s1`
+ note s2 = \<open>s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits cT) ClassCast) s1\<close>
hence s1_s2: "dom (locals (store s1)) \<subseteq> dom (locals (store s2))"
by simp
have "?NormalAssigned s2 A"
@@ -3675,14 +3675,14 @@
ultimately show ?case by (intro conjI)
next
case (Inst s0 e v s1 b iT Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Inst.prems obtain
da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
by (elim da_elim_cases)
from Inst.prems obtain eT where
wt_e: "Env\<turnstile>e\<Colon>-eT"
by (elim wt_elim_cases)
- note `PROP ?Hyp (In1l e) (Norm s0) s1`
+ note \<open>PROP ?Hyp (In1l e) (Norm s0) s1\<close>
with wt_e da_e G obtain
"?NormalAssigned s1 A" and
"?BreakAssigned (Norm s0) s1 A" and
@@ -3697,14 +3697,14 @@
thus ?case by simp
next
case (UnOp s0 e v s1 unop Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from UnOp.prems obtain
da_e: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
by (elim da_elim_cases)
from UnOp.prems obtain eT where
wt_e: "Env\<turnstile>e\<Colon>-eT"
by (elim wt_elim_cases)
- note `PROP ?Hyp (In1l e) (Norm s0) s1`
+ note \<open>PROP ?Hyp (In1l e) (Norm s0) s1\<close>
with wt_e da_e G obtain
"?NormalAssigned s1 A" and
"?BreakAssigned (Norm s0) s1 A" and
@@ -3713,7 +3713,7 @@
thus ?case by (intro conjI)
next
case (BinOp s0 e1 v1 s1 binop e2 v2 s2 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from BinOp.hyps
have
eval: "prg Env\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)\<rightarrow> s2"
@@ -3828,7 +3828,7 @@
where da_e1: "Env\<turnstile> dom (locals (snd (Norm s0))) \<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1"
and da_e2: "Env\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A"
by (elim da_elim_cases) (simp_all add: notAndOr)
- note `PROP ?Hyp (In1l e1) (Norm s0) s1`
+ note \<open>PROP ?Hyp (In1l e1) (Norm s0) s1\<close>
with wt_e1 da_e1 G normal_s1
obtain "?NormalAssigned s1 E1"
by simp
@@ -3880,20 +3880,20 @@
have "nrm A = dom (locals (store ((Norm s0)::state)))"
by (simp only: vn) (elim da_elim_cases,simp_all)
moreover
- from `G\<turnstile>Norm s0 \<midarrow>v=\<succ>(w, upd)\<rightarrow> s1`
+ from \<open>G\<turnstile>Norm s0 \<midarrow>v=\<succ>(w, upd)\<rightarrow> s1\<close>
have "s1=Norm s0"
by (simp only: vn) (elim eval_elim_cases,simp)
ultimately show ?thesis by simp
next
case False
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from False Acc.prems
have da_v: "Env\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>\<langle>v\<rangle>\<guillemotright> A"
by (elim da_elim_cases) simp_all
from Acc.prems obtain vT where
wt_v: "Env\<turnstile>v\<Colon>=vT"
by (elim wt_elim_cases)
- note `PROP ?Hyp (In2 v) (Norm s0) s1`
+ note \<open>PROP ?Hyp (In2 v) (Norm s0) s1\<close>
with wt_v da_v G obtain
"?NormalAssigned s1 A" and
"?BreakAssigned (Norm s0) s1 A" and
@@ -3903,7 +3903,7 @@
qed
next
case (Ass s0 var w upd s1 e v s2 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Ass.prems obtain varT eT where
wt_var: "Env\<turnstile>var\<Colon>=varT" and
wt_e: "Env\<turnstile>e\<Colon>-eT"
@@ -3918,8 +3918,8 @@
by (cases s2) (simp add: assign_def Let_def)
hence normal_s1: "normal s1"
by - (rule eval_no_abrupt_lemma [rule_format], rule Ass.hyps)
- note hyp_var = `PROP ?Hyp (In2 var) (Norm s0) s1`
- note hyp_e = `PROP ?Hyp (In1l e) s1 s2`
+ note hyp_var = \<open>PROP ?Hyp (In2 var) (Norm s0) s1\<close>
+ note hyp_e = \<open>PROP ?Hyp (In1l e) s1 s2\<close>
show "nrm A \<subseteq> dom (locals (store (assign upd v s2)))"
proof (cases "\<exists> vn. var = LVar vn")
case True
@@ -4017,7 +4017,7 @@
ultimately show ?case by (intro conjI)
next
case (Cond s0 e0 b s1 e1 e2 v s2 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
have "?NormalAssigned s2 A"
proof
assume normal_s2: "normal s2"
@@ -4140,7 +4140,7 @@
next
case (Call s0 e a s1 args vs s2 D mode statT mn pTs s3 s3' accC v s4
Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
have "?NormalAssigned (restore_lvars s2 s4) A"
proof
assume normal_restore_lvars: "normal (restore_lvars s2 s4)"
@@ -4154,9 +4154,9 @@
wt_e: "Env\<turnstile>e\<Colon>-eT" and
wt_args: "Env\<turnstile>args\<Colon>\<doteq>argsT"
by (elim wt_elim_cases)
- note s3 = `s3 = init_lvars G D \<lparr>name = mn, parTs = pTs\<rparr> mode a vs s2`
- note s3' = `s3' = check_method_access G accC statT mode
- \<lparr>name=mn,parTs=pTs\<rparr> a s3`
+ note s3 = \<open>s3 = init_lvars G D \<lparr>name = mn, parTs = pTs\<rparr> mode a vs s2\<close>
+ note s3' = \<open>s3' = check_method_access G accC statT mode
+ \<lparr>name=mn,parTs=pTs\<rparr> a s3\<close>
have normal_s2: "normal s2"
proof -
from normal_restore_lvars have "normal s4"
@@ -4170,7 +4170,7 @@
qed
then have normal_s1: "normal s1"
by - (rule eval_no_abrupt_lemma [rule_format], rule Call.hyps)
- note `PROP ?Hyp (In1l e) (Norm s0) s1`
+ note \<open>PROP ?Hyp (In1l e) (Norm s0) s1\<close>
with da_e wt_e G normal_s1
have "nrm E \<subseteq> dom (locals (store s1))"
by simp
@@ -4178,7 +4178,7 @@
da_args': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<guillemotright> A'" and
nrm_A_A': "nrm A \<subseteq> nrm A'"
by (rule da_weakenE) iprover
- note `PROP ?Hyp (In3 args) s1 s2`
+ note \<open>PROP ?Hyp (In3 args) s1 s2\<close>
with da_args' wt_args G normal_s2
have "nrm A' \<subseteq> dom (locals (store s2))"
by simp
@@ -4212,7 +4212,7 @@
ultimately show ?case by (intro conjI)
next
case (Methd s0 D sig v s1 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Methd.prems obtain m where
m: "methd (prg Env) D sig = Some m" and
da_body: "Env\<turnstile>(dom (locals (store ((Norm s0)::state))))
@@ -4222,7 +4222,7 @@
isCls: "is_class (prg Env) D" and
wt_body: "Env \<turnstile>In1l (Body (declclass m) (stmt (mbody (mthd m))))\<Colon>T"
by - (erule wt_elim_cases,simp)
- note `PROP ?Hyp (In1l (body G D sig)) (Norm s0) s1`
+ note \<open>PROP ?Hyp (In1l (body G D sig)) (Norm s0) s1\<close>
moreover
from wt_body have "Env\<turnstile>In1l (body G D sig)\<Colon>T"
using isCls m G by (simp add: body_def2)
@@ -4234,7 +4234,7 @@
using G by simp
next
case (Body s0 D s1 c s2 s3 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
from Body.prems
have nrm_A: "nrm A = dom (locals (store ((Norm s0)::state)))"
by (elim da_elim_cases) simp
@@ -4260,14 +4260,14 @@
thus ?case by simp
next
case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
have "?NormalAssigned s3 A"
proof
assume normal_s3: "normal s3"
show "nrm A \<subseteq> dom (locals (store s3))"
proof -
- note fvar = `(v, s2') = fvar statDeclC stat fn a s2` and
- s3 = `s3 = check_field_access G accC statDeclC fn stat a s2'`
+ note fvar = \<open>(v, s2') = fvar statDeclC stat fn a s2\<close> and
+ s3 = \<open>s3 = check_field_access G accC statDeclC fn stat a s2'\<close>
from FVar.prems
have da_e: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>\<langle>e\<rangle>\<guillemotright> A"
by (elim da_elim_cases)
@@ -4290,7 +4290,7 @@
show "normal s2"
by (cases s2) (simp add: fvar_def2)
qed
- note `PROP ?Hyp (In1l e) s1 s2`
+ note \<open>PROP ?Hyp (In1l e) s1 s2\<close>
with da_e' wt_e G normal_s2
have "nrm A' \<subseteq> dom (locals (store s2))"
by simp
@@ -4332,13 +4332,13 @@
ultimately show ?case by (intro conjI)
next
case (AVar s0 e1 a s1 e2 i s2 v s2' Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
have "?NormalAssigned s2' A"
proof
assume normal_s2': "normal s2'"
show "nrm A \<subseteq> dom (locals (store s2'))"
proof -
- note avar = `(v, s2') = avar G i a s2`
+ note avar = \<open>(v, s2') = avar G i a s2\<close>
from AVar.prems obtain E1 where
da_e1: "Env\<turnstile> (dom (locals (store ((Norm s0)::state))))\<guillemotright>\<langle>e1\<rangle>\<guillemotright> E1" and
da_e2: "Env\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A"
@@ -4352,14 +4352,14 @@
by (cases s2) (simp add: avar_def2)
hence "normal s1"
by - (rule eval_no_abrupt_lemma [rule_format], rule AVar, rule normal_s2)
- moreover note `PROP ?Hyp (In1l e1) (Norm s0) s1`
+ moreover note \<open>PROP ?Hyp (In1l e1) (Norm s0) s1\<close>
ultimately have "nrm E1 \<subseteq> dom (locals (store s1))"
using da_e1 wt_e1 G by simp
with da_e2 obtain A' where
da_e2': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e2\<rangle>\<guillemotright> A'" and
nrm_A_A': "nrm A \<subseteq> nrm A'"
by (rule da_weakenE) iprover
- note `PROP ?Hyp (In1l e2) s1 s2`
+ note \<open>PROP ?Hyp (In1l e2) s1 s2\<close>
with da_e2' wt_e2 G normal_s2
have "nrm A' \<subseteq> dom (locals (store s2))"
by simp
@@ -4404,7 +4404,7 @@
thus ?case by simp
next
case (Cons s0 e v s1 es vs s2 Env T A)
- note G = `prg Env = G`
+ note G = \<open>prg Env = G\<close>
have "?NormalAssigned s2 A"
proof
assume normal_s2: "normal s2"
@@ -4420,14 +4420,14 @@
by (elim wt_elim_cases)
have "normal s1"
by - (rule eval_no_abrupt_lemma [rule_format], rule Cons.hyps, rule normal_s2)
- moreover note `PROP ?Hyp (In1l e) (Norm s0) s1`
+ moreover note \<open>PROP ?Hyp (In1l e) (Norm s0) s1\<close>
ultimately have "nrm E \<subseteq> dom (locals (store s1))"
using da_e wt_e G by simp
with da_es obtain A' where
da_es': "Env\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>es\<rangle>\<guillemotright> A'" and
nrm_A_A': "nrm A \<subseteq> nrm A'"
by (rule da_weakenE) iprover
- note `PROP ?Hyp (In3 es) s1 s2`
+ note \<open>PROP ?Hyp (In3 es) s1 s2\<close>
with da_es' wt_es G normal_s2
have "nrm A' \<subseteq> dom (locals (store s2))"
by simp
--- a/src/HOL/Bali/Eval.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Eval.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,13 +1,13 @@
(* Title: HOL/Bali/Eval.thy
Author: David von Oheimb
*)
-subsection {* Operational evaluation (big-step) semantics of Java expressions and
+subsection \<open>Operational evaluation (big-step) semantics of Java expressions and
statements
-*}
+\<close>
theory Eval imports State DeclConcepts begin
-text {*
+text \<open>
improvements over Java Specification 1.0:
\begin{itemize}
@@ -57,10 +57,10 @@
\end{itemize}
\item the rules are defined carefully in order to be applicable even in not
type-correct situations (yielding undefined values),
- e.g. @{text "the_Addr (Val (Bool b)) = undefined"}.
+ e.g. \<open>the_Addr (Val (Bool b)) = undefined\<close>.
\begin{itemize}
\item[++] fewer rules
- \item[-] less readable because of auxiliary functions like @{text the_Addr}
+ \item[-] less readable because of auxiliary functions like \<open>the_Addr\<close>
\end{itemize}
Alternative: "defensive" evaluation throwing some InternalError exception
in case of (impossible, for correct programs) type mismatches
@@ -81,7 +81,7 @@
(whether there is enough memory to allocate the exception) in
evaluation rules
\end{itemize}
-\item unfortunately @{text new_Addr} is not directly executable because of
+\item unfortunately \<open>new_Addr\<close> is not directly executable because of
Hilbert operator.
\end{itemize}
simplifications:
@@ -93,7 +93,7 @@
modelled
\item exceptions in initializations not replaced by ExceptionInInitializerError
\end{itemize}
-*}
+\<close>
type_synonym vvar = "val \<times> (val \<Rightarrow> state \<Rightarrow> state)"
@@ -102,11 +102,11 @@
(type) "vvar" <= (type) "val \<times> (val \<Rightarrow> state \<Rightarrow> state)"
(type) "vals" <= (type) "(val, vvar, val list) sum3"
-text {* To avoid redundancy and to reduce the number of rules, there is only
+text \<open>To avoid redundancy and to reduce the number of rules, there is only
one evaluation rule for each syntactic term. This is also true for variables
- (e.g. see the rules below for @{text LVar}, @{text FVar} and @{text AVar}).
+ (e.g. see the rules below for \<open>LVar\<close>, \<open>FVar\<close> and \<open>AVar\<close>).
So evaluation of a variable must capture both possible further uses:
- read (rule @{text Acc}) or write (rule @{text Ass}) to the variable.
+ read (rule \<open>Acc\<close>) or write (rule \<open>Ass\<close>) to the variable.
Therefor a variable evaluates to a special value @{term vvar}, which is
a pair, consisting of the current value (for later read access) and an update
function (for later write access). Because
@@ -122,7 +122,7 @@
such a generic update function, but only to store the address and the kind
of variable (array (+ element type), local variable or field) for later
assignment.
-*}
+\<close>
abbreviation
dummy_res :: "vals" ("\<diamondsuit>")
@@ -355,7 +355,7 @@
-lemma init_lvars_def2: --{* better suited for simplification *}
+lemma init_lvars_def2: \<comment>\<open>better suited for simplification\<close>
"init_lvars G C sig mode a' pvs (x,s) =
set_lvars
(\<lambda> k.
@@ -378,7 +378,7 @@
(let m = the (methd G C sig)
in Body (declclass m) (stmt (mbody (mthd m))))"
-lemma body_def2: --{* better suited for simplification *}
+lemma body_def2: \<comment>\<open>better suited for simplification\<close>
"body G C sig = Body (declclass (the (methd G C sig)))
(stmt (mbody (mthd (the (methd G C sig)))))"
apply (unfold body_def Let_def)
@@ -412,7 +412,7 @@
,upd_gobj oref n v s))
in ((the (cs n),f),abupd (raise_if (\<not>i in_bounds k) IndOutBound \<circ> np a') s))"
-lemma fvar_def2: --{* better suited for simplification *}
+lemma fvar_def2: \<comment>\<open>better suited for simplification\<close>
"fvar C stat fn a' s =
((the
(values
@@ -427,7 +427,7 @@
apply (simp (no_asm) add: Let_def split_beta)
done
-lemma avar_def2: --{* better suited for simplification *}
+lemma avar_def2: \<comment>\<open>better suited for simplification\<close>
"avar G i' a' s =
((the ((snd(snd(the_Arr (globs (store s) (Heap (the_Addr a'))))))
(Inr (the_Intg i')))
@@ -471,7 +471,7 @@
inductive
halloc :: "[prog,state,obj_tag,loc,state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>halloc _\<succ>_\<rightarrow> _"[61,61,61,61,61]60) for G::prog
-where --{* allocating objects on the heap, cf. 12.5 *}
+where \<comment>\<open>allocating objects on the heap, cf. 12.5\<close>
Abrupt:
"G\<turnstile>(Some x,s) \<midarrow>halloc oi\<succ>undefined\<rightarrow> (Some x,s)"
@@ -483,8 +483,8 @@
G\<turnstile>Norm s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,init_obj G oi' (Heap a) s)"
inductive sxalloc :: "[prog,state,state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>sxalloc\<rightarrow> _"[61,61,61]60) for G::prog
-where --{* allocating exception objects for
- standard exceptions (other than OutOfMemory) *}
+where \<comment>\<open>allocating exception objects for
+ standard exceptions (other than OutOfMemory)\<close>
Norm: "G\<turnstile> Norm s \<midarrow>sxalloc\<rightarrow> Norm s"
@@ -513,42 +513,42 @@
| "G\<turnstile>s \<midarrow>e=\<succ>vf\<rightarrow> s' \<equiv> G\<turnstile>s \<midarrow>In2 e\<succ>\<rightarrow> (In2 vf, s')"
| "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<rightarrow> s' \<equiv> G\<turnstile>s \<midarrow>In3 e\<succ>\<rightarrow> (In3 v, s')"
---{* propagation of abrupt completion *}
+\<comment>\<open>propagation of abrupt completion\<close>
- --{* cf. 14.1, 15.5 *}
+ \<comment>\<open>cf. 14.1, 15.5\<close>
| Abrupt:
"G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (undefined3 t, (Some xc, s))"
---{* execution of statements *}
+\<comment>\<open>execution of statements\<close>
- --{* cf. 14.5 *}
+ \<comment>\<open>cf. 14.5\<close>
| Skip: "G\<turnstile>Norm s \<midarrow>Skip\<rightarrow> Norm s"
- --{* cf. 14.7 *}
+ \<comment>\<open>cf. 14.7\<close>
| Expr: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>Expr e\<rightarrow> s1"
| Lab: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<rightarrow> s1\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<rightarrow> abupd (absorb l) s1"
- --{* cf. 14.2 *}
+ \<comment>\<open>cf. 14.2\<close>
| Comp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<rightarrow> s1;
G\<turnstile> s1 \<midarrow>c2 \<rightarrow> s2\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>c1;; c2\<rightarrow> s2"
- --{* cf. 14.8.2 *}
+ \<comment>\<open>cf. 14.8.2\<close>
| If: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
G\<turnstile> s1\<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<rightarrow> s2"
- --{* cf. 14.10, 14.10.1 *}
+ \<comment>\<open>cf. 14.10, 14.10.1\<close>
- --{* A continue jump from the while body @{term c} is handled by
+ \<comment>\<open>A continue jump from the while body @{term c} is handled by
this rule. If a continue jump with the proper label was invoked inside
@{term c} this label (Cont l) is deleted out of the abrupt component of
the state before the iterative evaluation of the while statement.
- A break jump is handled by the Lab Statement @{text "Lab l (while\<dots>)"}.
- *}
+ A break jump is handled by the Lab Statement \<open>Lab l (while\<dots>)\<close>.
+\<close>
| Loop: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
if the_Bool b
then (G\<turnstile>s1 \<midarrow>c\<rightarrow> s2 \<and>
@@ -558,16 +558,16 @@
| Jmp: "G\<turnstile>Norm s \<midarrow>Jmp j\<rightarrow> (Some (Jump j), s)"
- --{* cf. 14.16 *}
+ \<comment>\<open>cf. 14.16\<close>
| Throw: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>Throw e\<rightarrow> abupd (throw a') s1"
- --{* cf. 14.18.1 *}
+ \<comment>\<open>cf. 14.18.1\<close>
| Try: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2;
if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(C vn) c2\<rightarrow> s3"
- --{* cf. 14.18.2 *}
+ \<comment>\<open>cf. 14.18.2\<close>
| Fin: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1,s1);
G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2;
s3=(if (\<exists> err. x1=Some (Error err))
@@ -575,7 +575,7 @@
else abupd (abrupt_if (x1\<noteq>None) x1) s2) \<rbrakk>
\<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<rightarrow> s3"
- --{* cf. 12.4.2, 8.5 *}
+ \<comment>\<open>cf. 12.4.2, 8.5\<close>
| Init: "\<lbrakk>the (class G C) = c;
if inited C (globs s0) then s3 = Norm s0
else (G\<turnstile>Norm (init_class_obj G C s0)
@@ -583,20 +583,20 @@
G\<turnstile>set_lvars empty s1 \<midarrow>init c\<rightarrow> s2 \<and> s3 = restore_lvars s1 s2)\<rbrakk>
\<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s3"
- --{* This class initialisation rule is a little bit inaccurate. Look at the
+ \<comment>\<open>This class initialisation rule is a little bit inaccurate. Look at the
exact sequence:
(1) The current class object (the static fields) are initialised
- (@{text init_class_obj}),
+ (\<open>init_class_obj\<close>),
(2) the superclasses are initialised,
(3) the static initialiser of the current class is invoked.
More precisely we should expect another ordering, namely 2 1 3.
But we can't just naively toggle 1 and 2. By calling
- @{text init_class_obj}
+ \<open>init_class_obj\<close>
before initialising the superclasses, we also implicitly record that
we have started to initialise the current class (by setting an
value for the class object). This becomes
crucial for the completeness proof of the axiomatic semantics
- @{text "AxCompl.thy"}. Static initialisation requires an induction on
+ \<open>AxCompl.thy\<close>. Static initialisation requires an induction on
the number of classes not yet initialised (or to be more precise,
classes were the initialisation has not yet begun).
So we could first assign a dummy value to the class before
@@ -604,30 +604,30 @@
But as long as we don't take memory overflow into account
when allocating class objects, we can leave things as they are for
convenience.
- *}
---{* evaluation of expressions *}
+\<close>
+\<comment>\<open>evaluation of expressions\<close>
- --{* cf. 15.8.1, 12.4.1 *}
+ \<comment>\<open>cf. 15.8.1, 12.4.1\<close>
| NewC: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1;
G\<turnstile> s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<rightarrow> s2"
- --{* cf. 15.9.1, 12.4.1 *}
+ \<comment>\<open>cf. 15.9.1, 12.4.1\<close>
| NewA: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<rightarrow> s2;
G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<rightarrow> s3"
- --{* cf. 15.15 *}
+ \<comment>\<open>cf. 15.15\<close>
| Cast: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
s2 = abupd (raise_if (\<not>G,store s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<rightarrow> s2"
- --{* cf. 15.19.2 *}
+ \<comment>\<open>cf. 15.19.2\<close>
| Inst: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<rightarrow> s1"
- --{* cf. 15.7.1 *}
+ \<comment>\<open>cf. 15.7.1\<close>
| Lit: "G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<rightarrow> Norm s"
| UnOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk>
@@ -639,25 +639,25 @@
\<rbrakk>
\<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)\<rightarrow> s2"
- --{* cf. 15.10.2 *}
+ \<comment>\<open>cf. 15.10.2\<close>
| Super: "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<rightarrow> Norm s"
- --{* cf. 15.2 *}
+ \<comment>\<open>cf. 15.2\<close>
| Acc: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<rightarrow> s1\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<rightarrow> s1"
- --{* cf. 15.25.1 *}
+ \<comment>\<open>cf. 15.25.1\<close>
| Ass: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<rightarrow> s1;
G\<turnstile> s1 \<midarrow>e-\<succ>v \<rightarrow> s2\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>va:=e-\<succ>v\<rightarrow> assign f v s2"
- --{* cf. 15.24 *}
+ \<comment>\<open>cf. 15.24\<close>
| Cond: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1;
G\<turnstile> s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<rightarrow> s2"
--- {* The interplay of @{term Call}, @{term Methd} and @{term Body}:
+\<comment> \<open>The interplay of @{term Call}, @{term Methd} and @{term Body}:
Method invocation is split up into these three rules:
\begin{itemize}
\item [@{term Call}] Calculates the target address and evaluates the
@@ -674,8 +674,8 @@
initialisation. Without class initialisation we
could just evaluate the body statement.
\end{itemize}
- *}
- --{* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5 *}
+\<close>
+ \<comment>\<open>cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5\<close>
| Call:
"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1; G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2;
D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>;
@@ -684,10 +684,10 @@
G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<rightarrow> s4\<rbrakk>
\<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<rightarrow> (restore_lvars s2 s4)"
---{* The accessibility check is after @{term init_lvars}, to keep it simple.
+\<comment>\<open>The accessibility check is after @{term init_lvars}, to keep it simple.
@{term init_lvars} already tests for the absence of a null-pointer
reference in case of an instance method invocation.
-*}
+\<close>
| Methd: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<rightarrow> s1"
@@ -699,40 +699,40 @@
else s2)\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)
\<rightarrow>abupd (absorb Ret) s3"
- --{* cf. 14.15, 12.4.1 *}
- --{* We filter out a break/continue in @{term s2}, so that we can proof
+ \<comment>\<open>cf. 14.15, 12.4.1\<close>
+ \<comment>\<open>We filter out a break/continue in @{term s2}, so that we can proof
definite assignment
correct, without the need of conformance of the state. By this the
- different parts of the typesafety proof can be disentangled a little. *}
+ different parts of the typesafety proof can be disentangled a little.\<close>
---{* evaluation of variables *}
+\<comment>\<open>evaluation of variables\<close>
- --{* cf. 15.13.1, 15.7.2 *}
+ \<comment>\<open>cf. 15.13.1, 15.7.2\<close>
| LVar: "G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<rightarrow> Norm s"
- --{* cf. 15.10.1, 12.4.1 *}
+ \<comment>\<open>cf. 15.10.1, 12.4.1\<close>
| FVar: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2;
(v,s2') = fvar statDeclC stat fn a s2;
s3 = check_field_access G accC statDeclC fn stat a s2' \<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>v\<rightarrow> s3"
- --{* The accessibility check is after @{term fvar}, to keep it simple.
+ \<comment>\<open>The accessibility check is after @{term fvar}, to keep it simple.
@{term fvar} already tests for the absence of a null-pointer reference
in case of an instance field
- *}
+\<close>
- --{* cf. 15.12.1, 15.25.1 *}
+ \<comment>\<open>cf. 15.12.1, 15.25.1\<close>
| AVar: "\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<rightarrow> s2;
(v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<rightarrow> s2'"
---{* evaluation of expression lists *}
+\<comment>\<open>evaluation of expression lists\<close>
- --{* cf. 15.11.4.2 *}
+ \<comment>\<open>cf. 15.11.4.2\<close>
| Nil:
"G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<rightarrow> Norm s0"
- --{* cf. 15.6.4 *}
+ \<comment>\<open>cf. 15.6.4\<close>
| Cons: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e -\<succ> v \<rightarrow> s1;
G\<turnstile> s1 \<midarrow>es\<doteq>\<succ>vs\<rightarrow> s2\<rbrakk> \<Longrightarrow>
G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<rightarrow> s2"
@@ -744,13 +744,13 @@
29(AVar),24(Call)]
*)
-ML {*
+ML \<open>
ML_Thms.bind_thm ("eval_induct", rearrange_prems
[0,1,2,8,4,30,31,27,15,16,
17,18,19,20,21,3,5,25,26,23,6,
7,11,9,13,14,12,22,10,28,
29,24] @{thm eval.induct})
-*}
+\<close>
declare split_if [split del] split_if_asm [split del]
@@ -780,7 +780,7 @@
declare not_None_eq [simp del] (* IntDef.Zero_def [simp del] *)
declare split_paired_All [simp del] split_paired_Ex [simp del]
-setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\<close>
inductive_cases eval_cases: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v, s')"
@@ -818,7 +818,7 @@
"G\<turnstile>Norm s \<midarrow>In1r (Init C) \<succ>\<rightarrow> (x, s')"
declare not_None_eq [simp] (* IntDef.Zero_def [simp] *)
declare split_paired_All [simp] split_paired_Ex [simp]
-declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *}
+declaration \<open>K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac)))\<close>
declare split_if [split] split_if_asm [split]
option.split [split] option.split_asm [split]
@@ -837,12 +837,12 @@
apply auto
done
-text {* The following simplification procedures set up the proper injections of
+text \<open>The following simplification procedures set up the proper injections of
terms and their corresponding values in the evaluation relation:
E.g. an expression
(injection @{term In1l} into terms) always evaluates to ordinary values
(injection @{term In1} into generalised values @{term vals}).
-*}
+\<close>
lemma eval_expr_eq: "G\<turnstile>s \<midarrow>In1l t\<succ>\<rightarrow> (w, s') = (\<exists>v. w=In1 v \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<rightarrow> s')"
by (auto, frule eval_Inj_elim, auto)
@@ -856,40 +856,40 @@
lemma eval_stmt_eq: "G\<turnstile>s \<midarrow>In1r t\<succ>\<rightarrow> (w, s') = (w=\<diamondsuit> \<and> G\<turnstile>s \<midarrow>t \<rightarrow> s')"
by (auto, frule eval_Inj_elim, auto, frule eval_Inj_elim, auto)
-simproc_setup eval_expr ("G\<turnstile>s \<midarrow>In1l t\<succ>\<rightarrow> (w, s')") = {*
+simproc_setup eval_expr ("G\<turnstile>s \<midarrow>In1l t\<succ>\<rightarrow> (w, s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
- | _ => SOME (mk_meta_eq @{thm eval_expr_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm eval_expr_eq}))\<close>
-simproc_setup eval_var ("G\<turnstile>s \<midarrow>In2 t\<succ>\<rightarrow> (w, s')") = {*
+simproc_setup eval_var ("G\<turnstile>s \<midarrow>In2 t\<succ>\<rightarrow> (w, s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
- | _ => SOME (mk_meta_eq @{thm eval_var_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm eval_var_eq}))\<close>
-simproc_setup eval_exprs ("G\<turnstile>s \<midarrow>In3 t\<succ>\<rightarrow> (w, s')") = {*
+simproc_setup eval_exprs ("G\<turnstile>s \<midarrow>In3 t\<succ>\<rightarrow> (w, s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
- | _ => SOME (mk_meta_eq @{thm eval_exprs_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm eval_exprs_eq}))\<close>
-simproc_setup eval_stmt ("G\<turnstile>s \<midarrow>In1r t\<succ>\<rightarrow> (w, s')") = {*
+simproc_setup eval_stmt ("G\<turnstile>s \<midarrow>In1r t\<succ>\<rightarrow> (w, s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
- | _ => SOME (mk_meta_eq @{thm eval_stmt_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm eval_stmt_eq}))\<close>
-ML {*
+ML \<open>
ML_Thms.bind_thms ("AbruptIs", sum3_instantiate @{context} @{thm eval.Abrupt})
-*}
+\<close>
declare halloc.Abrupt [intro!] eval.Abrupt [intro!] AbruptIs [intro!]
-text{* @{text Callee},@{text InsInitE}, @{text InsInitV}, @{text FinA} are only
+text\<open>\<open>Callee\<close>,\<open>InsInitE\<close>, \<open>InsInitV\<close>, \<open>FinA\<close> are only
used in smallstep semantics, not in the bigstep semantics. So their is no
valid evaluation of these terms
-*}
+\<close>
lemma eval_Callee: "G\<turnstile>Norm s\<midarrow>Callee l e-\<succ>v\<rightarrow> s' = False"
@@ -952,12 +952,12 @@
apply (frule eval_no_abrupt_lemma, auto)+
done
-simproc_setup eval_no_abrupt ("G\<turnstile>(x,s) \<midarrow>e\<succ>\<rightarrow> (w,Norm s')") = {*
+simproc_setup eval_no_abrupt ("G\<turnstile>(x,s) \<midarrow>e\<succ>\<rightarrow> (w,Norm s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ (Const (@{const_name Pair}, _) $ (Const (@{const_name None}, _)) $ _) $ _ $ _ $ _) => NONE
| _ => SOME (mk_meta_eq @{thm eval_no_abrupt}))
-*}
+\<close>
lemma eval_abrupt_lemma:
@@ -972,12 +972,12 @@
apply (frule eval_abrupt_lemma, auto)+
done
-simproc_setup eval_abrupt ("G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<rightarrow> (w,s')") = {*
+simproc_setup eval_abrupt ("G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<rightarrow> (w,s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ _ $ (Const (@{const_name Pair}, _) $ (Const (@{const_name Some}, _) $ _)$ _)) => NONE
| _ => SOME (mk_meta_eq @{thm eval_abrupt}))
-*}
+\<close>
lemma LitI: "G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else undefined)\<rightarrow> s"
apply (case_tac "s", case_tac "a = None")
@@ -1162,8 +1162,8 @@
lemma unique_eval [rule_format (no_asm)]:
"G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w, s') \<Longrightarrow> (\<forall>w' s''. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w', s'') \<longrightarrow> w' = w \<and> s'' = s')"
apply (erule eval_induct)
-apply (tactic {* ALLGOALS (EVERY'
- [strip_tac @{context}, rotate_tac ~1, eresolve_tac @{context} @{thms eval_elim_cases}]) *})
+apply (tactic \<open>ALLGOALS (EVERY'
+ [strip_tac @{context}, rotate_tac ~1, eresolve_tac @{context} @{thms eval_elim_cases}])\<close>)
(* 31 subgoals *)
prefer 28 (* Try *)
apply (simp (no_asm_use) only: split add: split_if_asm)
--- a/src/HOL/Bali/Evaln.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Evaln.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,14 +1,14 @@
(* Title: HOL/Bali/Evaln.thy
Author: David von Oheimb and Norbert Schirmer
*)
-subsection {* Operational evaluation (big-step) semantics of Java expressions and
+subsection \<open>Operational evaluation (big-step) semantics of Java expressions and
statements
-*}
+\<close>
theory Evaln imports TypeSafe begin
-text {*
+text \<open>
Variant of @{term eval} relation with counter for bounded recursive depth.
In principal @{term evaln} could replace @{term eval}.
@@ -25,7 +25,7 @@
@{term check_field_access} and @{term check_method_access} like @{term eval}.
If it would omit them @{term evaln} and @{term eval} would only be equivalent
for welltyped, and definitely assigned terms.
-*}
+\<close>
inductive
evaln :: "[prog, state, term, nat, vals, state] \<Rightarrow> bool"
@@ -46,12 +46,12 @@
| "G\<turnstile>s \<midarrow>e=\<succ>vf \<midarrow>n\<rightarrow> s' \<equiv> G\<turnstile>s \<midarrow>In2 e\<succ>\<midarrow>n\<rightarrow> (In2 vf, s')"
| "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<midarrow>n\<rightarrow> s' \<equiv> G\<turnstile>s \<midarrow>In3 e\<succ>\<midarrow>n\<rightarrow> (In3 v , s')"
---{* propagation of abrupt completion *}
+\<comment>\<open>propagation of abrupt completion\<close>
| Abrupt: "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (undefined3 t,(Some xc,s))"
---{* evaluation of variables *}
+\<comment>\<open>evaluation of variables\<close>
| LVar: "G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<midarrow>n\<rightarrow> Norm s"
@@ -67,7 +67,7 @@
---{* evaluation of expressions *}
+\<comment>\<open>evaluation of expressions\<close>
| NewC: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1;
G\<turnstile> s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
@@ -129,7 +129,7 @@
G\<turnstile>Norm s0 \<midarrow>Body D c
-\<succ>the (locals (store s2) Result)\<midarrow>n\<rightarrow>abupd (absorb Ret) s3"
---{* evaluation of expression lists *}
+\<comment>\<open>evaluation of expression lists\<close>
| Nil:
"G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<midarrow>n\<rightarrow> Norm s0"
@@ -139,7 +139,7 @@
G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<midarrow>n\<rightarrow> s2"
---{* execution of statements *}
+\<comment>\<open>execution of statements\<close>
| Skip: "G\<turnstile>Norm s \<midarrow>Skip\<midarrow>n\<rightarrow> Norm s"
@@ -197,7 +197,7 @@
option.split [split del] option.split_asm [split del]
not_None_eq [simp del]
split_paired_All [simp del] split_paired_Ex [simp del]
-setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\<close>
inductive_cases evaln_cases: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s')"
@@ -238,7 +238,7 @@
option.split [split] option.split_asm [split]
not_None_eq [simp]
split_paired_All [simp] split_paired_Ex [simp]
-declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *}
+declaration \<open>K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac)))\<close>
lemma evaln_Inj_elim: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (w,s') \<Longrightarrow> case t of In1 ec \<Rightarrow>
(case ec of Inl e \<Rightarrow> (\<exists>v. w = In1 v) | Inr c \<Rightarrow> w = \<diamondsuit>)
@@ -249,12 +249,12 @@
apply auto
done
-text {* The following simplification procedures set up the proper injections of
+text \<open>The following simplification procedures set up the proper injections of
terms and their corresponding values in the evaluation relation:
E.g. an expression
(injection @{term In1l} into terms) always evaluates to ordinary values
(injection @{term In1} into generalised values @{term vals}).
-*}
+\<close>
lemma evaln_expr_eq: "G\<turnstile>s \<midarrow>In1l t\<succ>\<midarrow>n\<rightarrow> (w, s') = (\<exists>v. w=In1 v \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<midarrow>n\<rightarrow> s')"
by (auto, frule evaln_Inj_elim, auto)
@@ -268,31 +268,31 @@
lemma evaln_stmt_eq: "G\<turnstile>s \<midarrow>In1r t\<succ>\<midarrow>n\<rightarrow> (w, s') = (w=\<diamondsuit> \<and> G\<turnstile>s \<midarrow>t \<midarrow>n\<rightarrow> s')"
by (auto, frule evaln_Inj_elim, auto, frule evaln_Inj_elim, auto)
-simproc_setup evaln_expr ("G\<turnstile>s \<midarrow>In1l t\<succ>\<midarrow>n\<rightarrow> (w, s')") = {*
+simproc_setup evaln_expr ("G\<turnstile>s \<midarrow>In1l t\<succ>\<midarrow>n\<rightarrow> (w, s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
- | _ => SOME (mk_meta_eq @{thm evaln_expr_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm evaln_expr_eq}))\<close>
-simproc_setup evaln_var ("G\<turnstile>s \<midarrow>In2 t\<succ>\<midarrow>n\<rightarrow> (w, s')") = {*
+simproc_setup evaln_var ("G\<turnstile>s \<midarrow>In2 t\<succ>\<midarrow>n\<rightarrow> (w, s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
- | _ => SOME (mk_meta_eq @{thm evaln_var_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm evaln_var_eq}))\<close>
-simproc_setup evaln_exprs ("G\<turnstile>s \<midarrow>In3 t\<succ>\<midarrow>n\<rightarrow> (w, s')") = {*
+simproc_setup evaln_exprs ("G\<turnstile>s \<midarrow>In3 t\<succ>\<midarrow>n\<rightarrow> (w, s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
- | _ => SOME (mk_meta_eq @{thm evaln_exprs_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm evaln_exprs_eq}))\<close>
-simproc_setup evaln_stmt ("G\<turnstile>s \<midarrow>In1r t\<succ>\<midarrow>n\<rightarrow> (w, s')") = {*
+simproc_setup evaln_stmt ("G\<turnstile>s \<midarrow>In1r t\<succ>\<midarrow>n\<rightarrow> (w, s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE
- | _ => SOME (mk_meta_eq @{thm evaln_stmt_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm evaln_stmt_eq}))\<close>
-ML {* ML_Thms.bind_thms ("evaln_AbruptIs", sum3_instantiate @{context} @{thm evaln.Abrupt}) *}
+ML \<open>ML_Thms.bind_thms ("evaln_AbruptIs", sum3_instantiate @{context} @{thm evaln.Abrupt})\<close>
declare evaln_AbruptIs [intro!]
lemma evaln_Callee: "G\<turnstile>Norm s\<midarrow>In1l (Callee l e)\<succ>\<midarrow>n\<rightarrow> (v,s') = False"
@@ -355,13 +355,13 @@
apply (frule evaln_abrupt_lemma, auto)+
done
-simproc_setup evaln_abrupt ("G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (w,s')") = {*
+simproc_setup evaln_abrupt ("G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<midarrow>n\<rightarrow> (w,s')") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ _ $ _ $ (Const (@{const_name Pair}, _) $ (Const (@{const_name Some},_) $ _)$ _))
=> NONE
| _ => SOME (mk_meta_eq @{thm evaln_abrupt}))
-*}
+\<close>
lemma evaln_LitI: "G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else undefined)\<midarrow>n\<rightarrow> s"
apply (case_tac "s", case_tac "a = None")
@@ -401,7 +401,7 @@
-subsubsection {* evaln implies eval *}
+subsubsection \<open>evaln implies eval\<close>
lemma evaln_eval:
assumes evaln: "G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1)"
@@ -409,7 +409,7 @@
using evaln
proof (induct)
case (Loop s0 e b n s1 c s2 l s3)
- note `G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1`
+ note \<open>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1\<close>
moreover
have "if the_Bool b
then (G\<turnstile>s1 \<midarrow>c\<rightarrow> s2) \<and>
@@ -419,16 +419,16 @@
ultimately show ?case by (rule eval.Loop)
next
case (Try s0 c1 n s1 s2 C vn c2 s3)
- note `G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1`
+ note \<open>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1\<close>
moreover
- note `G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2`
+ note \<open>G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2\<close>
moreover
have "if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2"
using Try.hyps by simp
ultimately show ?case by (rule eval.Try)
next
case (Init C c s0 s3 n s1 s2)
- note `the (class G C) = c`
+ note \<open>the (class G C) = c\<close>
moreover
have "if inited C (globs s0)
then s3 = Norm s0
@@ -448,10 +448,10 @@
lemma evaln_nonstrict [rule_format (no_asm), elim]:
"G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (w, s') \<Longrightarrow> \<forall>m. n\<le>m \<longrightarrow> G\<turnstile>s \<midarrow>t\<succ>\<midarrow>m\<rightarrow> (w, s')"
apply (erule evaln.induct)
-apply (tactic {* ALLGOALS (EVERY' [strip_tac @{context},
+apply (tactic \<open>ALLGOALS (EVERY' [strip_tac @{context},
TRY o eresolve_tac @{context} @{thms Suc_le_D_lemma},
REPEAT o smp_tac @{context} 1,
- resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW TRY o assume_tac @{context}]) *})
+ resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW TRY o assume_tac @{context}])\<close>)
(* 3 subgoals *)
apply (auto split del: split_if)
done
@@ -511,7 +511,7 @@
declare [[simproc del: wt_expr wt_var wt_exprs wt_stmt]]
-subsubsection {* eval implies evaln *}
+subsubsection \<open>eval implies evaln\<close>
lemma eval_evaln:
assumes eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)"
shows "\<exists>n. G\<turnstile>s0 \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v,s1)"
@@ -597,7 +597,7 @@
"G\<turnstile>Norm s0 \<midarrow>c1\<midarrow>n1\<rightarrow> s1"
by (iprover)
moreover
- note sxalloc = `G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2`
+ note sxalloc = \<open>G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2\<close>
moreover
from Try.hyps obtain n2 where
"if G,s2\<turnstile>catch catchC then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n2\<rightarrow> s3 else s3 = s2"
@@ -613,9 +613,9 @@
"G\<turnstile>Norm s1 \<midarrow>c2\<midarrow>n2\<rightarrow> s2"
by iprover
moreover
- note s3 = `s3 = (if \<exists>err. x1 = Some (Error err)
+ note s3 = \<open>s3 = (if \<exists>err. x1 = Some (Error err)
then (x1, s1)
- else abupd (abrupt_if (x1 \<noteq> None) x1) s2)`
+ else abupd (abrupt_if (x1 \<noteq> None) x1) s2)\<close>
ultimately
have
"G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<midarrow>max n1 n2\<rightarrow> s3"
@@ -623,7 +623,7 @@
then show ?case ..
next
case (Init C c s0 s3 s1 s2)
- note cls = `the (class G C) = c`
+ note cls = \<open>the (class G C) = c\<close>
moreover from Init.hyps obtain n where
"if inited C (globs s0) then s3 = Norm s0
else (G\<turnstile>Norm (init_class_obj G C s0)
@@ -650,7 +650,7 @@
"G\<turnstile>s1 \<midarrow>e-\<succ>i\<midarrow>n2\<rightarrow> s2"
by (iprover)
moreover
- note `G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3`
+ note \<open>G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3\<close>
ultimately
have "G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<midarrow>max n1 n2\<rightarrow> s3"
by (blast intro: evaln.NewA dest: evaln_max2)
@@ -661,7 +661,7 @@
"G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
by (iprover)
moreover
- note `s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits castT) ClassCast) s1`
+ note \<open>s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits castT) ClassCast) s1\<close>
ultimately
have "G\<turnstile>Norm s0 \<midarrow>Cast castT e-\<succ>v\<midarrow>n\<rightarrow> s2"
by (rule evaln.Cast)
@@ -672,7 +672,7 @@
"G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
by (iprover)
moreover
- note `b = (v \<noteq> Null \<and> G,snd s1\<turnstile>v fits RefT T)`
+ note \<open>b = (v \<noteq> Null \<and> G,snd s1\<turnstile>v fits RefT T)\<close>
ultimately
have "G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<midarrow>n\<rightarrow> s1"
by (rule evaln.Inst)
@@ -742,12 +742,12 @@
"G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<midarrow>n2\<rightarrow> s2"
by iprover
moreover
- note `invDeclC = invocation_declclass G mode (store s2) a' statT
- \<lparr>name=mn,parTs=pTs'\<rparr>`
+ note \<open>invDeclC = invocation_declclass G mode (store s2) a' statT
+ \<lparr>name=mn,parTs=pTs'\<rparr>\<close>
moreover
- note `s3 = init_lvars G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> mode a' vs s2`
+ note \<open>s3 = init_lvars G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> mode a' vs s2\<close>
moreover
- note `s3'=check_method_access G accC' statT mode \<lparr>name=mn,parTs=pTs'\<rparr> a' s3`
+ note \<open>s3'=check_method_access G accC' statT mode \<lparr>name=mn,parTs=pTs'\<rparr> a' s3\<close>
moreover
from Call.hyps
obtain m where
@@ -773,10 +773,10 @@
evaln_c: "G\<turnstile>s1 \<midarrow>c\<midarrow>n2\<rightarrow> s2"
by (iprover)
moreover
- note `s3 = (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
+ note \<open>s3 = (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
fst s2 = Some (Jump (Cont l))
then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2
- else s2)`
+ else s2)\<close>
ultimately
have
"G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)\<midarrow>max n1 n2
@@ -796,8 +796,8 @@
"G\<turnstile>s1 \<midarrow>e-\<succ>a\<midarrow>n2\<rightarrow> s2"
by iprover
moreover
- note `s3 = check_field_access G accC statDeclC fn stat a s2'`
- and `(v, s2') = fvar statDeclC stat fn a s2`
+ note \<open>s3 = check_field_access G accC statDeclC fn stat a s2'\<close>
+ and \<open>(v, s2') = fvar statDeclC stat fn a s2\<close>
ultimately
have "G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>v\<midarrow>max n1 n2\<rightarrow> s3"
by (iprover intro: evaln.FVar dest: evaln_max2)
@@ -809,7 +809,7 @@
"G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n2\<rightarrow> s2"
by iprover
moreover
- note `(v, s2') = avar G i a s2`
+ note \<open>(v, s2') = avar G i a s2\<close>
ultimately
have "G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<midarrow>max n1 n2\<rightarrow> s2'"
by (blast intro!: evaln.AVar dest: evaln_max2)
--- a/src/HOL/Bali/Example.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Example.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,13 +1,13 @@
(* Title: HOL/Bali/Example.thy
Author: David von Oheimb
*)
-subsection {* Example Bali program *}
+subsection \<open>Example Bali program\<close>
theory Example
imports Eval WellForm
begin
-text {*
+text \<open>
The following example Bali program includes:
\begin{itemize}
\item class and interface declarations with inheritance, hiding of fields,
@@ -52,7 +52,7 @@
}
}
\end{verbatim}
-*}
+\<close>
declare widen.null [intro]
lemma wf_fdecl_def2: "\<And>fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))"
@@ -894,7 +894,7 @@
declare member_is_static_simp [simp]
declare wt.Skip [rule del] wt.Init [rule del]
-ML {* ML_Thms.bind_thms ("wt_intros", map (rewrite_rule @{context} @{thms id_def}) @{thms wt.intros}) *}
+ML \<open>ML_Thms.bind_thms ("wt_intros", map (rewrite_rule @{context} @{thms id_def}) @{thms wt.intros})\<close>
lemmas wtIs = wt_Call wt_Super wt_FVar wt_StatRef wt_intros
lemmas daIs = assigned.select_convs da_Skip da_NewC da_Lit da_Super da.intros
@@ -1187,9 +1187,9 @@
declare BaseCl_def [simp] ExtCl_def [simp] Ext_foo_def [simp]
Base_foo_defs [simp]
-ML {* ML_Thms.bind_thms ("eval_intros", map
+ML \<open>ML_Thms.bind_thms ("eval_intros", map
(simplify (@{context} delsimps @{thms Skip_eq} addsimps @{thms lvar_def}) o
- rewrite_rule @{context} [@{thm assign_def}, @{thm Let_def}]) @{thms eval.intros}) *}
+ rewrite_rule @{context} [@{thm assign_def}, @{thm Let_def}]) @{thms eval.intros})\<close>
lemmas eval_Is = eval_Init eval_StatRef AbruptIs eval_intros
axiomatization
--- a/src/HOL/Bali/Name.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Name.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,22 +1,22 @@
(* Title: HOL/Bali/Name.thy
Author: David von Oheimb
*)
-subsection {* Java names *}
+subsection \<open>Java names\<close>
theory Name imports Basis begin
(* cf. 6.5 *)
-typedecl tnam --{* ordinary type name, i.e. class or interface name *}
-typedecl pname --{* package name *}
-typedecl mname --{* method name *}
-typedecl vname --{* variable or field name *}
-typedecl label --{* label as destination of break or continue *}
+typedecl tnam \<comment>\<open>ordinary type name, i.e. class or interface name\<close>
+typedecl pname \<comment>\<open>package name\<close>
+typedecl mname \<comment>\<open>method name\<close>
+typedecl vname \<comment>\<open>variable or field name\<close>
+typedecl label \<comment>\<open>label as destination of break or continue\<close>
-datatype ename --{* expression name *}
+datatype ename \<comment>\<open>expression name\<close>
= VNam vname
- | Res --{* special name to model the return value of methods *}
+ | Res \<comment>\<open>special name to model the return value of methods\<close>
-datatype lname --{* names for local variables and the This pointer *}
+datatype lname \<comment>\<open>names for local variables and the This pointer\<close>
= EName ename
| This
abbreviation VName :: "vname \<Rightarrow> lname"
@@ -25,7 +25,7 @@
abbreviation Result :: lname
where "Result == EName Res"
-datatype xname --{* names of standard exceptions *}
+datatype xname \<comment>\<open>names of standard exceptions\<close>
= Throwable
| NullPointer | OutOfMemory | ClassCast
| NegArrSize | IndOutBound | ArrStore
@@ -39,12 +39,12 @@
done
-datatype tname --{* type names for standard classes and other type names *}
+datatype tname \<comment>\<open>type names for standard classes and other type names\<close>
= Object'
| SXcpt' xname
| TName tnam
-record qtname = --{* qualified tname cf. 6.5.3, 6.5.4*}
+record qtname = \<comment>\<open>qualified tname cf. 6.5.3, 6.5.4\<close>
pid :: pname
tid :: tname
@@ -82,7 +82,7 @@
(type) "'a qtname_scheme" <= (type) "\<lparr>pid::pname,tid::tname,\<dots>::'a\<rparr>"
-axiomatization java_lang::pname --{* package java.lang *}
+axiomatization java_lang::pname \<comment>\<open>package java.lang\<close>
definition
Object :: qtname
--- a/src/HOL/Bali/State.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/State.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,35 +1,35 @@
(* Title: HOL/Bali/State.thy
Author: David von Oheimb
*)
-subsection {* State for evaluation of Java expressions and statements *}
+subsection \<open>State for evaluation of Java expressions and statements\<close>
theory State
imports DeclConcepts
begin
-text {*
+text \<open>
design issues:
\begin{itemize}
\item all kinds of objects (class instances, arrays, and class objects)
are handeled via a general object abstraction
\item the heap and the map for class objects are combined into a single table
- @{text "(recall (loc, obj) table \<times> (qtname, obj) table ~= (loc + qtname, obj) table)"}
+ \<open>(recall (loc, obj) table \<times> (qtname, obj) table ~= (loc + qtname, obj) table)\<close>
\end{itemize}
-*}
+\<close>
subsubsection "objects"
-datatype obj_tag = --{* tag for generic object *}
- CInst qtname --{* class instance *}
- | Arr ty int --{* array with component type and length *}
- --{* | CStat qtname the tag is irrelevant for a class object,
+datatype obj_tag = \<comment>\<open>tag for generic object\<close>
+ CInst qtname \<comment>\<open>class instance\<close>
+ | Arr ty int \<comment>\<open>array with component type and length\<close>
+ \<comment>\<open>| CStat qtname the tag is irrelevant for a class object,
i.e. the static fields of a class,
since its type is given already by the reference to
- it (see below) *}
+ it (see below)\<close>
-type_synonym vn = "fspec + int" --{* variable name *}
+type_synonym vn = "fspec + int" \<comment>\<open>variable name\<close>
record obj =
- tag :: "obj_tag" --{* generalized object *}
+ tag :: "obj_tag" \<comment>\<open>generalized object\<close>
"values" :: "(vn, val) table"
translations
@@ -130,7 +130,7 @@
subsubsection "object references"
-type_synonym oref = "loc + qtname" --{* generalized object reference *}
+type_synonym oref = "loc + qtname" \<comment>\<open>generalized object reference\<close>
syntax
Heap :: "loc \<Rightarrow> oref"
Stat :: "qtname \<Rightarrow> oref"
@@ -213,7 +213,7 @@
subsubsection "stores"
-type_synonym globs --{* global variables: heap and static variables *}
+type_synonym globs \<comment>\<open>global variables: heap and static variables\<close>
= "(oref , obj) table"
type_synonym heap
= "(loc , obj) table"
@@ -580,7 +580,7 @@
subsubsection "full program state"
type_synonym
- state = "abopt \<times> st" --{* state including abruption information *}
+ state = "abopt \<times> st" \<comment>\<open>state including abruption information\<close>
translations
(type) "abopt" <= (type) "abrupt option"
@@ -727,7 +727,7 @@
apply (simp (no_asm))
done
-subsubsection {* @{text error_free} *}
+subsubsection \<open>\<open>error_free\<close>\<close>
definition
error_free :: "state \<Rightarrow> bool"
--- a/src/HOL/Bali/Table.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Table.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,11 +1,11 @@
(* Title: HOL/Bali/Table.thy
Author: David von Oheimb
*)
-subsection {* Abstract tables and their implementation as lists *}
+subsection \<open>Abstract tables and their implementation as lists\<close>
theory Table imports Basis begin
-text {*
+text \<open>
design issues:
\begin{itemize}
\item definition of table: infinite map vs. list vs. finite set
@@ -27,18 +27,18 @@
\item[-] sometimes awkward case distinctions, alleviated by operator 'the'
\end{itemize}
\end{itemize}
-*}
+\<close>
-type_synonym ('a, 'b) table --{* table with key type 'a and contents type 'b *}
+type_synonym ('a, 'b) table \<comment>\<open>table with key type 'a and contents type 'b\<close>
= "'a \<rightharpoonup> 'b"
-type_synonym ('a, 'b) tables --{* non-unique table with key 'a and contents 'b *}
+type_synonym ('a, 'b) tables \<comment>\<open>non-unique table with key 'a and contents 'b\<close>
= "'a \<Rightarrow> 'b set"
subsubsection "map of / table of"
abbreviation
- table_of :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) table" --{* concrete table *}
+ table_of :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) table" \<comment>\<open>concrete table\<close>
where "table_of \<equiv> map_of"
translations
@@ -49,12 +49,12 @@
by (simp add: map_add_def)
-subsubsection {* Conditional Override *}
+subsubsection \<open>Conditional Override\<close>
definition cond_override :: "('b \<Rightarrow>'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b) table" where
---{* when merging tables old and new, only override an entry of table old when
- the condition cond holds *}
+\<comment>\<open>when merging tables old and new, only override an entry of table old when
+ the condition cond holds\<close>
"cond_override cond old new =
(\<lambda>k.
(case new k of
@@ -95,7 +95,7 @@
by (rule finite_UnI)
-subsubsection {* Filter on Tables *}
+subsubsection \<open>Filter on Tables\<close>
definition filter_tab :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) table \<Rightarrow> ('a, 'b) table"
where
@@ -179,7 +179,7 @@
by (auto simp add: fun_eq_iff cond_override_def filter_tab_def )
-subsubsection {* Misc *}
+subsubsection \<open>Misc\<close>
lemma Ball_set_table: "(\<forall> (x,y)\<in> set l. P x y) \<Longrightarrow> \<forall> x. \<forall> y\<in> map_of l x: P x y"
apply (erule rev_mp)
@@ -276,13 +276,13 @@
where "(t hidings s entails R) = (\<forall>k. \<forall>x\<in>t k. \<forall>y\<in>s k. R x y)"
definition
- --{* variant for unique table: *}
+ \<comment>\<open>variant for unique table:\<close>
hiding_entails :: "('a, 'b) table \<Rightarrow> ('a, 'c) table \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"
("_ hiding _ entails _" 20)
where "(t hiding s entails R) = (\<forall>k. \<forall>x\<in>t k: \<forall>y\<in>s k: R x y)"
definition
- --{* variant for a unique table and conditional overriding: *}
+ \<comment>\<open>variant for a unique table and conditional overriding:\<close>
cond_hiding_entails :: "('a, 'b) table \<Rightarrow> ('a, 'c) table
\<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"
("_ hiding _ under _ entails _" 20)
--- a/src/HOL/Bali/Term.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Term.thy Sat Jan 02 18:48:45 2016 +0100
@@ -2,11 +2,11 @@
Author: David von Oheimb
*)
-subsection {* Java expressions and statements *}
+subsection \<open>Java expressions and statements\<close>
theory Term imports Value Table begin
-text {*
+text \<open>
design issues:
\begin{itemize}
\item invocation frames for local variables could be reduced to special static
@@ -42,185 +42,185 @@
\item no synchronized statements
\item no secondary forms of if, while (e.g. no for) (may be easily simulated)
\item no switch (may be simulated with if)
-\item the @{text try_catch_finally} statement is divided into the
- @{text try_catch} statement
+\item the \<open>try_catch_finally\<close> statement is divided into the
+ \<open>try_catch\<close> statement
and a finally statement, which may be considered as try..finally with
empty catch
-\item the @{text try_catch} statement has exactly one catch clause;
+\item the \<open>try_catch\<close> statement has exactly one catch clause;
multiple ones can be
simulated with instanceof
-\item the compiler is supposed to add the annotations {@{text _}} during
+\item the compiler is supposed to add the annotations {\<open>_\<close>} during
type-checking. This
transformation is left out as its result is checked by the type rules anyway
\end{itemize}
-*}
+\<close>
-type_synonym locals = "(lname, val) table" --{* local variables *}
+type_synonym locals = "(lname, val) table" \<comment>\<open>local variables\<close>
datatype jump
- = Break label --{* break *}
- | Cont label --{* continue *}
- | Ret --{* return from method *}
+ = Break label \<comment>\<open>break\<close>
+ | Cont label \<comment>\<open>continue\<close>
+ | Ret \<comment>\<open>return from method\<close>
-datatype xcpt --{* exception *}
- = Loc loc --{* location of allocated execption object *}
- | Std xname --{* intermediate standard exception, see Eval.thy *}
+datatype xcpt \<comment>\<open>exception\<close>
+ = Loc loc \<comment>\<open>location of allocated execption object\<close>
+ | Std xname \<comment>\<open>intermediate standard exception, see Eval.thy\<close>
datatype error
- = AccessViolation --{* Access to a member that isn't permitted *}
- | CrossMethodJump --{* Method exits with a break or continue *}
+ = AccessViolation \<comment>\<open>Access to a member that isn't permitted\<close>
+ | CrossMethodJump \<comment>\<open>Method exits with a break or continue\<close>
-datatype abrupt --{* abrupt completion *}
- = Xcpt xcpt --{* exception *}
- | Jump jump --{* break, continue, return *}
- | Error error -- {* runtime errors, we wan't to detect and proof absent
- in welltyped programms *}
+datatype abrupt \<comment>\<open>abrupt completion\<close>
+ = Xcpt xcpt \<comment>\<open>exception\<close>
+ | Jump jump \<comment>\<open>break, continue, return\<close>
+ | Error error \<comment> \<open>runtime errors, we wan't to detect and proof absent
+ in welltyped programms\<close>
type_synonym
abopt = "abrupt option"
-text {* Local variable store and exception.
+text \<open>Local variable store and exception.
Anticipation of State.thy used by smallstep semantics. For a method call,
we save the local variables of the caller in the term Callee to restore them
after method return. Also an exception must be restored after the finally
-statement *}
+statement\<close>
translations
(type) "locals" <= (type) "(lname, val) table"
-datatype inv_mode --{* invocation mode for method calls *}
- = Static --{* static *}
- | SuperM --{* super *}
- | IntVir --{* interface or virtual *}
+datatype inv_mode \<comment>\<open>invocation mode for method calls\<close>
+ = Static \<comment>\<open>static\<close>
+ | SuperM \<comment>\<open>super\<close>
+ | IntVir \<comment>\<open>interface or virtual\<close>
-record sig = --{* signature of a method, cf. 8.4.2 *}
- name ::"mname" --{* acutally belongs to Decl.thy *}
+record sig = \<comment>\<open>signature of a method, cf. 8.4.2\<close>
+ name ::"mname" \<comment>\<open>acutally belongs to Decl.thy\<close>
parTs::"ty list"
translations
(type) "sig" <= (type) "\<lparr>name::mname,parTs::ty list\<rparr>"
(type) "sig" <= (type) "\<lparr>name::mname,parTs::ty list,\<dots>::'a\<rparr>"
---{* function codes for unary operations *}
-datatype unop = UPlus -- {*{\tt +} unary plus*}
- | UMinus -- {*{\tt -} unary minus*}
- | UBitNot -- {*{\tt ~} bitwise NOT*}
- | UNot -- {*{\tt !} logical complement*}
+\<comment>\<open>function codes for unary operations\<close>
+datatype unop = UPlus \<comment> \<open>{\tt +} unary plus\<close>
+ | UMinus \<comment> \<open>{\tt -} unary minus\<close>
+ | UBitNot \<comment> \<open>{\tt ~} bitwise NOT\<close>
+ | UNot \<comment> \<open>{\tt !} logical complement\<close>
---{* function codes for binary operations *}
-datatype binop = Mul -- {*{\tt * } multiplication*}
- | Div -- {*{\tt /} division*}
- | Mod -- {*{\tt \%} remainder*}
- | Plus -- {*{\tt +} addition*}
- | Minus -- {*{\tt -} subtraction*}
- | LShift -- {*{\tt <<} left shift*}
- | RShift -- {*{\tt >>} signed right shift*}
- | RShiftU -- {*{\tt >>>} unsigned right shift*}
- | Less -- {*{\tt <} less than*}
- | Le -- {*{\tt <=} less than or equal*}
- | Greater -- {*{\tt >} greater than*}
- | Ge -- {*{\tt >=} greater than or equal*}
- | Eq -- {*{\tt ==} equal*}
- | Neq -- {*{\tt !=} not equal*}
- | BitAnd -- {*{\tt \&} bitwise AND*}
- | And -- {*{\tt \&} boolean AND*}
- | BitXor -- {*{\texttt \^} bitwise Xor*}
- | Xor -- {*{\texttt \^} boolean Xor*}
- | BitOr -- {*{\tt |} bitwise Or*}
- | Or -- {*{\tt |} boolean Or*}
- | CondAnd -- {*{\tt \&\&} conditional And*}
- | CondOr -- {*{\tt ||} conditional Or *}
-text{* The boolean operators {\tt \&} and {\tt |} strictly evaluate both
+\<comment>\<open>function codes for binary operations\<close>
+datatype binop = Mul \<comment> \<open>{\tt * } multiplication\<close>
+ | Div \<comment> \<open>{\tt /} division\<close>
+ | Mod \<comment> \<open>{\tt \%} remainder\<close>
+ | Plus \<comment> \<open>{\tt +} addition\<close>
+ | Minus \<comment> \<open>{\tt -} subtraction\<close>
+ | LShift \<comment> \<open>{\tt <<} left shift\<close>
+ | RShift \<comment> \<open>{\tt >>} signed right shift\<close>
+ | RShiftU \<comment> \<open>{\tt >>>} unsigned right shift\<close>
+ | Less \<comment> \<open>{\tt <} less than\<close>
+ | Le \<comment> \<open>{\tt <=} less than or equal\<close>
+ | Greater \<comment> \<open>{\tt >} greater than\<close>
+ | Ge \<comment> \<open>{\tt >=} greater than or equal\<close>
+ | Eq \<comment> \<open>{\tt ==} equal\<close>
+ | Neq \<comment> \<open>{\tt !=} not equal\<close>
+ | BitAnd \<comment> \<open>{\tt \&} bitwise AND\<close>
+ | And \<comment> \<open>{\tt \&} boolean AND\<close>
+ | BitXor \<comment> \<open>{\texttt \^} bitwise Xor\<close>
+ | Xor \<comment> \<open>{\texttt \^} boolean Xor\<close>
+ | BitOr \<comment> \<open>{\tt |} bitwise Or\<close>
+ | Or \<comment> \<open>{\tt |} boolean Or\<close>
+ | CondAnd \<comment> \<open>{\tt \&\&} conditional And\<close>
+ | CondOr \<comment> \<open>{\tt ||} conditional Or\<close>
+text\<open>The boolean operators {\tt \&} and {\tt |} strictly evaluate both
of their arguments. The conditional operators {\tt \&\&} and {\tt ||} only
evaluate the second argument if the value of the whole expression isn't
allready determined by the first argument.
e.g.: {\tt false \&\& e} e is not evaluated;
{\tt true || e} e is not evaluated;
-*}
+\<close>
datatype var
- = LVar lname --{* local variable (incl. parameters) *}
+ = LVar lname \<comment>\<open>local variable (incl. parameters)\<close>
| FVar qtname qtname bool expr vname ("{_,_,_}_.._"[10,10,10,85,99]90)
- --{* class field *}
- --{* @{term "{accC,statDeclC,stat}e..fn"} *}
- --{* @{text accC}: accessing class (static class were *}
- --{* the code is declared. Annotation only needed for *}
- --{* evaluation to check accessibility) *}
- --{* @{text statDeclC}: static declaration class of field*}
- --{* @{text stat}: static or instance field?*}
- --{* @{text e}: reference to object*}
- --{* @{text fn}: field name*}
+ \<comment>\<open>class field\<close>
+ \<comment>\<open>@{term "{accC,statDeclC,stat}e..fn"}\<close>
+ \<comment>\<open>\<open>accC\<close>: accessing class (static class were\<close>
+ \<comment>\<open>the code is declared. Annotation only needed for\<close>
+ \<comment>\<open>evaluation to check accessibility)\<close>
+ \<comment>\<open>\<open>statDeclC\<close>: static declaration class of field\<close>
+ \<comment>\<open>\<open>stat\<close>: static or instance field?\<close>
+ \<comment>\<open>\<open>e\<close>: reference to object\<close>
+ \<comment>\<open>\<open>fn\<close>: field name\<close>
| AVar expr expr ("_.[_]"[90,10 ]90)
- --{* array component *}
- --{* @{term "e1.[e2]"}: e1 array reference; e2 index *}
+ \<comment>\<open>array component\<close>
+ \<comment>\<open>@{term "e1.[e2]"}: e1 array reference; e2 index\<close>
| InsInitV stmt var
- --{* insertion of initialization before evaluation *}
- --{* of var (technical term for smallstep semantics.)*}
+ \<comment>\<open>insertion of initialization before evaluation\<close>
+ \<comment>\<open>of var (technical term for smallstep semantics.)\<close>
and expr
- = NewC qtname --{* class instance creation *}
+ = NewC qtname \<comment>\<open>class instance creation\<close>
| NewA ty expr ("New _[_]"[99,10 ]85)
- --{* array creation *}
- | Cast ty expr --{* type cast *}
+ \<comment>\<open>array creation\<close>
+ | Cast ty expr \<comment>\<open>type cast\<close>
| Inst expr ref_ty ("_ InstOf _"[85,99] 85)
- --{* instanceof *}
- | Lit val --{* literal value, references not allowed *}
- | UnOp unop expr --{* unary operation *}
- | BinOp binop expr expr --{* binary operation *}
+ \<comment>\<open>instanceof\<close>
+ | Lit val \<comment>\<open>literal value, references not allowed\<close>
+ | UnOp unop expr \<comment>\<open>unary operation\<close>
+ | BinOp binop expr expr \<comment>\<open>binary operation\<close>
- | Super --{* special Super keyword *}
- | Acc var --{* variable access *}
+ | Super \<comment>\<open>special Super keyword\<close>
+ | Acc var \<comment>\<open>variable access\<close>
| Ass var expr ("_:=_" [90,85 ]85)
- --{* variable assign *}
- | Cond expr expr expr ("_ ? _ : _" [85,85,80]80) --{* conditional *}
+ \<comment>\<open>variable assign\<close>
+ | Cond expr expr expr ("_ ? _ : _" [85,85,80]80) \<comment>\<open>conditional\<close>
| Call qtname ref_ty inv_mode expr mname "(ty list)" "(expr list)"
("{_,_,_}_\<cdot>_'( {_}_')"[10,10,10,85,99,10,10]85)
- --{* method call *}
- --{* @{term "{accC,statT,mode}e\<cdot>mn({pTs}args)"} " *}
- --{* @{text accC}: accessing class (static class were *}
- --{* the call code is declared. Annotation only needed for*}
- --{* evaluation to check accessibility) *}
- --{* @{text statT}: static declaration class/interface of *}
- --{* method *}
- --{* @{text mode}: invocation mode *}
- --{* @{text e}: reference to object*}
- --{* @{text mn}: field name*}
- --{* @{text pTs}: types of parameters *}
- --{* @{text args}: the actual parameters/arguments *}
- | Methd qtname sig --{* (folded) method (see below) *}
- | Body qtname stmt --{* (unfolded) method body *}
+ \<comment>\<open>method call\<close>
+ \<comment>\<open>@{term "{accC,statT,mode}e\<cdot>mn({pTs}args)"} "\<close>
+ \<comment>\<open>\<open>accC\<close>: accessing class (static class were\<close>
+ \<comment>\<open>the call code is declared. Annotation only needed for\<close>
+ \<comment>\<open>evaluation to check accessibility)\<close>
+ \<comment>\<open>\<open>statT\<close>: static declaration class/interface of\<close>
+ \<comment>\<open>method\<close>
+ \<comment>\<open>\<open>mode\<close>: invocation mode\<close>
+ \<comment>\<open>\<open>e\<close>: reference to object\<close>
+ \<comment>\<open>\<open>mn\<close>: field name\<close>
+ \<comment>\<open>\<open>pTs\<close>: types of parameters\<close>
+ \<comment>\<open>\<open>args\<close>: the actual parameters/arguments\<close>
+ | Methd qtname sig \<comment>\<open>(folded) method (see below)\<close>
+ | Body qtname stmt \<comment>\<open>(unfolded) method body\<close>
| InsInitE stmt expr
- --{* insertion of initialization before *}
- --{* evaluation of expr (technical term for smallstep sem.) *}
- | Callee locals expr --{* save callers locals in callee-Frame *}
- --{* (technical term for smallstep semantics) *}
+ \<comment>\<open>insertion of initialization before\<close>
+ \<comment>\<open>evaluation of expr (technical term for smallstep sem.)\<close>
+ | Callee locals expr \<comment>\<open>save callers locals in callee-Frame\<close>
+ \<comment>\<open>(technical term for smallstep semantics)\<close>
and stmt
- = Skip --{* empty statement *}
- | Expr expr --{* expression statement *}
+ = Skip \<comment>\<open>empty statement\<close>
+ | Expr expr \<comment>\<open>expression statement\<close>
| Lab jump stmt ("_\<bullet> _" [ 99,66]66)
- --{* labeled statement; handles break *}
+ \<comment>\<open>labeled statement; handles break\<close>
| Comp stmt stmt ("_;; _" [ 66,65]65)
| If' expr stmt stmt ("If'(_') _ Else _" [ 80,79,79]70)
| Loop label expr stmt ("_\<bullet> While'(_') _" [ 99,80,79]70)
- | Jmp jump --{* break, continue, return *}
+ | Jmp jump \<comment>\<open>break, continue, return\<close>
| Throw expr
| TryC stmt qtname vname stmt ("Try _ Catch'(_ _') _" [79,99,80,79]70)
- --{* @{term "Try c1 Catch(C vn) c2"} *}
- --{* @{text c1}: block were exception may be thrown *}
- --{* @{text C}: execption class to catch *}
- --{* @{text vn}: local name for exception used in @{text c2}*}
- --{* @{text c2}: block to execute when exception is cateched*}
+ \<comment>\<open>@{term "Try c1 Catch(C vn) c2"}\<close>
+ \<comment>\<open>\<open>c1\<close>: block were exception may be thrown\<close>
+ \<comment>\<open>\<open>C\<close>: execption class to catch\<close>
+ \<comment>\<open>\<open>vn\<close>: local name for exception used in \<open>c2\<close>\<close>
+ \<comment>\<open>\<open>c2\<close>: block to execute when exception is cateched\<close>
| Fin stmt stmt ("_ Finally _" [ 79,79]70)
- | FinA abopt stmt --{* Save abruption of first statement *}
- --{* technical term for smallstep sem.) *}
- | Init qtname --{* class initialization *}
+ | FinA abopt stmt \<comment>\<open>Save abruption of first statement\<close>
+ \<comment>\<open>technical term for smallstep sem.)\<close>
+ | Init qtname \<comment>\<open>class initialization\<close>
datatype_compat var expr stmt
-text {*
+text \<open>
The expressions Methd and Body are artificial program constructs, in the
sense that they are not used to define a concrete Bali program. In the
operational semantic's they are "generated on the fly"
@@ -235,7 +235,7 @@
frame stack.
The InsInitV/E terms are only used by the smallstep semantics to model the
intermediate steps of class-initialisation.
-*}
+\<close>
type_synonym "term" = "(expr+stmt,var,expr list) sum3"
translations
@@ -254,7 +254,7 @@
abbreviation
Return :: "expr \<Rightarrow> stmt"
- where "Return e == Expr (Ass (LVar (EName Res)) e);; Jmp Ret" --{* \tt Res := e;; Jmp Ret *}
+ where "Return e == Expr (Ass (LVar (EName Res)) e);; Jmp Ret" \<comment>\<open>\tt Res := e;; Jmp Ret\<close>
abbreviation
StatRef :: "ref_ty \<Rightarrow> expr"
@@ -264,14 +264,14 @@
is_stmt :: "term \<Rightarrow> bool"
where "is_stmt t = (\<exists>c. t=In1r c)"
-ML {* ML_Thms.bind_thms ("is_stmt_rews", sum3_instantiate @{context} @{thm is_stmt_def}) *}
+ML \<open>ML_Thms.bind_thms ("is_stmt_rews", sum3_instantiate @{context} @{thm is_stmt_def})\<close>
declare is_stmt_rews [simp]
-text {*
+text \<open>
Here is some syntactic stuff to handle the injections of statements,
expressions, variables and expression lists into general terms.
-*}
+\<close>
abbreviation (input)
expr_inj_term :: "expr \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>e" 1000)
@@ -289,22 +289,22 @@
lst_inj_term :: "expr list \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>l" 1000)
where "\<langle>es\<rangle>\<^sub>l == In3 es"
-text {* It seems to be more elegant to have an overloaded injection like the
+text \<open>It seems to be more elegant to have an overloaded injection like the
following.
-*}
+\<close>
class inj_term =
fixes inj_term:: "'a \<Rightarrow> term" ("\<langle>_\<rangle>" 1000)
-text {* How this overloaded injections work can be seen in the theory
-@{text DefiniteAssignment}. Other big inductive relations on
-terms defined in theories @{text WellType}, @{text Eval}, @{text Evaln} and
-@{text AxSem} don't follow this convention right now, but introduce subtle
+text \<open>How this overloaded injections work can be seen in the theory
+\<open>DefiniteAssignment\<close>. Other big inductive relations on
+terms defined in theories \<open>WellType\<close>, \<open>Eval\<close>, \<open>Evaln\<close> and
+\<open>AxSem\<close> don't follow this convention right now, but introduce subtle
syntactic sugar in the relations themselves to make a distinction on
expressions, statements and so on. So unfortunately you will encounter a
mixture of dealing with these injections. The abbreviations above are used
as bridge between the different conventions.
-*}
+\<close>
instantiation stmt :: inj_term
begin
@@ -427,15 +427,15 @@
apply auto
done
-subsubsection {* Evaluation of unary operations *}
+subsubsection \<open>Evaluation of unary operations\<close>
primrec eval_unop :: "unop \<Rightarrow> val \<Rightarrow> val"
where
"eval_unop UPlus v = Intg (the_Intg v)"
| "eval_unop UMinus v = Intg (- (the_Intg v))"
-| "eval_unop UBitNot v = Intg 42" -- "FIXME: Not yet implemented"
+| "eval_unop UBitNot v = Intg 42" \<comment> "FIXME: Not yet implemented"
| "eval_unop UNot v = Bool (\<not> the_Bool v)"
-subsubsection {* Evaluation of binary operations *}
+subsubsection \<open>Evaluation of binary operations\<close>
primrec eval_binop :: "binop \<Rightarrow> val \<Rightarrow> val \<Rightarrow> val"
where
"eval_binop Mul v1 v2 = Intg ((the_Intg v1) * (the_Intg v2))"
@@ -444,10 +444,10 @@
| "eval_binop Plus v1 v2 = Intg ((the_Intg v1) + (the_Intg v2))"
| "eval_binop Minus v1 v2 = Intg ((the_Intg v1) - (the_Intg v2))"
--- "Be aware of the explicit coercion of the shift distance to nat"
+\<comment> "Be aware of the explicit coercion of the shift distance to nat"
| "eval_binop LShift v1 v2 = Intg ((the_Intg v1) * (2^(nat (the_Intg v2))))"
| "eval_binop RShift v1 v2 = Intg ((the_Intg v1) div (2^(nat (the_Intg v2))))"
-| "eval_binop RShiftU v1 v2 = Intg 42" --"FIXME: Not yet implemented"
+| "eval_binop RShiftU v1 v2 = Intg 42" \<comment>"FIXME: Not yet implemented"
| "eval_binop Less v1 v2 = Bool ((the_Intg v1) < (the_Intg v2))"
| "eval_binop Le v1 v2 = Bool ((the_Intg v1) \<le> (the_Intg v2))"
@@ -456,11 +456,11 @@
| "eval_binop Eq v1 v2 = Bool (v1=v2)"
| "eval_binop Neq v1 v2 = Bool (v1\<noteq>v2)"
-| "eval_binop BitAnd v1 v2 = Intg 42" -- "FIXME: Not yet implemented"
+| "eval_binop BitAnd v1 v2 = Intg 42" \<comment> "FIXME: Not yet implemented"
| "eval_binop And v1 v2 = Bool ((the_Bool v1) \<and> (the_Bool v2))"
-| "eval_binop BitXor v1 v2 = Intg 42" -- "FIXME: Not yet implemented"
+| "eval_binop BitXor v1 v2 = Intg 42" \<comment> "FIXME: Not yet implemented"
| "eval_binop Xor v1 v2 = Bool ((the_Bool v1) \<noteq> (the_Bool v2))"
-| "eval_binop BitOr v1 v2 = Intg 42" -- "FIXME: Not yet implemented"
+| "eval_binop BitOr v1 v2 = Intg 42" \<comment> "FIXME: Not yet implemented"
| "eval_binop Or v1 v2 = Bool ((the_Bool v1) \<or> (the_Bool v2))"
| "eval_binop CondAnd v1 v2 = Bool ((the_Bool v1) \<and> (the_Bool v2))"
| "eval_binop CondOr v1 v2 = Bool ((the_Bool v1) \<or> (the_Bool v2))"
@@ -469,8 +469,8 @@
need_second_arg :: "binop \<Rightarrow> val \<Rightarrow> bool" where
"need_second_arg binop v1 = (\<not> ((binop=CondAnd \<and> \<not> the_Bool v1) \<or>
(binop=CondOr \<and> the_Bool v1)))"
-text {* @{term CondAnd} and @{term CondOr} only evalulate the second argument
- if the value isn't already determined by the first argument*}
+text \<open>@{term CondAnd} and @{term CondOr} only evalulate the second argument
+ if the value isn't already determined by the first argument\<close>
lemma need_second_arg_CondAnd [simp]: "need_second_arg CondAnd (Bool b) = b"
by (simp add: need_second_arg_def)
--- a/src/HOL/Bali/Trans.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Trans.thy Sat Jan 02 18:48:45 2016 +0100
@@ -236,14 +236,14 @@
| InsInitFVar:
"G\<turnstile>(\<langle>InsInitV Skip ({accC,statDeclC,stat}Lit a..fn)\<rangle>,Norm s)
\<mapsto>1 (\<langle>{accC,statDeclC,stat}Lit a..fn\<rangle>,Norm s)"
--- {* Notice, that we do not have literal values for @{text vars}.
-The rules for accessing variables (@{text Acc}) and assigning to variables
-(@{text Ass}), test this with the predicate @{text groundVar}. After
-initialisation is done and the @{text FVar} is evaluated, we can't just
-throw away the @{text InsInitFVar} term and return a literal value, as in the
-cases of @{text New} or @{text NewC}. Instead we just return the evaluated
-@{text FVar} and test for initialisation in the rule @{text FVar}.
-*}
+\<comment> \<open>Notice, that we do not have literal values for \<open>vars\<close>.
+The rules for accessing variables (\<open>Acc\<close>) and assigning to variables
+(\<open>Ass\<close>), test this with the predicate \<open>groundVar\<close>. After
+initialisation is done and the \<open>FVar\<close> is evaluated, we can't just
+throw away the \<open>InsInitFVar\<close> term and return a literal value, as in the
+cases of \<open>New\<close> or \<open>NewC\<close>. Instead we just return the evaluated
+\<open>FVar\<close> and test for initialisation in the rule \<open>FVar\<close>.
+\<close>
| AVarE1: "\<lbrakk>G\<turnstile>(\<langle>e1\<rangle>,Norm s) \<mapsto>1 (\<langle>e1'\<rangle>,s')\<rbrakk>
@@ -258,7 +258,7 @@
(* evaluation of expression lists *)
- -- {* @{text Nil} is fully evaluated *}
+ \<comment> \<open>\<open>Nil\<close> is fully evaluated\<close>
| ConsHd: "\<lbrakk>G\<turnstile>(\<langle>e::expr\<rangle>,Norm s) \<mapsto>1 (\<langle>e'::expr\<rangle>,s')\<rbrakk>
\<Longrightarrow>
@@ -339,8 +339,8 @@
\<mapsto>1 (\<langle>(if C = Object then Skip else (Init (super c)));;
Expr (Callee (locals s) (InsInitE (init c) SKIP))\<rangle>
,Norm (init_class_obj G C s))"
--- {* @{text InsInitE} is just used as trick to embed the statement
-@{text "init c"} into an expression*}
+\<comment> \<open>\<open>InsInitE\<close> is just used as trick to embed the statement
+\<open>init c\<close> into an expression\<close>
| InsInitESKIP:
"G\<turnstile>(\<langle>InsInitE Skip SKIP\<rangle>,Norm s) \<mapsto>1 (\<langle>SKIP\<rangle>,Norm s)"
--- a/src/HOL/Bali/Type.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Type.thy Sat Jan 02 18:48:45 2016 +0100
@@ -2,33 +2,33 @@
Author: David von Oheimb
*)
-subsection {* Java types *}
+subsection \<open>Java types\<close>
theory Type imports Name begin
-text {*
+text \<open>
simplifications:
\begin{itemize}
\item only the most important primitive types
\item the null type is regarded as reference type
\end{itemize}
-*}
+\<close>
-datatype prim_ty --{* primitive type, cf. 4.2 *}
- = Void --{* result type of void methods *}
+datatype prim_ty \<comment>\<open>primitive type, cf. 4.2\<close>
+ = Void \<comment>\<open>result type of void methods\<close>
| Boolean
| Integer
-datatype ref_ty --{* reference type, cf. 4.3 *}
- = NullT --{* null type, cf. 4.1 *}
- | IfaceT qtname --{* interface type *}
- | ClassT qtname --{* class type *}
- | ArrayT ty --{* array type *}
+datatype ref_ty \<comment>\<open>reference type, cf. 4.3\<close>
+ = NullT \<comment>\<open>null type, cf. 4.1\<close>
+ | IfaceT qtname \<comment>\<open>interface type\<close>
+ | ClassT qtname \<comment>\<open>class type\<close>
+ | ArrayT ty \<comment>\<open>array type\<close>
-and ty --{* any type, cf. 4.1 *}
- = PrimT prim_ty --{* primitive type *}
- | RefT ref_ty --{* reference type *}
+and ty \<comment>\<open>any type, cf. 4.1\<close>
+ = PrimT prim_ty \<comment>\<open>primitive type\<close>
+ | RefT ref_ty \<comment>\<open>reference type\<close>
abbreviation "NT == RefT NullT"
abbreviation "Iface I == RefT (IfaceT I)"
--- a/src/HOL/Bali/TypeRel.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/TypeRel.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,11 +1,11 @@
(* Title: HOL/Bali/TypeRel.thy
Author: David von Oheimb
*)
-subsection {* The relations between Java types *}
+subsection \<open>The relations between Java types\<close>
theory TypeRel imports Decl begin
-text {*
+text \<open>
simplifications:
\begin{itemize}
\item subinterface, subclass and widening relation includes identity
@@ -19,11 +19,11 @@
\end{itemize}
design issues:
\begin{itemize}
-\item the type relations do not require @{text is_type} for their arguments
-\item the subint1 and subcls1 relations imply @{text is_iface}/@{text is_class}
+\item the type relations do not require \<open>is_type\<close> for their arguments
+\item the subint1 and subcls1 relations imply \<open>is_iface\<close>/\<open>is_class\<close>
for their first arguments, which is required for their finiteness
\end{itemize}
-*}
+\<close>
(*subint1, in Decl.thy*) (* direct subinterface *)
(*subint , by translation*) (* subinterface (+ identity) *)
@@ -32,8 +32,8 @@
(*subclseq, by translation*) (* subclass + identity *)
definition
- implmt1 :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{* direct implementation *}
- --{* direct implementation, cf. 8.1.3 *}
+ implmt1 :: "prog \<Rightarrow> (qtname \<times> qtname) set" \<comment>\<open>direct implementation\<close>
+ \<comment>\<open>direct implementation, cf. 8.1.3\<close>
where "implmt1 G = {(C,I). C\<noteq>Object \<and> (\<exists>c\<in>class G C: I\<in>set (superIfs c))}"
@@ -43,7 +43,7 @@
abbreviation
subint_syntax :: "prog => [qtname, qtname] => bool" ("_\<turnstile>_\<preceq>I _" [71,71,71] 70)
- where "G\<turnstile>I \<preceq>I J == (I,J) \<in>(subint1 G)^*" --{* cf. 9.1.3 *}
+ where "G\<turnstile>I \<preceq>I J == (I,J) \<in>(subint1 G)^*" \<comment>\<open>cf. 9.1.3\<close>
abbreviation
implmt1_syntax :: "prog => [qtname, qtname] => bool" ("_\<turnstile>_\<leadsto>1_" [71,71,71] 70)
@@ -334,7 +334,7 @@
apply auto
done
-inductive --{* implementation, cf. 8.1.4 *}
+inductive \<comment>\<open>implementation, cf. 8.1.4\<close>
implmt :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" ("_\<turnstile>_\<leadsto>_" [71,71,71] 70)
for G :: prog
where
@@ -369,13 +369,13 @@
subsubsection "widening relation"
inductive
- --{*widening, viz. method invocation conversion, cf. 5.3
- i.e. kind of syntactic subtyping *}
+ \<comment>\<open>widening, viz. method invocation conversion, cf. 5.3
+ i.e. kind of syntactic subtyping\<close>
widen :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_\<preceq>_" [71,71,71] 70)
for G :: prog
where
- refl: "G\<turnstile>T\<preceq>T" --{*identity conversion, cf. 5.1.1 *}
-| subint: "G\<turnstile>I\<preceq>I J \<Longrightarrow> G\<turnstile>Iface I\<preceq> Iface J" --{*wid.ref.conv.,cf. 5.1.4 *}
+ refl: "G\<turnstile>T\<preceq>T" \<comment>\<open>identity conversion, cf. 5.1.1\<close>
+| subint: "G\<turnstile>I\<preceq>I J \<Longrightarrow> G\<turnstile>Iface I\<preceq> Iface J" \<comment>\<open>wid.ref.conv.,cf. 5.1.4\<close>
| int_obj: "G\<turnstile>Iface I\<preceq> Class Object"
| subcls: "G\<turnstile>C\<preceq>\<^sub>C D \<Longrightarrow> G\<turnstile>Class C\<preceq> Class D"
| implmt: "G\<turnstile>C\<leadsto>I \<Longrightarrow> G\<turnstile>Class C\<preceq> Iface I"
@@ -400,11 +400,11 @@
apply (ind_cases "G\<turnstile>S\<preceq>PrimT x")
by auto
-text {*
+text \<open>
These widening lemmata hold in Bali but are to strong for ordinary
Java. They would not work for real Java Integral Types, like short,
long, int. These lemmata are just for documentation and are not used.
- *}
+\<close>
lemma widen_PrimT_strong: "G\<turnstile>PrimT x\<preceq>T \<Longrightarrow> T = PrimT x"
by (ind_cases "G\<turnstile>PrimT x\<preceq>T") simp_all
@@ -412,7 +412,7 @@
lemma widen_PrimT2_strong: "G\<turnstile>S\<preceq>PrimT x \<Longrightarrow> S = PrimT x"
by (ind_cases "G\<turnstile>S\<preceq>PrimT x") simp_all
-text {* Specialized versions for booleans also would work for real Java *}
+text \<open>Specialized versions for booleans also would work for real Java\<close>
lemma widen_Boolean: "G\<turnstile>PrimT Boolean\<preceq>T \<Longrightarrow> T = PrimT Boolean"
by (ind_cases "G\<turnstile>PrimT Boolean\<preceq>T") simp_all
@@ -594,7 +594,7 @@
*)
(* more detailed than necessary for type-safety, see above rules. *)
-inductive --{* narrowing reference conversion, cf. 5.1.5 *}
+inductive \<comment>\<open>narrowing reference conversion, cf. 5.1.5\<close>
narrow :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_\<succ>_" [71,71,71] 70)
for G :: prog
where
@@ -624,18 +624,18 @@
\<exists>t. S=PrimT t \<and> G\<turnstile>PrimT t\<preceq>PrimT pt"
by (ind_cases "G\<turnstile>S\<succ>PrimT pt")
-text {*
+text \<open>
These narrowing lemmata hold in Bali but are to strong for ordinary
Java. They would not work for real Java Integral Types, like short,
long, int. These lemmata are just for documentation and are not used.
- *}
+\<close>
lemma narrow_PrimT_strong: "G\<turnstile>PrimT pt\<succ>T \<Longrightarrow> T=PrimT pt"
by (ind_cases "G\<turnstile>PrimT pt\<succ>T")
lemma narrow_PrimT2_strong: "G\<turnstile>S\<succ>PrimT pt \<Longrightarrow> S=PrimT pt"
by (ind_cases "G\<turnstile>S\<succ>PrimT pt")
-text {* Specialized versions for booleans also would work for real Java *}
+text \<open>Specialized versions for booleans also would work for real Java\<close>
lemma narrow_Boolean: "G\<turnstile>PrimT Boolean\<succ>T \<Longrightarrow> T=PrimT Boolean"
by (ind_cases "G\<turnstile>PrimT Boolean\<succ>T")
@@ -645,7 +645,7 @@
subsubsection "casting relation"
-inductive --{* casting conversion, cf. 5.5 *}
+inductive \<comment>\<open>casting conversion, cf. 5.5\<close>
cast :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_\<preceq>? _" [71,71,71] 70)
for G :: prog
where
--- a/src/HOL/Bali/TypeSafe.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/TypeSafe.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,7 +1,7 @@
(* Title: HOL/Bali/TypeSafe.thy
Author: David von Oheimb and Norbert Schirmer
*)
-subsection {* The type soundness proof for Java *}
+subsection \<open>The type soundness proof for Java\<close>
theory TypeSafe
imports DefiniteAssignmentCorrect Conform
@@ -114,7 +114,7 @@
else G,s\<turnstile>the_In1 v\<Colon>\<preceq>T
| Inr Ts \<Rightarrow> list_all2 (conf G s) (the_In3 v) Ts)"
-text {*
+text \<open>
With @{term rconf} we describe the conformance of the result value of a term.
This definition gets rather complicated because of the relations between the
injections of the different terms, types and values. The main case distinction
@@ -132,7 +132,7 @@
local variables are allowed to be @{term None}, since the definedness is not
ensured by conformance but by definite assignment. Field and array variables
must contain a value.
-*}
+\<close>
@@ -588,8 +588,8 @@
qed
corollary DynT_mheadsE [consumes 7]:
---{* Same as @{text DynT_mheadsD} but better suited for application in
-typesafety proof *}
+\<comment>\<open>Same as \<open>DynT_mheadsD\<close> but better suited for application in
+typesafety proof\<close>
assumes invC_compatible: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
and wf: "wf_prog G"
and wt_e: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT"
@@ -726,8 +726,8 @@
declare split_paired_All [simp del] split_paired_Ex [simp del]
declare split_if [split del] split_if_asm [split del]
option.split [split del] option.split_asm [split del]
-setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *}
-setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\<close>
+setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
lemma FVar_lemma:
"\<lbrakk>((v, f), Norm s2') = fvar statDeclC (static field) fn a (x2, s2);
@@ -755,8 +755,8 @@
declare split_paired_All [simp] split_paired_Ex [simp]
declare split_if [split] split_if_asm [split]
option.split [split] option.split_asm [split]
-setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac)) *}
-setup {* map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac)) *}
+setup \<open>map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\<close>
+setup \<open>map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac))\<close>
lemma AVar_lemma1: "\<lbrakk>globs s (Inl a) = Some obj;tag obj=Arr ty i;
@@ -871,8 +871,8 @@
declare split_paired_All [simp del] split_paired_Ex [simp del]
declare split_if [split del] split_if_asm [split del]
option.split [split del] option.split_asm [split del]
-setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *}
-setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\<close>
+setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
lemma conforms_init_lvars:
"\<lbrakk>wf_mhead G (pid declC) sig (mhead (mthd dm)); wf_prog G;
@@ -924,8 +924,8 @@
declare split_paired_All [simp] split_paired_Ex [simp]
declare split_if [split] split_if_asm [split]
option.split [split] option.split_asm [split]
-setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac)) *}
-setup {* map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac)) *}
+setup \<open>map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\<close>
+setup \<open>map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac))\<close>
subsection "accessibility"
@@ -1169,7 +1169,7 @@
case Nil thus ?case by simp
next
case (Cons p ps tab qs)
- from `length (p#ps) = length qs`
+ from \<open>length (p#ps) = length qs\<close>
obtain q qs' where qs: "qs=q#qs'" and eq_length: "length ps=length qs'"
by (cases qs) auto
from eq_length have "(tab(p\<mapsto>q))(ps[\<mapsto>]qs'@zs)=(tab(p\<mapsto>q))(ps[\<mapsto>]qs')"
@@ -1185,7 +1185,7 @@
case Nil thus ?case by simp
next
case (Cons p ps tab qs x y)
- from `length (p#ps) = length qs`
+ from \<open>length (p#ps) = length qs\<close>
obtain q qs' where qs: "qs=q#qs'" and eq_length: "length ps=length qs'"
by (cases qs) auto
from eq_length
@@ -1278,14 +1278,14 @@
case Nil thus ?case by simp
next
case (Cons x xs tab ys z)
- note z = `tab vn = Some z`
+ note z = \<open>tab vn = Some z\<close>
show ?case
proof (cases ys)
case Nil
with z show ?thesis by simp
next
case (Cons y ys')
- note ys = `ys = y#ys'`
+ note ys = \<open>ys = y#ys'\<close>
from z obtain z' where "(tab(x\<mapsto>y)) vn = Some z'"
by (rule map_upd_Some_expand [of tab,elim_format]) blast
hence "\<exists>z. ((tab(x\<mapsto>y))(xs[\<mapsto>]ys')) vn = Some z"
@@ -1320,14 +1320,14 @@
case Nil thus ?case by simp
next
case (Cons x xs tab tab' ys z)
- note some = `(tab(x # xs[\<mapsto>]ys)) vn = Some z`
- note tab_not_z = `tab vn \<noteq> Some z`
+ note some = \<open>(tab(x # xs[\<mapsto>]ys)) vn = Some z\<close>
+ note tab_not_z = \<open>tab vn \<noteq> Some z\<close>
show ?case
proof (cases ys)
case Nil with some tab_not_z show ?thesis by simp
next
case (Cons y tl)
- note ys = `ys = y#tl`
+ note ys = \<open>ys = y#tl\<close>
show ?thesis
proof (cases "(tab(x\<mapsto>y)) vn \<noteq> Some z")
case True
@@ -1412,15 +1412,15 @@
case Nil thus ?case by simp
next
case (Cons x xs tab tab' ys)
- note tab_vn = `(tab(x # xs[\<mapsto>]ys)) vn = Some el`
- note tab'_vn = `(tab'(x # xs[\<mapsto>]ys)) vn = None`
+ note tab_vn = \<open>(tab(x # xs[\<mapsto>]ys)) vn = Some el\<close>
+ note tab'_vn = \<open>(tab'(x # xs[\<mapsto>]ys)) vn = None\<close>
show ?case
proof (cases ys)
case Nil
with tab_vn show ?thesis by simp
next
case (Cons y tl)
- note ys = `ys=y#tl`
+ note ys = \<open>ys=y#tl\<close>
with tab_vn tab'_vn
have "(tab(x\<mapsto>y)) vn = Some el"
by - (rule Cons.hyps,auto)
@@ -1497,7 +1497,7 @@
next
case (Cons x xs tab ys)
note Hyp = Cons.hyps
- note len = `length (x#xs)=length ys`
+ note len = \<open>length (x#xs)=length ys\<close>
show ?case
proof (cases ys)
case Nil with len show ?thesis by simp
@@ -1728,7 +1728,7 @@
?ErrorFree s0 s1")
proof (induct)
case (Abrupt xc s t L accC T A)
- from `(Some xc, s)\<Colon>\<preceq>(G,L)`
+ from \<open>(Some xc, s)\<Colon>\<preceq>(G,L)\<close>
show "(Some xc, s)\<Colon>\<preceq>(G,L) \<and>
(normal (Some xc, s)
\<longrightarrow> G,L,store (Some xc,s)\<turnstile>t\<succ>undefined3 t\<Colon>\<preceq>T) \<and>
@@ -1736,19 +1736,19 @@
by simp
next
case (Skip s L accC T A)
- from `Norm s\<Colon>\<preceq>(G, L)` and
- `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r Skip\<Colon>T`
+ from \<open>Norm s\<Colon>\<preceq>(G, L)\<close> and
+ \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r Skip\<Colon>T\<close>
show "Norm s\<Colon>\<preceq>(G, L) \<and>
(normal (Norm s) \<longrightarrow> G,L,store (Norm s)\<turnstile>In1r Skip\<succ>\<diamondsuit>\<Colon>\<preceq>T) \<and>
(error_free (Norm s) = error_free (Norm s))"
by simp
next
case (Expr s0 e v s1 L accC T A)
- note `G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1`
- note hyp = `PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
+ note \<open>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<close>
+ note hyp = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
moreover
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Expr e)\<Colon>T`
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Expr e)\<Colon>T\<close>
then obtain eT
where "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l e\<Colon>eT"
by (rule wt_elim_cases) blast
@@ -1766,10 +1766,10 @@
by (simp)
next
case (Lab s0 c s1 l L accC T A)
- note hyp = `PROP ?TypeSafe (Norm s0) s1 (In1r c) \<diamondsuit>`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
+ note hyp = \<open>PROP ?TypeSafe (Norm s0) s1 (In1r c) \<diamondsuit>\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
moreover
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (l\<bullet> c)\<Colon>T`
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (l\<bullet> c)\<Colon>T\<close>
then have "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
by (rule wt_elim_cases) blast
moreover from Lab.prems obtain C where
@@ -1789,12 +1789,12 @@
by (simp)
next
case (Comp s0 c1 s1 c2 s2 L accC T A)
- note eval_c1 = `G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1`
- note eval_c2 = `G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2`
- note hyp_c1 = `PROP ?TypeSafe (Norm s0) s1 (In1r c1) \<diamondsuit>`
- note hyp_c2 = `PROP ?TypeSafe s1 s2 (In1r c2) \<diamondsuit>`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (c1;; c2)\<Colon>T`
+ note eval_c1 = \<open>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1\<close>
+ note eval_c2 = \<open>G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2\<close>
+ note hyp_c1 = \<open>PROP ?TypeSafe (Norm s0) s1 (In1r c1) \<diamondsuit>\<close>
+ note hyp_c2 = \<open>PROP ?TypeSafe s1 s2 (In1r c2) \<diamondsuit>\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (c1;; c2)\<Colon>T\<close>
then obtain wt_c1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
wt_c2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c2\<Colon>\<surd>"
by (rule wt_elim_cases) blast
@@ -1835,13 +1835,13 @@
qed
next
case (If s0 e b s1 c1 c2 s2 L accC T A)
- note eval_e = `G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1`
- note eval_then_else = `G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2`
- note hyp_e = `PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 b)`
+ note eval_e = \<open>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1\<close>
+ note eval_then_else = \<open>G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2\<close>
+ note hyp_e = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 b)\<close>
note hyp_then_else =
- `PROP ?TypeSafe s1 s2 (In1r (if the_Bool b then c1 else c2)) \<diamondsuit>`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (If(e) c1 Else c2)\<Colon>T`
+ \<open>PROP ?TypeSafe s1 s2 (In1r (if the_Bool b then c1 else c2)) \<diamondsuit>\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (If(e) c1 Else c2)\<Colon>T\<close>
then obtain
wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
wt_then_else: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
@@ -1899,7 +1899,7 @@
with wt show ?thesis
by simp
qed
- -- {* Note that we don't have to show that @{term b} really is a boolean
+ \<comment> \<open>Note that we don't have to show that @{term b} really is a boolean
value. With @{term the_Bool} we enforce to get a value of boolean
type. So execution will be type safe, even if b would be
a string, for example. We might not expect such a behaviour to be
@@ -1907,18 +1907,18 @@
the evaulation rule, so that it only has a type safe evaluation if
we actually get a boolean value for the condition. That b is actually
a boolean value is part of @{term hyp_e}. See also Loop
- *}
+\<close>
next
case (Loop s0 e b s1 c s2 l s3 L accC T A)
- note eval_e = `G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1`
- note hyp_e = `PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 b)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T`
+ note eval_e = \<open>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1\<close>
+ note hyp_e = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 b)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (l\<bullet> While(e) c)\<Colon>T\<close>
then obtain wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
by (rule wt_elim_cases) blast
- note da = `\<lparr>prg=G, cls=accC, lcl=L\<rparr>
- \<turnstile> dom (locals(store ((Norm s0)::state))) \<guillemotright>In1r (l\<bullet> While(e) c)\<guillemotright> A`
+ note da = \<open>\<lparr>prg=G, cls=accC, lcl=L\<rparr>
+ \<turnstile> dom (locals(store ((Norm s0)::state))) \<guillemotright>In1r (l\<bullet> While(e) c)\<guillemotright> A\<close>
then
obtain E C where
da_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>
@@ -2041,7 +2041,7 @@
qed
next
case (Jmp s j L accC T A)
- note `Norm s\<Colon>\<preceq>(G, L)`
+ note \<open>Norm s\<Colon>\<preceq>(G, L)\<close>
moreover
from Jmp.prems
have "j=Ret \<longrightarrow> Result \<in> dom (locals (store ((Norm s)::state)))"
@@ -2055,10 +2055,10 @@
by simp
next
case (Throw s0 e a s1 L accC T A)
- note `G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1`
- note hyp = `PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 a)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Throw e)\<Colon>T`
+ note \<open>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1\<close>
+ note hyp = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 a)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Throw e)\<Colon>T\<close>
then obtain tn
where wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-Class tn" and
throwable: "G\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable"
@@ -2083,11 +2083,11 @@
by simp
next
case (Try s0 c1 s1 s2 catchC vn c2 s3 L accC T A)
- note eval_c1 = `G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1`
- note sx_alloc = `G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2`
- note hyp_c1 = `PROP ?TypeSafe (Norm s0) s1 (In1r c1) \<diamondsuit>`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>In1r (Try c1 Catch(catchC vn) c2)\<Colon>T`
+ note eval_c1 = \<open>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1\<close>
+ note sx_alloc = \<open>G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2\<close>
+ note hyp_c1 = \<open>PROP ?TypeSafe (Norm s0) s1 (In1r c1) \<diamondsuit>\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>In1r (Try c1 Catch(catchC vn) c2)\<Colon>T\<close>
then obtain
wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
wt_c2: "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class catchC)\<rparr>\<turnstile>c2\<Colon>\<surd>" and
@@ -2165,7 +2165,7 @@
have "(dom (locals (store ((Norm s0)::state))) \<union> {VName vn})
\<subseteq> dom (locals (store (new_xcpt_var vn s2)))"
proof -
- from `G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1`
+ from \<open>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1\<close>
have "dom (locals (store ((Norm s0)::state)))
\<subseteq> dom (locals (store s1))"
by (rule dom_locals_eval_mono_elim)
@@ -2200,15 +2200,15 @@
qed
next
case (Fin s0 c1 x1 s1 c2 s2 s3 L accC T A)
- note eval_c1 = `G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1, s1)`
- note eval_c2 = `G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2`
- note s3 = `s3 = (if \<exists>err. x1 = Some (Error err)
+ note eval_c1 = \<open>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1, s1)\<close>
+ note eval_c2 = \<open>G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2\<close>
+ note s3 = \<open>s3 = (if \<exists>err. x1 = Some (Error err)
then (x1, s1)
- else abupd (abrupt_if (x1 \<noteq> None) x1) s2)`
- note hyp_c1 = `PROP ?TypeSafe (Norm s0) (x1,s1) (In1r c1) \<diamondsuit>`
- note hyp_c2 = `PROP ?TypeSafe (Norm s1) s2 (In1r c2) \<diamondsuit>`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (c1 Finally c2)\<Colon>T`
+ else abupd (abrupt_if (x1 \<noteq> None) x1) s2)\<close>
+ note hyp_c1 = \<open>PROP ?TypeSafe (Norm s0) (x1,s1) (In1r c1) \<diamondsuit>\<close>
+ note hyp_c2 = \<open>PROP ?TypeSafe (Norm s1) s2 (In1r c2) \<diamondsuit>\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (c1 Finally c2)\<Colon>T\<close>
then obtain
wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
wt_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c2\<Colon>\<surd>"
@@ -2270,9 +2270,9 @@
qed
next
case (Init C c s0 s3 s1 s2 L accC T A)
- note cls = `the (class G C) = c`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Init C)\<Colon>T`
+ note cls = \<open>the (class G C) = c\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1r (Init C)\<Colon>T\<close>
with cls
have cls_C: "class G C = Some c"
by - (erule wt_elim_cases, auto)
@@ -2376,12 +2376,12 @@
qed
next
case (NewC s0 C s1 a s2 L accC T A)
- note `G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1`
- note halloc = `G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2`
- note hyp = `PROP ?TypeSafe (Norm s0) s1 (In1r (Init C)) \<diamondsuit>`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
+ note \<open>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1\<close>
+ note halloc = \<open>G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2\<close>
+ note hyp = \<open>PROP ?TypeSafe (Norm s0) s1 (In1r (Init C)) \<diamondsuit>\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
moreover
- note wt = `\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>In1l (NewC C)\<Colon>T`
+ note wt = \<open>\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>In1l (NewC C)\<Colon>T\<close>
then obtain is_cls_C: "is_class G C" and
T: "T=Inl (Class C)"
by (rule wt_elim_cases) (auto dest: is_acc_classD)
@@ -2408,13 +2408,13 @@
by auto
next
case (NewA s0 elT s1 e i s2 a s3 L accC T A)
- note eval_init = `G\<turnstile>Norm s0 \<midarrow>init_comp_ty elT\<rightarrow> s1`
- note eval_e = `G\<turnstile>s1 \<midarrow>e-\<succ>i\<rightarrow> s2`
- note halloc = `G\<turnstile>abupd (check_neg i) s2\<midarrow>halloc Arr elT (the_Intg i)\<succ>a\<rightarrow> s3`
- note hyp_init = `PROP ?TypeSafe (Norm s0) s1 (In1r (init_comp_ty elT)) \<diamondsuit>`
- note hyp_size = `PROP ?TypeSafe s1 s2 (In1l e) (In1 i)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (New elT[e])\<Colon>T`
+ note eval_init = \<open>G\<turnstile>Norm s0 \<midarrow>init_comp_ty elT\<rightarrow> s1\<close>
+ note eval_e = \<open>G\<turnstile>s1 \<midarrow>e-\<succ>i\<rightarrow> s2\<close>
+ note halloc = \<open>G\<turnstile>abupd (check_neg i) s2\<midarrow>halloc Arr elT (the_Intg i)\<succ>a\<rightarrow> s3\<close>
+ note hyp_init = \<open>PROP ?TypeSafe (Norm s0) s1 (In1r (init_comp_ty elT)) \<diamondsuit>\<close>
+ note hyp_size = \<open>PROP ?TypeSafe s1 s2 (In1l e) (In1 i)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (New elT[e])\<Colon>T\<close>
then obtain
wt_init: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>init_comp_ty elT\<Colon>\<surd>" and
wt_size: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-PrimT Integer" and
@@ -2479,11 +2479,11 @@
by simp
next
case (Cast s0 e v s1 s2 castT L accC T A)
- note `G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1`
- note s2 = `s2 = abupd (raise_if (\<not> G,store s1\<turnstile>v fits castT) ClassCast) s1`
- note hyp = `PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (Cast castT e)\<Colon>T`
+ note \<open>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<close>
+ note s2 = \<open>s2 = abupd (raise_if (\<not> G,store s1\<turnstile>v fits castT) ClassCast) s1\<close>
+ note hyp = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (Cast castT e)\<Colon>T\<close>
then obtain eT
where wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" and
eT: "G\<turnstile>eT\<preceq>? castT" and
@@ -2525,8 +2525,8 @@
by blast
next
case (Inst s0 e v s1 b instT L accC T A)
- note hyp = `PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
+ note hyp = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
from Inst.prems obtain eT
where wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-RefT eT" and
T: "T=Inl (PrimT Boolean)"
@@ -2549,9 +2549,9 @@
intro: conf_litval simp add: empty_dt_def)
next
case (UnOp s0 e v s1 unop L accC T A)
- note hyp = `PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (UnOp unop e)\<Colon>T`
+ note hyp = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (UnOp unop e)\<Colon>T\<close>
then obtain eT
where wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" and
wt_unop: "wt_unop unop eT" and
@@ -2576,15 +2576,15 @@
by simp
next
case (BinOp s0 e1 v1 s1 binop e2 v2 s2 L accC T A)
- note eval_e1 = `G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1`
- note eval_e2 = `G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
- else In1r Skip)\<succ>\<rightarrow> (In1 v2, s2)`
- note hyp_e1 = `PROP ?TypeSafe (Norm s0) s1 (In1l e1) (In1 v1)`
- note hyp_e2 = `PROP ?TypeSafe s1 s2
+ note eval_e1 = \<open>G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1\<close>
+ note eval_e2 = \<open>G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then In1l e2
+ else In1r Skip)\<succ>\<rightarrow> (In1 v2, s2)\<close>
+ note hyp_e1 = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e1) (In1 v1)\<close>
+ note hyp_e2 = \<open>PROP ?TypeSafe s1 s2
(if need_second_arg binop v1 then In1l e2 else In1r Skip)
- (In1 v2)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (BinOp binop e1 e2)\<Colon>T`
+ (In1 v2)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (BinOp binop e1 e2)\<Colon>T\<close>
then obtain e1T e2T where
wt_e1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e1\<Colon>-e1T" and
wt_e2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e2\<Colon>-e2T" and
@@ -2597,8 +2597,8 @@
daSkip: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store s1)) \<guillemotright>In1r Skip\<guillemotright> S"
by (auto intro: da_Skip [simplified] assigned.select_convs)
- note da = `\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store ((Norm s0::state))))
- \<guillemotright>\<langle>BinOp binop e1 e2\<rangle>\<^sub>e\<guillemotright> A`
+ note da = \<open>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store ((Norm s0::state))))
+ \<guillemotright>\<langle>BinOp binop e1 e2\<rangle>\<^sub>e\<guillemotright> A\<close>
then obtain E1 where
da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store ((Norm s0)::state))) \<guillemotright>In1l e1\<guillemotright> E1"
@@ -2612,20 +2612,20 @@
have conf_v:
"G,L,snd s2\<turnstile>In1l (BinOp binop e1 e2)\<succ>In1 (eval_binop binop v1 v2)\<Colon>\<preceq>T"
by (cases binop) auto
- -- {* Note that we don't use the information that v1 really is compatible
+ \<comment> \<open>Note that we don't use the information that v1 really is compatible
with the expected type e1T and v2 is compatible with e2T,
- because @{text eval_binop} will anyway produce an output of
+ because \<open>eval_binop\<close> will anyway produce an output of
the right type.
So evaluating the addition of an integer with a string is type
safe. This is a little bit annoying since we may regard such a
behaviour as not type safe.
- If we want to avoid this we can redefine @{text eval_binop} so that
+ If we want to avoid this we can redefine \<open>eval_binop\<close> so that
it only produces a output of proper type if it is assigned to
values of the expected types, and arbitrary if the inputs have
unexpected types. The proof can easily be adapted since we
have the hypothesis that the values have a proper type.
This also applies to unary operations.
- *}
+\<close>
from eval_e1 have
s0_s1:"dom (locals (store ((Norm s0)::state))) \<subseteq> dom (locals (store s1))"
by (rule dom_locals_eval_mono_elim)
@@ -2667,8 +2667,8 @@
qed
next
case (Super s L accC T A)
- note conf_s = `Norm s\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l Super\<Colon>T`
+ note conf_s = \<open>Norm s\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l Super\<Colon>T\<close>
then obtain C c where
C: "L This = Some (Class C)" and
neq_Obj: "C\<noteq>Object" and
@@ -2692,8 +2692,8 @@
by simp
next
case (Acc s0 v w upd s1 L accC T A)
- note hyp = `PROP ?TypeSafe (Norm s0) s1 (In2 v) (In2 (w,upd))`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
+ note hyp = \<open>PROP ?TypeSafe (Norm s0) s1 (In2 v) (In2 (w,upd))\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
from Acc.prems obtain vT where
wt_v: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>v\<Colon>=vT" and
T: "T=Inl vT"
@@ -2712,7 +2712,7 @@
also
have "dom (locals s0) \<subseteq> dom (locals (store s1))"
proof -
- from `G\<turnstile>Norm s0 \<midarrow>v=\<succ>(w, upd)\<rightarrow> s1`
+ from \<open>G\<turnstile>Norm s0 \<midarrow>v=\<succ>(w, upd)\<rightarrow> s1\<close>
show ?thesis
by (rule dom_locals_eval_mono_elim) simp
qed
@@ -2732,12 +2732,12 @@
by simp
next
case (Ass s0 var w upd s1 e v s2 L accC T A)
- note eval_var = `G\<turnstile>Norm s0 \<midarrow>var=\<succ>(w, upd)\<rightarrow> s1`
- note eval_e = `G\<turnstile>s1 \<midarrow>e-\<succ>v\<rightarrow> s2`
- note hyp_var = `PROP ?TypeSafe (Norm s0) s1 (In2 var) (In2 (w,upd))`
- note hyp_e = `PROP ?TypeSafe s1 s2 (In1l e) (In1 v)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (var:=e)\<Colon>T`
+ note eval_var = \<open>G\<turnstile>Norm s0 \<midarrow>var=\<succ>(w, upd)\<rightarrow> s1\<close>
+ note eval_e = \<open>G\<turnstile>s1 \<midarrow>e-\<succ>v\<rightarrow> s2\<close>
+ note hyp_var = \<open>PROP ?TypeSafe (Norm s0) s1 (In2 var) (In2 (w,upd))\<close>
+ note hyp_e = \<open>PROP ?TypeSafe s1 s2 (In1l e) (In1 v)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (var:=e)\<Colon>T\<close>
then obtain varT eT where
wt_var: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>var\<Colon>=varT" and
wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" and
@@ -2889,13 +2889,13 @@
qed
next
case (Cond s0 e0 b s1 e1 e2 v s2 L accC T A)
- note eval_e0 = `G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1`
- note eval_e1_e2 = `G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2`
- note hyp_e0 = `PROP ?TypeSafe (Norm s0) s1 (In1l e0) (In1 b)`
- note hyp_if = `PROP ?TypeSafe s1 s2
- (In1l (if the_Bool b then e1 else e2)) (In1 v)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (e0 ? e1 : e2)\<Colon>T`
+ note eval_e0 = \<open>G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1\<close>
+ note eval_e1_e2 = \<open>G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2\<close>
+ note hyp_e0 = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e0) (In1 b)\<close>
+ note hyp_if = \<open>PROP ?TypeSafe s1 s2
+ (In1l (if the_Bool b then e1 else e2)) (In1 v)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (e0 ? e1 : e2)\<Colon>T\<close>
then obtain T1 T2 statT where
wt_e0: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e0\<Colon>-PrimT Boolean" and
wt_e1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e1\<Colon>-T1" and
@@ -2983,24 +2983,24 @@
next
case (Call s0 e a s1 args vs s2 invDeclC mode statT mn pTs' s3 s3' accC'
v s4 L accC T A)
- note eval_e = `G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1`
- note eval_args = `G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2`
- note invDeclC = `invDeclC
+ note eval_e = \<open>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1\<close>
+ note eval_args = \<open>G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2\<close>
+ note invDeclC = \<open>invDeclC
= invocation_declclass G mode (store s2) a statT
- \<lparr>name = mn, parTs = pTs'\<rparr>`
+ \<lparr>name = mn, parTs = pTs'\<rparr>\<close>
note init_lvars =
- `s3 = init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> mode a vs s2`
- note check = `s3' =
- check_method_access G accC' statT mode \<lparr>name = mn, parTs = pTs'\<rparr> a s3`
+ \<open>s3 = init_lvars G invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> mode a vs s2\<close>
+ note check = \<open>s3' =
+ check_method_access G accC' statT mode \<lparr>name = mn, parTs = pTs'\<rparr> a s3\<close>
note eval_methd =
- `G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<rightarrow> s4`
- note hyp_e = `PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 a)`
- note hyp_args = `PROP ?TypeSafe s1 s2 (In3 args) (In3 vs)`
- note hyp_methd = `PROP ?TypeSafe s3' s4
- (In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)) (In1 v)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg=G, cls=accC, lcl=L\<rparr>
- \<turnstile>In1l ({accC',statT,mode}e\<cdot>mn( {pTs'}args))\<Colon>T`
+ \<open>G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>-\<succ>v\<rightarrow> s4\<close>
+ note hyp_e = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 a)\<close>
+ note hyp_args = \<open>PROP ?TypeSafe s1 s2 (In3 args) (In3 vs)\<close>
+ note hyp_methd = \<open>PROP ?TypeSafe s3' s4
+ (In1l (Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)) (In1 v)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg=G, cls=accC, lcl=L\<rparr>
+ \<turnstile>In1l ({accC',statT,mode}e\<cdot>mn( {pTs'}args))\<Colon>T\<close>
from wt obtain pTs statDeclT statM where
wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT" and
wt_args: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>args\<Colon>\<doteq>pTs" and
@@ -3325,10 +3325,10 @@
qed
next
case (Methd s0 D sig v s1 L accC T A)
- note `G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<rightarrow> s1`
- note hyp = `PROP ?TypeSafe (Norm s0) s1 (In1l (body G D sig)) (In1 v)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (Methd D sig)\<Colon>T`
+ note \<open>G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<rightarrow> s1\<close>
+ note hyp = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l (body G D sig)) (In1 v)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (Methd D sig)\<Colon>T\<close>
then obtain m bodyT where
D: "is_class G D" and
m: "methd G D sig = Some m" and
@@ -3350,12 +3350,12 @@
by (auto simp add: Let_def body_def)
next
case (Body s0 D s1 c s2 s3 L accC T A)
- note eval_init = `G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1`
- note eval_c = `G\<turnstile>s1 \<midarrow>c\<rightarrow> s2`
- note hyp_init = `PROP ?TypeSafe (Norm s0) s1 (In1r (Init D)) \<diamondsuit>`
- note hyp_c = `PROP ?TypeSafe s1 s2 (In1r c) \<diamondsuit>`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (Body D c)\<Colon>T`
+ note eval_init = \<open>G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1\<close>
+ note eval_c = \<open>G\<turnstile>s1 \<midarrow>c\<rightarrow> s2\<close>
+ note hyp_init = \<open>PROP ?TypeSafe (Norm s0) s1 (In1r (Init D)) \<diamondsuit>\<close>
+ note hyp_c = \<open>PROP ?TypeSafe s1 s2 (In1r c) \<diamondsuit>\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In1l (Body D c)\<Colon>T\<close>
then obtain bodyT where
iscls_D: "is_class G D" and
wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>" and
@@ -3413,10 +3413,10 @@
have "\<And> j. abrupt s2 = Some (Jump j) \<Longrightarrow> j=Ret"
by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
moreover
- note `s3 =
+ note \<open>s3 =
(if \<exists>l. abrupt s2 = Some (Jump (Break l)) \<or>
abrupt s2 = Some (Jump (Cont l))
- then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)`
+ then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)\<close>
ultimately show ?thesis
by force
qed
@@ -3451,8 +3451,8 @@
by (cases s2) (auto intro: conforms_locals)
next
case (LVar s vn L accC T)
- note conf_s = `Norm s\<Colon>\<preceq>(G, L)` and
- wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In2 (LVar vn)\<Colon>T`
+ note conf_s = \<open>Norm s\<Colon>\<preceq>(G, L)\<close> and
+ wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In2 (LVar vn)\<Colon>T\<close>
then obtain vnT where
vnT: "L vn = Some vnT" and
T: "T=Inl vnT"
@@ -3474,14 +3474,14 @@
by (simp add: lvar_def)
next
case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC L accC' T A)
- note eval_init = `G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1`
- note eval_e = `G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2`
- note fvar = `(v, s2') = fvar statDeclC stat fn a s2`
- note check = `s3 = check_field_access G accC statDeclC fn stat a s2'`
- note hyp_init = `PROP ?TypeSafe (Norm s0) s1 (In1r (Init statDeclC)) \<diamondsuit>`
- note hyp_e = `PROP ?TypeSafe s1 s2 (In1l e) (In1 a)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile>In2 ({accC,statDeclC,stat}e..fn)\<Colon>T`
+ note eval_init = \<open>G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1\<close>
+ note eval_e = \<open>G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2\<close>
+ note fvar = \<open>(v, s2') = fvar statDeclC stat fn a s2\<close>
+ note check = \<open>s3 = check_field_access G accC statDeclC fn stat a s2'\<close>
+ note hyp_init = \<open>PROP ?TypeSafe (Norm s0) s1 (In1r (Init statDeclC)) \<diamondsuit>\<close>
+ note hyp_e = \<open>PROP ?TypeSafe s1 s2 (In1l e) (In1 a)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile>In2 ({accC,statDeclC,stat}e..fn)\<Colon>T\<close>
then obtain statC f where
wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-Class statC" and
accfield: "accfield G accC statC fn = Some (statDeclC,f)" and
@@ -3590,13 +3590,13 @@
by auto
next
case (AVar s0 e1 a s1 e2 i s2 v s2' L accC T A)
- note eval_e1 = `G\<turnstile>Norm s0 \<midarrow>e1-\<succ>a\<rightarrow> s1`
- note eval_e2 = `G\<turnstile>s1 \<midarrow>e2-\<succ>i\<rightarrow> s2`
- note hyp_e1 = `PROP ?TypeSafe (Norm s0) s1 (In1l e1) (In1 a)`
- note hyp_e2 = `PROP ?TypeSafe s1 s2 (In1l e2) (In1 i)`
- note avar = `(v, s2') = avar G i a s2`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In2 (e1.[e2])\<Colon>T`
+ note eval_e1 = \<open>G\<turnstile>Norm s0 \<midarrow>e1-\<succ>a\<rightarrow> s1\<close>
+ note eval_e2 = \<open>G\<turnstile>s1 \<midarrow>e2-\<succ>i\<rightarrow> s2\<close>
+ note hyp_e1 = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e1) (In1 a)\<close>
+ note hyp_e2 = \<open>PROP ?TypeSafe s1 s2 (In1l e2) (In1 i)\<close>
+ note avar = \<open>(v, s2') = avar G i a s2\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In2 (e1.[e2])\<Colon>T\<close>
then obtain elemT
where wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-elemT.[]" and
wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-PrimT Integer" and
@@ -3690,12 +3690,12 @@
by (auto elim!: wt_elim_cases)
next
case (Cons s0 e v s1 es vs s2 L accC T A)
- note eval_e = `G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1`
- note eval_es = `G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<rightarrow> s2`
- note hyp_e = `PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)`
- note hyp_es = `PROP ?TypeSafe s1 s2 (In3 es) (In3 vs)`
- note conf_s0 = `Norm s0\<Colon>\<preceq>(G, L)`
- note wt = `\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In3 (e # es)\<Colon>T`
+ note eval_e = \<open>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<close>
+ note eval_es = \<open>G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<rightarrow> s2\<close>
+ note hyp_e = \<open>PROP ?TypeSafe (Norm s0) s1 (In1l e) (In1 v)\<close>
+ note hyp_es = \<open>PROP ?TypeSafe s1 s2 (In3 es) (In3 vs)\<close>
+ note conf_s0 = \<open>Norm s0\<Colon>\<preceq>(G, L)\<close>
+ note wt = \<open>\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>In3 (e # es)\<Colon>T\<close>
then obtain eT esT where
wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT" and
wt_es: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>es\<Colon>\<doteq>esT" and
@@ -3749,10 +3749,10 @@
from this and conf_s0 wt da show ?thesis .
qed
-text {*
+text \<open>
-*} (* dummy text command to break paragraph for latex;
+\<close> (* dummy text command to break paragraph for latex;
large paragraphs exhaust memory of debian pdflatex *)
corollary eval_type_soundE [consumes 5]:
@@ -3823,7 +3823,7 @@
subsection "Ideas for the future"
-text {* In the type soundness proof and the correctness proof of
+text \<open>In the type soundness proof and the correctness proof of
definite assignment we perform induction on the evaluation relation with the
further preconditions that the term is welltyped and definitely assigned. During
the proofs we have to establish the welltypedness and definite assignment of
@@ -3833,7 +3833,7 @@
evaluation of a wellformed term, were these propagations is already done, once
and forever.
Then we can do the proofs with this rule and can enjoy the time we have saved.
-Here is a first and incomplete sketch of such a rule.*}
+Here is a first and incomplete sketch of such a rule.\<close>
theorem wellformed_eval_induct [consumes 4, case_names Abrupt Skip Expr Lab
Comp If]:
assumes eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)"
@@ -3926,8 +3926,8 @@
by (rule lab)
next
case (Comp s0 c1 s1 c2 s2 L accC T A)
- note eval_c1 = `G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1`
- note eval_c2 = `G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2`
+ note eval_c1 = \<open>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1\<close>
+ note eval_c2 = \<open>G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2\<close>
from Comp.prems obtain
wt_c1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
wt_c2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c2\<Colon>\<surd>"
@@ -3969,8 +3969,8 @@
by (rule comp) iprover+
next
case (If s0 e b s1 c1 c2 s2 L accC T A)
- note eval_e = `G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1`
- note eval_then_else = `G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2`
+ note eval_e = \<open>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1\<close>
+ note eval_then_else = \<open>G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2\<close>
from If.prems
obtain
wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
--- a/src/HOL/Bali/Value.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/Value.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,20 +1,20 @@
(* Title: HOL/Bali/Value.thy
Author: David von Oheimb
*)
-subsection {* Java values *}
+subsection \<open>Java values\<close>
theory Value imports Type begin
-typedecl loc --{* locations, i.e. abstract references on objects *}
+typedecl loc \<comment>\<open>locations, i.e. abstract references on objects\<close>
datatype val
- = Unit --{* dummy result value of void methods *}
- | Bool bool --{* Boolean value *}
- | Intg int --{* integer value *}
- | Null --{* null reference *}
- | Addr loc --{* addresses, i.e. locations of objects *}
+ = Unit \<comment>\<open>dummy result value of void methods\<close>
+ | Bool bool \<comment>\<open>Boolean value\<close>
+ | Intg int \<comment>\<open>integer value\<close>
+ | Null \<comment>\<open>null reference\<close>
+ | Addr loc \<comment>\<open>addresses, i.e. locations of objects\<close>
primrec the_Bool :: "val \<Rightarrow> bool"
@@ -36,13 +36,13 @@
| "typeof dt Null = Some NT"
| "typeof dt (Addr a) = dt a"
-primrec defpval :: "prim_ty \<Rightarrow> val" --{* default value for primitive types *}
+primrec defpval :: "prim_ty \<Rightarrow> val" \<comment>\<open>default value for primitive types\<close>
where
"defpval Void = Unit"
| "defpval Boolean = Bool False"
| "defpval Integer = Intg 0"
-primrec default_val :: "ty \<Rightarrow> val" --{* default value for all types *}
+primrec default_val :: "ty \<Rightarrow> val" \<comment>\<open>default value for all types\<close>
where
"default_val (PrimT pt) = defpval pt"
| "default_val (RefT r ) = Null"
--- a/src/HOL/Bali/WellForm.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/WellForm.thy Sat Jan 02 18:48:45 2016 +0100
@@ -2,10 +2,10 @@
Author: David von Oheimb and Norbert Schirmer
*)
-subsection {* Well-formedness of Java programs *}
+subsection \<open>Well-formedness of Java programs\<close>
theory WellForm imports DefiniteAssignment begin
-text {*
+text \<open>
For static checks on expressions and statements, see WellType.thy
improvements over Java Specification 1.0 (cf. 8.4.6.3, 8.4.6.4, 9.4.1):
@@ -25,11 +25,11 @@
\item Object and standard exceptions are assumed to be declared like normal
classes
\end{itemize}
-*}
+\<close>
subsubsection "well-formed field declarations"
-text {* well-formed field declaration (common part for classes and interfaces),
- cf. 8.3 and (9.3) *}
+text \<open>well-formed field declaration (common part for classes and interfaces),
+ cf. 8.3 and (9.3)\<close>
definition
wf_fdecl :: "prog \<Rightarrow> pname \<Rightarrow> fdecl \<Rightarrow> bool"
@@ -46,7 +46,7 @@
(*well-formed method declaration,cf. 8.4, 8.4.1, 8.4.3, 8.4.5, 14.3.2, (9.4)*)
(* cf. 14.15, 15.7.2, for scope issues cf. 8.4.1 and 14.3.2 *)
-text {*
+text \<open>
A method head is wellformed if:
\begin{itemize}
\item the signature and the method head agree in the number of parameters
@@ -54,7 +54,7 @@
\item the result type is visible
\item the parameter names are unique
\end{itemize}
-*}
+\<close>
definition
wf_mhead :: "prog \<Rightarrow> pname \<Rightarrow> sig \<Rightarrow> mhead \<Rightarrow> bool" where
"wf_mhead G P = (\<lambda> sig mh. length (parTs sig) = length (pars mh) \<and>
@@ -63,7 +63,7 @@
distinct (pars mh))"
-text {*
+text \<open>
A method declaration is wellformed if:
\begin{itemize}
\item the method head is wellformed
@@ -76,7 +76,7 @@
the parameters the special result variable (Res) and This are assoziated
with there types.
\end{itemize}
-*}
+\<close>
definition
callee_lcl :: "qtname \<Rightarrow> sig \<Rightarrow> methd \<Rightarrow> lenv" where
@@ -205,7 +205,7 @@
subsubsection "well-formed interface declarations"
(* well-formed interface declaration, cf. 9.1, 9.1.2.1, 9.1.3, 9.4 *)
-text {*
+text \<open>
A interface declaration is wellformed if:
\begin{itemize}
\item the interface hierarchy is wellstructured
@@ -219,7 +219,7 @@
\item the result type of a method overriding a set of methods defined in the
superinterfaces widens to each of the corresponding result types
\end{itemize}
-*}
+\<close>
definition
wf_idecl :: "prog \<Rightarrow> idecl \<Rightarrow> bool" where
"wf_idecl G =
@@ -277,7 +277,7 @@
(* well-formed class declaration, cf. 8.1, 8.1.2.1, 8.1.2.2, 8.1.3, 8.1.4 and
class method declaration, cf. 8.4.3.3, 8.4.6.1, 8.4.6.2, 8.4.6.3, 8.4.6.4 *)
-text {*
+text \<open>
A class declaration is wellformed if:
\begin{itemize}
\item there is no interface with the same name
@@ -320,7 +320,7 @@
\end{itemize}
\end{itemize}
-*}
+\<close>
(* to Table *)
definition
entails :: "('a,'b) table \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool" ("_ entails _" 20)
@@ -503,7 +503,7 @@
subsubsection "well-formed programs"
(* well-formed program, cf. 8.1, 9.1 *)
-text {*
+text \<open>
A program declaration is wellformed if:
\begin{itemize}
\item the class ObjectC of Object is defined
@@ -512,12 +512,12 @@
necessary since every interface automatically inherits from Object.
We must know, that every time a Object method is "overriden" by an
interface method this is also overriden by the class implementing the
- the interface (see @{text "implement_dynmethd and class_mheadsD"})
+ the interface (see \<open>implement_dynmethd and class_mheadsD\<close>)
\item all standard Exceptions are defined
\item all defined interfaces are wellformed
\item all defined classes are wellformed
\end{itemize}
-*}
+\<close>
definition
wf_prog :: "prog \<Rightarrow> bool" where
"wf_prog G = (let is = ifaces G; cs = classes G in
@@ -811,7 +811,7 @@
by (cases new, cases old) auto
qed
-text {* Compare this lemma about static
+text \<open>Compare this lemma about static
overriding @{term "G \<turnstile>new overrides\<^sub>S old"} with the definition of
dynamic overriding @{term "G \<turnstile>new overrides old"}.
Conforming result types and restrictions on the access modifiers of the old
@@ -820,8 +820,8 @@
no restrictions on the access modifiers but enforces confrom result types
as precondition. But with some efford we can guarantee the access modifier
restriction for dynamic overriding, too. See lemma
-@{text wf_prog_dyn_override_prop}.
-*}
+\<open>wf_prog_dyn_override_prop\<close>.
+\<close>
lemma wf_prog_stat_overridesD:
assumes stat_override: "G \<turnstile>new overrides\<^sub>S old" and wf: "wf_prog G"
shows
@@ -1758,7 +1758,7 @@
then show ?case by (blast elim: bexI')
next
case (step I SI)
- from `G\<turnstile>I \<prec>I1 SI`
+ from \<open>G\<turnstile>I \<prec>I1 SI\<close>
obtain i where
ifI: "iface G I = Some i" and
SI: "SI \<in> set (isuperIfs i)"
@@ -2048,7 +2048,7 @@
qed
qed
-text {*
+text \<open>
Which dynamic classes are valid to look up a member of a distinct static type?
We have to distinct class members (named static members in Java)
from instance members. Class members are global to all Objects of a class,
@@ -2076,7 +2076,7 @@
methods at all, we have to lookup methods in the base class Object.
The limitation to classes in the field column is artificial and comes out
-of the typing rule for the field access (see rule @{text "FVar"} in the
+of the typing rule for the field access (see rule \<open>FVar\<close> in the
welltyping relation @{term "wt"} in theory WellType).
I stems out of the fact, that Object
indeed has no non private fields. So interfaces and arrays can actually
@@ -2091,7 +2091,7 @@
Iface Object
Class dynC
Array Object
-*}
+\<close>
primrec valid_lookup_cls:: "prog \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> bool \<Rightarrow> bool"
("_,_ \<turnstile> _ valid'_lookup'_cls'_for _" [61,61,61,61] 60)
where
@@ -2127,8 +2127,8 @@
qed
declare split_paired_All [simp del] split_paired_Ex [simp del]
-setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *}
-setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\<close>
+setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
lemma dynamic_mheadsD:
"\<lbrakk>emh \<in> mheads G S statT sig;
@@ -2257,8 +2257,8 @@
qed
qed
declare split_paired_All [simp] split_paired_Ex [simp]
-setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac)) *}
-setup {* map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac)) *}
+setup \<open>map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\<close>
+setup \<open>map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac))\<close>
(* Tactical version *)
(*
@@ -2401,8 +2401,8 @@
declare split_paired_All [simp del] split_paired_Ex [simp del]
-setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *}
-setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\<close>
+setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
lemma wt_is_type: "E,dt\<Turnstile>v\<Colon>T \<Longrightarrow> wf_prog (prg E) \<longrightarrow>
dt=empty_dt \<longrightarrow> (case T of
@@ -2426,8 +2426,8 @@
)
done
declare split_paired_All [simp] split_paired_Ex [simp]
-setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac)) *}
-setup {* map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac)) *}
+setup \<open>map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\<close>
+setup \<open>map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac))\<close>
lemma ty_expr_is_type:
"\<lbrakk>E\<turnstile>e\<Colon>-T; wf_prog (prg E)\<rbrakk> \<Longrightarrow> is_type (prg E) T"
@@ -2855,7 +2855,7 @@
dest: acc_modi_le_Dests)
qed
-subsubsection {* Properties of dynamic accessibility *}
+subsubsection \<open>Properties of dynamic accessibility\<close>
lemma dyn_accessible_Private:
assumes dyn_acc: "G \<turnstile> m in C dyn_accessible_from accC" and
@@ -2866,7 +2866,7 @@
show ?thesis
proof (induct)
case (Immediate m C)
- from `G \<turnstile> m in C permits_acc_from accC` and `accmodi m = Private`
+ from \<open>G \<turnstile> m in C permits_acc_from accC\<close> and \<open>accmodi m = Private\<close>
show ?case
by (simp add: permits_acc_def)
next
@@ -2876,9 +2876,9 @@
qed
qed
-text {* @{text dyn_accessible_Package} only works with the @{text wf_prog} assumption.
+text \<open>\<open>dyn_accessible_Package\<close> only works with the \<open>wf_prog\<close> assumption.
Without it. it is easy to leaf the Package!
-*}
+\<close>
lemma dyn_accessible_Package:
"\<lbrakk>G \<turnstile> m in C dyn_accessible_from accC; accmodi m = Package;
wf_prog G\<rbrakk>
@@ -2919,8 +2919,8 @@
qed
qed
-text {* For fields we don't need the wellformedness of the program, since
-there is no overriding *}
+text \<open>For fields we don't need the wellformedness of the program, since
+there is no overriding\<close>
lemma dyn_accessible_field_Package:
assumes dyn_acc: "G \<turnstile> f in C dyn_accessible_from accC" and
pack: "accmodi f = Package" and
@@ -2931,7 +2931,7 @@
show ?thesis
proof (induct)
case (Immediate f C)
- from `G \<turnstile> f in C permits_acc_from accC` and `accmodi f = Package`
+ from \<open>G \<turnstile> f in C permits_acc_from accC\<close> and \<open>accmodi f = Package\<close>
show ?case
by (simp add: permits_acc_def)
next
@@ -2940,9 +2940,9 @@
qed
qed
-text {* @{text dyn_accessible_instance_field_Protected} only works for fields
+text \<open>\<open>dyn_accessible_instance_field_Protected\<close> only works for fields
since methods can break the package bounds due to overriding
-*}
+\<close>
lemma dyn_accessible_instance_field_Protected:
assumes dyn_acc: "G \<turnstile> f in C dyn_accessible_from accC" and
prot: "accmodi f = Protected" and
@@ -2955,7 +2955,7 @@
show ?thesis
proof (induct)
case (Immediate f C)
- note `G \<turnstile> f in C permits_acc_from accC`
+ note \<open>G \<turnstile> f in C permits_acc_from accC\<close>
moreover
assume "accmodi f = Protected" and "is_field f" and "\<not> is_static f" and
"pid (declclass f) \<noteq> pid accC"
@@ -2983,12 +2983,12 @@
assume "accmodi f = Protected" and "is_field f" and "is_static f" and
"pid (declclass f) \<noteq> pid accC"
moreover
- note `G \<turnstile> f in C permits_acc_from accC`
+ note \<open>G \<turnstile> f in C permits_acc_from accC\<close>
ultimately
have "G\<turnstile> accC \<preceq>\<^sub>C declclass f"
by (auto simp add: permits_acc_def)
moreover
- from `G \<turnstile> f member_in C`
+ from \<open>G \<turnstile> f member_in C\<close>
have "G\<turnstile>C \<preceq>\<^sub>C declclass f"
by (rule member_in_class_relation)
ultimately show ?case
--- a/src/HOL/Bali/WellType.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Bali/WellType.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1,13 +1,13 @@
(* Title: HOL/Bali/WellType.thy
Author: David von Oheimb
*)
-subsection {* Well-typedness of Java programs *}
+subsection \<open>Well-typedness of Java programs\<close>
theory WellType
imports DeclConcepts
begin
-text {*
+text \<open>
improvements over Java Specification 1.0:
\begin{itemize}
\item methods of Object can be called upon references of interface or array type
@@ -26,15 +26,15 @@
the dynamic type of objects. Therefore, they can be used for both
checking static types and determining runtime types in transition semantics.
\end{itemize}
-*}
+\<close>
type_synonym lenv
- = "(lname, ty) table" --{* local variables, including This and Result*}
+ = "(lname, ty) table" \<comment>\<open>local variables, including This and Result\<close>
record env =
- prg:: "prog" --{* program *}
- cls:: "qtname" --{* current package and class name *}
- lcl:: "lenv" --{* local environment *}
+ prg:: "prog" \<comment>\<open>program\<close>
+ cls:: "qtname" \<comment>\<open>current package and class name\<close>
+ lcl:: "lenv" \<comment>\<open>local environment\<close>
translations
(type) "lenv" <= (type) "(lname, ty) table"
@@ -44,7 +44,7 @@
abbreviation
- pkg :: "env \<Rightarrow> pname" --{* select the current package from an environment *}
+ pkg :: "env \<Rightarrow> pname" \<comment>\<open>select the current package from an environment\<close>
where "pkg e == pid (cls e)"
subsubsection "Static overloading: maximally specific methods "
@@ -52,7 +52,7 @@
type_synonym
emhead = "ref_ty \<times> mhead"
---{* Some mnemotic selectors for emhead *}
+\<comment>\<open>Some mnemotic selectors for emhead\<close>
definition
"declrefT" :: "emhead \<Rightarrow> ref_ty"
where "declrefT = fst"
@@ -107,20 +107,20 @@
| "mheads G S (ArrayT T) = accObjectmheads G S (ArrayT T)"
definition
- --{* applicable methods, cf. 15.11.2.1 *}
+ \<comment>\<open>applicable methods, cf. 15.11.2.1\<close>
appl_methds :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list) set" where
"appl_methds G S rt = (\<lambda> sig.
{(mh,pTs') |mh pTs'. mh \<in> mheads G S rt \<lparr>name=name sig,parTs=pTs'\<rparr> \<and>
G\<turnstile>(parTs sig)[\<preceq>]pTs'})"
definition
- --{* more specific methods, cf. 15.11.2.2 *}
+ \<comment>\<open>more specific methods, cf. 15.11.2.2\<close>
more_spec :: "prog \<Rightarrow> emhead \<times> ty list \<Rightarrow> emhead \<times> ty list \<Rightarrow> bool" where
"more_spec G = (\<lambda>(mh,pTs). \<lambda>(mh',pTs'). G\<turnstile>pTs[\<preceq>]pTs')"
(*more_spec G \<equiv>\<lambda>((d,h),pTs). \<lambda>((d',h'),pTs'). G\<turnstile>RefT d\<preceq>RefT d'\<and>G\<turnstile>pTs[\<preceq>]pTs'*)
definition
- --{* maximally specific methods, cf. 15.11.2.2 *}
+ \<comment>\<open>maximally specific methods, cf. 15.11.2.2\<close>
max_spec :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list) set" where
"max_spec G S rt sig = {m. m \<in>appl_methds G S rt sig \<and>
(\<forall>m'\<in>appl_methds G S rt sig. more_spec G m' m \<longrightarrow> m'=m)}"
@@ -262,13 +262,13 @@
| "E,dt\<Turnstile>e\<Colon>=T \<equiv> E,dt\<Turnstile>In2 e\<Colon>Inl T"
| "E,dt\<Turnstile>e\<Colon>\<doteq>T \<equiv> E,dt\<Turnstile>In3 e\<Colon>Inr T"
---{* well-typed statements *}
+\<comment>\<open>well-typed statements\<close>
| Skip: "E,dt\<Turnstile>Skip\<Colon>\<surd>"
| Expr: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-T\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>Expr e\<Colon>\<surd>"
- --{* cf. 14.6 *}
+ \<comment>\<open>cf. 14.6\<close>
| Lab: "E,dt\<Turnstile>c\<Colon>\<surd> \<Longrightarrow>
E,dt\<Turnstile>l\<bullet> c\<Colon>\<surd>"
@@ -276,62 +276,62 @@
E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>c1;; c2\<Colon>\<surd>"
- --{* cf. 14.8 *}
+ \<comment>\<open>cf. 14.8\<close>
| If: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
E,dt\<Turnstile>c1\<Colon>\<surd>;
E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>If(e) c1 Else c2\<Colon>\<surd>"
- --{* cf. 14.10 *}
+ \<comment>\<open>cf. 14.10\<close>
| Loop: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
E,dt\<Turnstile>c\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>l\<bullet> While(e) c\<Colon>\<surd>"
- --{* cf. 14.13, 14.15, 14.16 *}
+ \<comment>\<open>cf. 14.13, 14.15, 14.16\<close>
| Jmp: "E,dt\<Turnstile>Jmp jump\<Colon>\<surd>"
- --{* cf. 14.16 *}
+ \<comment>\<open>cf. 14.16\<close>
| Throw: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class tn;
prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>Throw e\<Colon>\<surd>"
- --{* cf. 14.18 *}
+ \<comment>\<open>cf. 14.18\<close>
| Try: "\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable;
lcl E (VName vn)=None; E \<lparr>lcl := lcl E(VName vn\<mapsto>Class tn)\<rparr>,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk>
\<Longrightarrow>
E,dt\<Turnstile>Try c1 Catch(tn vn) c2\<Colon>\<surd>"
- --{* cf. 14.18 *}
+ \<comment>\<open>cf. 14.18\<close>
| Fin: "\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>c1 Finally c2\<Colon>\<surd>"
| Init: "\<lbrakk>is_class (prg E) C\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>Init C\<Colon>\<surd>"
- --{* @{term Init} is created on the fly during evaluation (see Eval.thy).
+ \<comment>\<open>@{term Init} is created on the fly during evaluation (see Eval.thy).
The class isn't necessarily accessible from the points @{term Init}
is called. Therefor we only demand @{term is_class} and not
@{term is_acc_class} here.
- *}
+\<close>
---{* well-typed expressions *}
+\<comment>\<open>well-typed expressions\<close>
- --{* cf. 15.8 *}
+ \<comment>\<open>cf. 15.8\<close>
| NewC: "\<lbrakk>is_acc_class (prg E) (pkg E) C\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>NewC C\<Colon>-Class C"
- --{* cf. 15.9 *}
+ \<comment>\<open>cf. 15.9\<close>
| NewA: "\<lbrakk>is_acc_type (prg E) (pkg E) T;
E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>New T[i]\<Colon>-T.[]"
- --{* cf. 15.15 *}
+ \<comment>\<open>cf. 15.15\<close>
| Cast: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-T; is_acc_type (prg E) (pkg E) T';
prg E\<turnstile>T\<preceq>? T'\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>Cast T' e\<Colon>-T'"
- --{* cf. 15.19.2 *}
+ \<comment>\<open>cf. 15.19.2\<close>
| Inst: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT T; is_acc_type (prg E) (pkg E) (RefT T');
prg E\<turnstile>RefT T\<preceq>? RefT T'\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>e InstOf T'\<Colon>-PrimT Boolean"
- --{* cf. 15.7.1 *}
+ \<comment>\<open>cf. 15.7.1\<close>
| Lit: "\<lbrakk>typeof dt x = Some T\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>Lit x\<Colon>-T"
@@ -344,28 +344,28 @@
\<Longrightarrow>
E,dt\<Turnstile>BinOp binop e1 e2\<Colon>-T"
- --{* cf. 15.10.2, 15.11.1 *}
+ \<comment>\<open>cf. 15.10.2, 15.11.1\<close>
| Super: "\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object;
class (prg E) C = Some c\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>Super\<Colon>-Class (super c)"
- --{* cf. 15.13.1, 15.10.1, 15.12 *}
+ \<comment>\<open>cf. 15.13.1, 15.10.1, 15.12\<close>
| Acc: "\<lbrakk>E,dt\<Turnstile>va\<Colon>=T\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>Acc va\<Colon>-T"
- --{* cf. 15.25, 15.25.1 *}
+ \<comment>\<open>cf. 15.25, 15.25.1\<close>
| Ass: "\<lbrakk>E,dt\<Turnstile>va\<Colon>=T; va \<noteq> LVar This;
E,dt\<Turnstile>v \<Colon>-T';
prg E\<turnstile>T'\<preceq>T\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>va:=v\<Colon>-T'"
- --{* cf. 15.24 *}
+ \<comment>\<open>cf. 15.24\<close>
| Cond: "\<lbrakk>E,dt\<Turnstile>e0\<Colon>-PrimT Boolean;
E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2;
prg E\<turnstile>T1\<preceq>T2 \<and> T = T2 \<or> prg E\<turnstile>T2\<preceq>T1 \<and> T = T1\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>e0 ? e1 : e2\<Colon>-T"
- --{* cf. 15.11.1, 15.11.2, 15.11.3 *}
+ \<comment>\<open>cf. 15.11.1, 15.11.2, 15.11.3\<close>
| Call: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT;
E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;
max_spec (prg E) (cls E) statT \<lparr>name=mn,parTs=pTs\<rparr>
@@ -377,7 +377,7 @@
methd (prg E) C sig = Some m;
E,dt\<Turnstile>Body (declclass m) (stmt (mbody (mthd m)))\<Colon>-T\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>Methd C sig\<Colon>-T"
- --{* The class @{term C} is the dynamic class of the method call
+ \<comment>\<open>The class @{term C} is the dynamic class of the method call
(cf. Eval.thy).
It hasn't got to be directly accessible from the current package
@{term "(pkg E)"}.
@@ -386,42 +386,42 @@
Note that l is just a dummy value. It is only used in the smallstep
semantics. To proof typesafety directly for the smallstep semantics
we would have to assume conformance of l here!
- *}
+\<close>
| Body: "\<lbrakk>is_class (prg E) D;
E,dt\<Turnstile>blk\<Colon>\<surd>;
(lcl E) Result = Some T;
is_type (prg E) T\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>Body D blk\<Colon>-T"
---{* The class @{term D} implementing the method must not directly be
+\<comment>\<open>The class @{term D} implementing the method must not directly be
accessible from the current package @{term "(pkg E)"}, but can also
be indirectly accessible due to inheritance (enshured in @{term Call})
The result type hasn't got to be accessible in Java! (If it is not
accessible you can only assign it to Object).
For dummy value l see rule @{term Methd}.
- *}
+\<close>
---{* well-typed variables *}
+\<comment>\<open>well-typed variables\<close>
- --{* cf. 15.13.1 *}
+ \<comment>\<open>cf. 15.13.1\<close>
| LVar: "\<lbrakk>lcl E vn = Some T; is_acc_type (prg E) (pkg E) T\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>LVar vn\<Colon>=T"
- --{* cf. 15.10.1 *}
+ \<comment>\<open>cf. 15.10.1\<close>
| FVar: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C;
accfield (prg E) (cls E) C fn = Some (statDeclC,f)\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>{cls E,statDeclC,is_static f}e..fn\<Colon>=(type f)"
- --{* cf. 15.12 *}
+ \<comment>\<open>cf. 15.12\<close>
| AVar: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-T.[];
E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>e.[i]\<Colon>=T"
---{* well-typed expression lists *}
+\<comment>\<open>well-typed expression lists\<close>
- --{* cf. 15.11.??? *}
+ \<comment>\<open>cf. 15.11.???\<close>
| Nil: "E,dt\<Turnstile>[]\<Colon>\<doteq>[]"
- --{* cf. 15.11.??? *}
+ \<comment>\<open>cf. 15.11.???\<close>
| Cons: "\<lbrakk>E,dt\<Turnstile>e \<Colon>-T;
E,dt\<Turnstile>es\<Colon>\<doteq>Ts\<rbrakk> \<Longrightarrow>
E,dt\<Turnstile>e#es\<Colon>\<doteq>T#Ts"
@@ -458,7 +458,7 @@
declare not_None_eq [simp del]
declare split_if [split del] split_if_asm [split del]
declare split_paired_All [simp del] split_paired_Ex [simp del]
-setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\<close>
inductive_cases wt_elim_cases [cases set]:
"E,dt\<Turnstile>In2 (LVar vn) \<Colon>T"
@@ -494,7 +494,7 @@
declare not_None_eq [simp]
declare split_if [split] split_if_asm [split]
declare split_paired_All [simp] split_paired_Ex [simp]
-setup {* map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac)) *}
+setup \<open>map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac))\<close>
lemma is_acc_class_is_accessible:
"is_acc_class G P C \<Longrightarrow> G\<turnstile>(Class C) accessible_in P"
@@ -522,9 +522,9 @@
done
-text {* Special versions of some typing rules, better suited to pattern
+text \<open>Special versions of some typing rules, better suited to pattern
match the conclusion (no selectors in the conclusion)
-*}
+\<close>
lemma wt_Call:
"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT; E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;
@@ -588,13 +588,13 @@
apply auto
done
---{* In the special syntax to distinguish the typing judgements for expressions,
+\<comment>\<open>In the special syntax to distinguish the typing judgements for expressions,
statements, variables and expression lists the kind of term corresponds
to the kind of type in the end e.g. An statement (injection @{term In3}
into terms, always has type void (injection @{term Inl} into the generalised
types. The following simplification procedures establish these kinds of
correlation.
- *}
+\<close>
lemma wt_expr_eq: "E,dt\<Turnstile>In1l t\<Colon>U = (\<exists>T. U=Inl T \<and> E,dt\<Turnstile>t\<Colon>-T)"
by (auto, frule wt_Inj_elim, auto)
@@ -608,29 +608,29 @@
lemma wt_stmt_eq: "E,dt\<Turnstile>In1r t\<Colon>U = (U=Inl(PrimT Void)\<and>E,dt\<Turnstile>t\<Colon>\<surd>)"
by (auto, frule wt_Inj_elim, auto, frule wt_Inj_elim, auto)
-simproc_setup wt_expr ("E,dt\<Turnstile>In1l t\<Colon>U") = {*
+simproc_setup wt_expr ("E,dt\<Turnstile>In1l t\<Colon>U") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE
- | _ => SOME (mk_meta_eq @{thm wt_expr_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm wt_expr_eq}))\<close>
-simproc_setup wt_var ("E,dt\<Turnstile>In2 t\<Colon>U") = {*
+simproc_setup wt_var ("E,dt\<Turnstile>In2 t\<Colon>U") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE
- | _ => SOME (mk_meta_eq @{thm wt_var_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm wt_var_eq}))\<close>
-simproc_setup wt_exprs ("E,dt\<Turnstile>In3 t\<Colon>U") = {*
+simproc_setup wt_exprs ("E,dt\<Turnstile>In3 t\<Colon>U") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE
- | _ => SOME (mk_meta_eq @{thm wt_exprs_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm wt_exprs_eq}))\<close>
-simproc_setup wt_stmt ("E,dt\<Turnstile>In1r t\<Colon>U") = {*
+simproc_setup wt_stmt ("E,dt\<Turnstile>In1r t\<Colon>U") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
(_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE
- | _ => SOME (mk_meta_eq @{thm wt_stmt_eq})) *}
+ | _ => SOME (mk_meta_eq @{thm wt_stmt_eq}))\<close>
lemma wt_elim_BinOp:
"\<lbrakk>E,dt\<Turnstile>In1l (BinOp binop e1 e2)\<Colon>T;
@@ -658,14 +658,14 @@
apply (simp_all (no_asm_use) split del: split_if_asm)
apply (safe del: disjE)
(* 17 subgoals *)
-apply (tactic {* ALLGOALS (fn i =>
+apply (tactic \<open>ALLGOALS (fn i =>
if i = 11 then EVERY'
[Rule_Insts.thin_tac @{context} "E,dt\<Turnstile>e0\<Colon>-PrimT Boolean" [(@{binding E}, NONE, NoSyn)],
Rule_Insts.thin_tac @{context} "E,dt\<Turnstile>e1\<Colon>-T1" [(@{binding E}, NONE, NoSyn), (@{binding T1}, NONE, NoSyn)],
Rule_Insts.thin_tac @{context} "E,dt\<Turnstile>e2\<Colon>-T2" [(@{binding E}, NONE, NoSyn), (@{binding T2}, NONE, NoSyn)]] i
- else Rule_Insts.thin_tac @{context} "All P" [(@{binding P}, NONE, NoSyn)] i) *})
+ else Rule_Insts.thin_tac @{context} "All P" [(@{binding P}, NONE, NoSyn)] i)\<close>)
(*apply (safe del: disjE elim!: wt_elim_cases)*)
-apply (tactic {*ALLGOALS (eresolve_tac @{context} @{thms wt_elim_cases})*})
+apply (tactic \<open>ALLGOALS (eresolve_tac @{context} @{thms wt_elim_cases})\<close>)
apply (simp_all (no_asm_use) split del: split_if_asm)
apply (erule_tac [12] V = "All P" for P in thin_rl) (* Call *)
apply (blast del: equalityCE dest: sym [THEN trans])+
--- a/src/HOL/Hoare/Arith2.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare/Arith2.thy Sat Jan 02 18:48:45 2016 +0100
@@ -21,7 +21,7 @@
| "fac (Suc n) = Suc n * fac n"
-subsubsection {* cd *}
+subsubsection \<open>cd\<close>
lemma cd_nnn: "0<n ==> cd n n n"
apply (simp add: cd_def)
@@ -48,7 +48,7 @@
done
-subsubsection {* gcd *}
+subsubsection \<open>gcd\<close>
lemma gcd_nnn: "0<n ==> n = gcd n n"
apply (unfold gcd_def)
@@ -79,7 +79,7 @@
done
-subsubsection {* pow *}
+subsubsection \<open>pow\<close>
lemma sq_pow_div2 [simp]:
"m mod 2 = 0 ==> ((n::nat)*n)^(m div 2) = n^m"
--- a/src/HOL/Hoare/Heap.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare/Heap.thy Sat Jan 02 18:48:45 2016 +0100
@@ -60,12 +60,12 @@
definition distPath :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool"
where "distPath h x as y \<longleftrightarrow> Path h x as y \<and> distinct as"
-text{* The term @{term"distPath h x as y"} expresses the fact that a
+text\<open>The term @{term"distPath h x as y"} expresses the fact that a
non-repeating path @{term as} connects location @{term x} to location
-@{term y} by means of the @{term h} field. In the case where @{text "x
-= y"}, and there is a cycle from @{term x} to itself, @{term as} can
+@{term y} by means of the @{term h} field. In the case where \<open>x
+= y\<close>, and there is a cycle from @{term x} to itself, @{term as} can
be both @{term "[]"} and the non-repeating list of nodes in the
-cycle. *}
+cycle.\<close>
lemma neq_dP: "p \<noteq> q \<Longrightarrow> Path h p Ps q \<Longrightarrow> distinct Ps \<Longrightarrow>
EX a Qs. p = Ref a & Ps = a#Qs & a \<notin> set Qs"
--- a/src/HOL/Hoare/HeapSyntaxAbort.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare/HeapSyntaxAbort.thy Sat Jan 02 18:48:45 2016 +0100
@@ -7,13 +7,13 @@
subsection "Field access and update"
-text{* Heap update @{text"p^.h := e"} is now guarded against @{term p}
+text\<open>Heap update \<open>p^.h := e\<close> is now guarded against @{term p}
being Null. However, @{term p} may still be illegal,
e.g. uninitialized or dangling. To guard against that, one needs a
more detailed model of the heap where allocated and free addresses are
distinguished, e.g. by making the heap a map, or by carrying the set
of free addresses around. This is needed anyway as soon as we want to
-reason about storage allocation/deallocation. *}
+reason about storage allocation/deallocation.\<close>
syntax
"_refupdate" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a ref \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)"
--- a/src/HOL/Hoare/Hoare_Logic.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare/Hoare_Logic.thy Sat Jan 02 18:48:45 2016 +0100
@@ -54,8 +54,8 @@
("{_} // _ // {_}" [0,55,0] 50)
ML_file "hoare_syntax.ML"
-parse_translation {* [(@{syntax_const "_hoare_vars"}, K Hoare_Syntax.hoare_vars_tr)] *}
-print_translation {* [(@{const_syntax Valid}, K (Hoare_Syntax.spec_tr' @{syntax_const "_hoare"}))] *}
+parse_translation \<open>[(@{syntax_const "_hoare_vars"}, K Hoare_Syntax.hoare_vars_tr)]\<close>
+print_translation \<open>[(@{const_syntax Valid}, K (Hoare_Syntax.spec_tr' @{syntax_const "_hoare"}))]\<close>
lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
@@ -92,16 +92,16 @@
lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
by blast
-lemmas AbortRule = SkipRule -- "dummy version"
+lemmas AbortRule = SkipRule \<comment> "dummy version"
ML_file "hoare_tac.ML"
-method_setup vcg = {*
- Scan.succeed (fn ctxt => SIMPLE_METHOD' (Hoare.hoare_tac ctxt (K all_tac))) *}
+method_setup vcg = \<open>
+ Scan.succeed (fn ctxt => SIMPLE_METHOD' (Hoare.hoare_tac ctxt (K all_tac)))\<close>
"verification condition generator"
-method_setup vcg_simp = {*
+method_setup vcg_simp = \<open>
Scan.succeed (fn ctxt =>
- SIMPLE_METHOD' (Hoare.hoare_tac ctxt (asm_full_simp_tac ctxt))) *}
+ SIMPLE_METHOD' (Hoare.hoare_tac ctxt (asm_full_simp_tac ctxt)))\<close>
"verification condition generator plus simplification"
end
--- a/src/HOL/Hoare/Hoare_Logic_Abort.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy Sat Jan 02 18:48:45 2016 +0100
@@ -56,9 +56,9 @@
("{_} // _ // {_}" [0,55,0] 50)
ML_file "hoare_syntax.ML"
-parse_translation {* [(@{syntax_const "_hoare_abort_vars"}, K Hoare_Syntax.hoare_vars_tr)] *}
+parse_translation \<open>[(@{syntax_const "_hoare_abort_vars"}, K Hoare_Syntax.hoare_vars_tr)]\<close>
print_translation
- {* [(@{const_syntax Valid}, K (Hoare_Syntax.spec_tr' @{syntax_const "_hoare_abort"}))] *}
+ \<open>[(@{const_syntax Valid}, K (Hoare_Syntax.spec_tr' @{syntax_const "_hoare_abort"}))]\<close>
(*** The proof rules ***)
@@ -99,20 +99,20 @@
by(auto simp:Valid_def)
-subsection {* Derivation of the proof rules and, most importantly, the VCG tactic *}
+subsection \<open>Derivation of the proof rules and, most importantly, the VCG tactic\<close>
lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
by blast
ML_file "hoare_tac.ML"
-method_setup vcg = {*
- Scan.succeed (fn ctxt => SIMPLE_METHOD' (Hoare.hoare_tac ctxt (K all_tac))) *}
+method_setup vcg = \<open>
+ Scan.succeed (fn ctxt => SIMPLE_METHOD' (Hoare.hoare_tac ctxt (K all_tac)))\<close>
"verification condition generator"
-method_setup vcg_simp = {*
+method_setup vcg_simp = \<open>
Scan.succeed (fn ctxt =>
- SIMPLE_METHOD' (Hoare.hoare_tac ctxt (asm_full_simp_tac ctxt))) *}
+ SIMPLE_METHOD' (Hoare.hoare_tac ctxt (asm_full_simp_tac ctxt)))\<close>
"verification condition generator plus simplification"
(* Special syntax for guarded statements and guarded array updates: *)
@@ -125,7 +125,7 @@
"a[i] := v" => "(i < CONST length a) \<rightarrow> (a := CONST list_update a i v)"
(* reverse translation not possible because of duplicate "a" *)
-text{* Note: there is no special syntax for guarded array access. Thus
-you must write @{text"j < length a \<rightarrow> a[i] := a!j"}. *}
+text\<open>Note: there is no special syntax for guarded array access. Thus
+you must write \<open>j < length a \<rightarrow> a[i] := a!j\<close>.\<close>
end
--- a/src/HOL/Hoare/Pointer_Examples.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare/Pointer_Examples.thy Sat Jan 02 18:48:45 2016 +0100
@@ -39,8 +39,8 @@
apply fastforce
done
-text{* And now with ghost variables @{term ps} and @{term qs}. Even
-``more automatic''. *}
+text\<open>And now with ghost variables @{term ps} and @{term qs}. Even
+``more automatic''.\<close>
lemma "VARS next p ps q qs r
{List next p Ps \<and> List next q Qs \<and> set Ps \<inter> set Qs = {} \<and>
@@ -95,7 +95,7 @@
qed
-text{* Finaly, the functional version. A bit more verbose, but automatic! *}
+text\<open>Finaly, the functional version. A bit more verbose, but automatic!\<close>
lemma "VARS tl p q r
{islist tl p \<and> islist tl q \<and>
@@ -115,11 +115,11 @@
subsection "Searching in a list"
-text{*What follows is a sequence of successively more intelligent proofs that
+text\<open>What follows is a sequence of successively more intelligent proofs that
a simple loop finds an element in a linked list.
We start with a proof based on the @{term List} predicate. This means it only
-works for acyclic lists. *}
+works for acyclic lists.\<close>
lemma "VARS tl p
{List tl p Ps \<and> X \<in> set Ps}
@@ -133,8 +133,8 @@
apply clarsimp
done
-text{*Using @{term Path} instead of @{term List} generalizes the correctness
-statement to cyclic lists as well: *}
+text\<open>Using @{term Path} instead of @{term List} generalizes the correctness
+statement to cyclic lists as well:\<close>
lemma "VARS tl p
{Path tl p Ps X}
@@ -148,9 +148,9 @@
apply clarsimp
done
-text{*Now it dawns on us that we do not need the list witness at all --- it
+text\<open>Now it dawns on us that we do not need the list witness at all --- it
suffices to talk about reachability, i.e.\ we can use relations directly. The
-first version uses a relation on @{typ"'a ref"}: *}
+first version uses a relation on @{typ"'a ref"}:\<close>
lemma "VARS tl p
{(p,X) \<in> {(Ref x,tl x) |x. True}^*}
@@ -166,7 +166,7 @@
apply(fast elim:converse_rtranclE)
done
-text{*Finally, a version based on a relation on type @{typ 'a}:*}
+text\<open>Finally, a version based on a relation on type @{typ 'a}:\<close>
lemma "VARS tl p
{p \<noteq> Null \<and> (addr p,X) \<in> {(x,y). tl x = Ref y}^*}
@@ -230,7 +230,7 @@
| "merge([],y#ys,f) = y # merge([],ys,f)"
| "merge([],[],f) = []"
-text{* Simplifies the proof a little: *}
+text\<open>Simplifies the proof a little:\<close>
lemma [simp]: "({} = insert a A \<inter> B) = (a \<notin> B & {} = A \<inter> B)"
by blast
@@ -287,7 +287,7 @@
apply(clarsimp simp add:List_app)
done
-text{* And now with ghost variables: *}
+text\<open>And now with ghost variables:\<close>
lemma "VARS elem next p q r s ps qs rs a
{List next p Ps \<and> List next q Qs \<and> set Ps \<inter> set Qs = {} \<and>
@@ -324,18 +324,18 @@
apply(clarsimp simp add:List_app)
done
-text{* The proof is a LOT simpler because it does not need
+text\<open>The proof is a LOT simpler because it does not need
instantiations anymore, but it is still not quite automatic, probably
-because of this wrong orientation business. *}
+because of this wrong orientation business.\<close>
-text{* More of the previous proof without ghost variables can be
+text\<open>More of the previous proof without ghost variables can be
automated, but the runtime goes up drastically. In general it is
usually more efficient to give the witness directly than to have it
found by proof.
Now we try a functional version of the abstraction relation @{term
Path}. Since the result is not that convincing, we do not prove any of
-the lemmas.*}
+the lemmas.\<close>
axiomatization
ispath :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" and
@@ -404,8 +404,8 @@
subsection "Cyclic list reversal"
-text{* We consider two algorithms for the reversal of circular lists.
-*}
+text\<open>We consider two algorithms for the reversal of circular lists.
+\<close>
lemma circular_list_rev_I:
"VARS next root p q tmp
@@ -428,7 +428,7 @@
apply clarsimp
done
-text{* In the beginning, we are able to assert @{term"distPath next
+text\<open>In the beginning, we are able to assert @{term"distPath next
root as root"}, with @{term"as"} set to @{term"[]"} or
@{term"[r,a,b,c]"}. Note that @{term"Path next root as root"} would
additionally give us an infinite number of lists with the recurring
@@ -450,7 +450,7 @@
It may come as a surprise to the reader that the simple algorithm for
acyclic list reversal, with modified annotations, works for cyclic
-lists as well: *}
+lists as well:\<close>
lemma circular_list_rev_II:
--- a/src/HOL/Hoare/Pointers0.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare/Pointers0.thy Sat Jan 02 18:48:45 2016 +0100
@@ -241,7 +241,7 @@
qed
-text{* Finaly, the functional version. A bit more verbose, but automatic! *}
+text\<open>Finaly, the functional version. A bit more verbose, but automatic!\<close>
lemma "VARS tl p q r
{islist tl p \<and> islist tl q \<and>
@@ -261,11 +261,11 @@
subsection "Searching in a list"
-text{*What follows is a sequence of successively more intelligent proofs that
+text\<open>What follows is a sequence of successively more intelligent proofs that
a simple loop finds an element in a linked list.
We start with a proof based on the @{term List} predicate. This means it only
-works for acyclic lists. *}
+works for acyclic lists.\<close>
lemma "VARS tl p
{List tl p Ps \<and> X \<in> set Ps}
@@ -282,8 +282,8 @@
apply clarsimp
done
-text{*Using @{term Path} instead of @{term List} generalizes the correctness
-statement to cyclic lists as well: *}
+text\<open>Using @{term Path} instead of @{term List} generalizes the correctness
+statement to cyclic lists as well:\<close>
lemma "VARS tl p
{Path tl p Ps X}
@@ -297,8 +297,8 @@
apply clarsimp
done
-text{*Now it dawns on us that we do not need the list witness at all --- it
-suffices to talk about reachability, i.e.\ we can use relations directly. *}
+text\<open>Now it dawns on us that we do not need the list witness at all --- it
+suffices to talk about reachability, i.e.\ we can use relations directly.\<close>
lemma "VARS tl p
{(p,X) \<in> {(x,y). y = tl x & x \<noteq> Null}^*}
--- a/src/HOL/Hoare/SchorrWaite.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare/SchorrWaite.thy Sat Jan 02 18:48:45 2016 +0100
@@ -8,10 +8,10 @@
theory SchorrWaite imports HeapSyntax begin
-section {* Machinery for the Schorr-Waite proof*}
+section \<open>Machinery for the Schorr-Waite proof\<close>
definition
- -- "Relations induced by a mapping"
+ \<comment> "Relations induced by a mapping"
rel :: "('a \<Rightarrow> 'a ref) \<Rightarrow> ('a \<times> 'a) set"
where "rel m = {(x,y). m x = Ref y}"
@@ -29,7 +29,7 @@
lemmas rel_defs = relS_def rel_def
-text {* Rewrite rules for relations induced by a mapping*}
+text \<open>Rewrite rules for relations induced by a mapping\<close>
lemma self_reachable: "b \<in> B \<Longrightarrow> b \<in> R\<^sup>* `` B"
apply blast
@@ -83,11 +83,11 @@
done
definition
- -- "Restriction of a relation"
+ \<comment> "Restriction of a relation"
restr ::"('a \<times> 'a) set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('a \<times> 'a) set" ("(_/ | _)" [50, 51] 50)
where "restr r m = {(x,y). (x,y) \<in> r \<and> \<not> m x}"
-text {* Rewrite rules for the restriction of a relation *}
+text \<open>Rewrite rules for the restriction of a relation\<close>
lemma restr_identity[simp]:
" (\<forall>x. \<not> m x) \<Longrightarrow> (R |m) = R"
@@ -115,11 +115,11 @@
done
definition
- -- "A short form for the stack mapping function for List"
+ \<comment> "A short form for the stack mapping function for List"
S :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a ref) \<Rightarrow> ('a \<Rightarrow> 'a ref) \<Rightarrow> ('a \<Rightarrow> 'a ref)"
where "S c l r = (\<lambda>x. if c x then r x else l x)"
-text {* Rewrite rules for Lists using S as their mapping *}
+text \<open>Rewrite rules for Lists using S as their mapping\<close>
lemma [rule_format,simp]:
"\<forall>p. a \<notin> set stack \<longrightarrow> List (S c l r) p stack = List (S (c(a:=x)) (l(a:=y)) (r(a:=z))) p stack"
@@ -146,7 +146,7 @@
done
primrec
- --"Recursive definition of what is means for a the graph/stack structure to be reconstructible"
+ \<comment>"Recursive definition of what is means for a the graph/stack structure to be reconstructible"
stkOk :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a ref) \<Rightarrow> ('a \<Rightarrow> 'a ref) \<Rightarrow> ('a \<Rightarrow> 'a ref) \<Rightarrow> ('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow>'a list \<Rightarrow> bool"
where
stkOk_nil: "stkOk c l r iL iR t [] = True"
@@ -155,7 +155,7 @@
iL p = (if c p then l p else t) \<and>
iR p = (if c p then t else r p))"
-text {* Rewrite rules for stkOk *}
+text \<open>Rewrite rules for stkOk\<close>
lemma [simp]: "\<And>t. \<lbrakk> x \<notin> set xs; Ref x\<noteq>t \<rbrakk> \<Longrightarrow>
stkOk (c(x := f)) l r iL iR t xs = stkOk c l r iL iR t xs"
@@ -194,7 +194,7 @@
done
-section{*The Schorr-Waite algorithm*}
+section\<open>The Schorr-Waite algorithm\<close>
theorem SchorrWaiteAlgorithm:
@@ -276,21 +276,21 @@
have "?popInv stack_tl"
proof -
- -- {*List property is maintained:*}
+ \<comment> \<open>List property is maintained:\<close>
from i1 p_notin_stack_tl ifB2
have poI1: "List (S c l (r(p \<rightarrow> t))) (p^.r) stack_tl"
by(simp add: addr_p_eq stack_eq, simp add: S_def)
moreover
- -- {*Everything on the stack is marked:*}
+ \<comment> \<open>Everything on the stack is marked:\<close>
from i2 have poI2: "\<forall> x \<in> set stack_tl. m x" by (simp add:stack_eq)
moreover
- -- {*Everything is still reachable:*}
+ \<comment> \<open>Everything is still reachable:\<close>
let "(R = reachable ?Ra ?A)" = "?I3"
let "?Rb" = "(relS {l, r(p \<rightarrow> t)})"
let "?B" = "{p, p^.r}"
- -- {*Our goal is @{text"R = reachable ?Rb ?B"}.*}
+ \<comment> \<open>Our goal is \<open>R = reachable ?Rb ?B\<close>.\<close>
have "?Ra\<^sup>* `` addrs ?A = ?Rb\<^sup>* `` addrs ?B" (is "?L = ?R")
proof
show "?L \<subseteq> ?R"
@@ -314,11 +314,11 @@
with i3 have poI3: "R = reachable ?Rb ?B" by (simp add:reachable_def)
moreover
- -- "If it is reachable and not marked, it is still reachable using..."
+ \<comment> "If it is reachable and not marked, it is still reachable using..."
let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Ra ?A" = ?I4
let "?Rb" = "relS {l, r(p \<rightarrow> t)} | m"
let "?B" = "{p} \<union> set (map (r(p \<rightarrow> t)) stack_tl)"
- -- {*Our goal is @{text"\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Rb ?B"}.*}
+ \<comment> \<open>Our goal is \<open>\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Rb ?B\<close>.\<close>
let ?T = "{t, p^.r}"
have "?Ra\<^sup>* `` addrs ?A \<subseteq> ?Rb\<^sup>* `` (addrs ?B \<union> addrs ?T)"
@@ -331,37 +331,37 @@
by (clarsimp simp:restr_def relS_def)
(fastforce simp add:rel_def Image_iff addrs_def dest:rel_upd1)
qed
- -- "We now bring a term from the right to the left of the subset relation."
+ \<comment> "We now bring a term from the right to the left of the subset relation."
hence subset: "?Ra\<^sup>* `` addrs ?A - ?Rb\<^sup>* `` addrs ?T \<subseteq> ?Rb\<^sup>* `` addrs ?B"
by blast
have poI4: "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Rb ?B"
proof (rule allI, rule impI)
fix x
assume a: "x \<in> R \<and> \<not> m x"
- -- {*First, a disjunction on @{term"p^.r"} used later in the proof*}
+ \<comment> \<open>First, a disjunction on @{term"p^.r"} used later in the proof\<close>
have pDisj:"p^.r = Null \<or> (p^.r \<noteq> Null \<and> p^.r^.m)" using poI1 poI2
by auto
- -- {*@{term x} belongs to the left hand side of @{thm[source] subset}:*}
+ \<comment> \<open>@{term x} belongs to the left hand side of @{thm[source] subset}:\<close>
have incl: "x \<in> ?Ra\<^sup>*``addrs ?A" using a i4 by (simp only:reachable_def, clarsimp)
have excl: "x \<notin> ?Rb\<^sup>*`` addrs ?T" using pDisj ifB1 a by (auto simp add:addrs_def)
- -- {*And therefore also belongs to the right hand side of @{thm[source]subset},*}
- -- {*which corresponds to our goal.*}
+ \<comment> \<open>And therefore also belongs to the right hand side of @{thm[source]subset},\<close>
+ \<comment> \<open>which corresponds to our goal.\<close>
from incl excl subset show "x \<in> reachable ?Rb ?B" by (auto simp add:reachable_def)
qed
moreover
- -- "If it is marked, then it is reachable"
+ \<comment> "If it is marked, then it is reachable"
from i5 have poI5: "\<forall>x. m x \<longrightarrow> x \<in> R" .
moreover
- -- {*If it is not on the stack, then its @{term l} and @{term r} fields are unchanged*}
+ \<comment> \<open>If it is not on the stack, then its @{term l} and @{term r} fields are unchanged\<close>
from i7 i6 ifB2
have poI6: "\<forall>x. x \<notin> set stack_tl \<longrightarrow> (r(p \<rightarrow> t)) x = iR x \<and> l x = iL x"
by(auto simp: addr_p_eq stack_eq fun_upd_apply)
moreover
- -- {*If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed*}
+ \<comment> \<open>If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed\<close>
from p_notin_stack_tl i7 have poI7: "stkOk c l (r(p \<rightarrow> t)) iL iR p stack_tl"
by (clarsimp simp:stack_eq addr_p_eq)
@@ -371,11 +371,11 @@
}
moreover
- -- "Proofs of the Swing and Push arm follow."
- -- "Since they are in principle simmilar to the Pop arm proof,"
- -- "we show fewer comments and use frequent pattern matching."
+ \<comment> "Proofs of the Swing and Push arm follow."
+ \<comment> "Since they are in principle simmilar to the Pop arm proof,"
+ \<comment> "we show fewer comments and use frequent pattern matching."
{
- -- "Swing arm"
+ \<comment> "Swing arm"
assume ifB1: "?ifB1" and nifB2: "\<not>?ifB2"
from ifB1 whileB have pNotNull: "p \<noteq> Null" by clarsimp
then obtain addr_p where addr_p_eq: "p = Ref addr_p" by clarsimp
@@ -387,18 +387,18 @@
have "?swInv stack"
proof -
- -- {*List property is maintained:*}
+ \<comment> \<open>List property is maintained:\<close>
from i1 p_notin_stack_tl nifB2
have swI1: "?swI1"
by (simp add:addr_p_eq stack_eq, simp add:S_def)
moreover
- -- {*Everything on the stack is marked:*}
+ \<comment> \<open>Everything on the stack is marked:\<close>
from i2
have swI2: "?swI2" .
moreover
- -- {*Everything is still reachable:*}
+ \<comment> \<open>Everything is still reachable:\<close>
let "R = reachable ?Ra ?A" = "?I3"
let "R = reachable ?Rb ?B" = "?swI3"
have "?Ra\<^sup>* `` addrs ?A = ?Rb\<^sup>* `` addrs ?B"
@@ -419,7 +419,7 @@
have swI3: "?swI3" by (simp add:reachable_def)
moreover
- -- "If it is reachable and not marked, it is still reachable using..."
+ \<comment> "If it is reachable and not marked, it is still reachable using..."
let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Ra ?A" = ?I4
let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Rb ?B" = ?swI4
let ?T = "{t}"
@@ -449,18 +449,18 @@
qed
moreover
- -- "If it is marked, then it is reachable"
+ \<comment> "If it is marked, then it is reachable"
from i5
have "?swI5" .
moreover
- -- {*If it is not on the stack, then its @{term l} and @{term r} fields are unchanged*}
+ \<comment> \<open>If it is not on the stack, then its @{term l} and @{term r} fields are unchanged\<close>
from i6 stack_eq
have "?swI6"
by clarsimp
moreover
- -- {*If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed*}
+ \<comment> \<open>If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed\<close>
from stackDist i7 nifB2
have "?swI7"
by (clarsimp simp:addr_p_eq stack_eq)
@@ -472,7 +472,7 @@
moreover
{
- -- "Push arm"
+ \<comment> "Push arm"
assume nifB1: "\<not>?ifB1"
from nifB1 whileB have tNotNull: "t \<noteq> Null" by clarsimp
then obtain addr_t where addr_t_eq: "t = Ref addr_t" by clarsimp
@@ -483,19 +483,19 @@
have "?puInv new_stack"
proof -
- -- {*List property is maintained:*}
+ \<comment> \<open>List property is maintained:\<close>
from i1 t_notin_stack
have puI1: "?puI1"
by (simp add:addr_t_eq new_stack_eq, simp add:S_def)
moreover
- -- {*Everything on the stack is marked:*}
+ \<comment> \<open>Everything on the stack is marked:\<close>
from i2
have puI2: "?puI2"
by (simp add:new_stack_eq fun_upd_apply)
moreover
- -- {*Everything is still reachable:*}
+ \<comment> \<open>Everything is still reachable:\<close>
let "R = reachable ?Ra ?A" = "?I3"
let "R = reachable ?Rb ?B" = "?puI3"
have "?Ra\<^sup>* `` addrs ?A = ?Rb\<^sup>* `` addrs ?B"
@@ -516,7 +516,7 @@
have puI3: "?puI3" by (simp add:reachable_def)
moreover
- -- "If it is reachable and not marked, it is still reachable using..."
+ \<comment> "If it is reachable and not marked, it is still reachable using..."
let "\<forall>x. x \<in> R \<and> \<not> m x \<longrightarrow> x \<in> reachable ?Ra ?A" = ?I4
let "\<forall>x. x \<in> R \<and> \<not> ?new_m x \<longrightarrow> x \<in> reachable ?Rb ?B" = ?puI4
let ?T = "{t}"
@@ -546,19 +546,19 @@
qed
moreover
- -- "If it is marked, then it is reachable"
+ \<comment> "If it is marked, then it is reachable"
from i5
have "?puI5"
by (auto simp:addrs_def i3 reachable_def addr_t_eq fun_upd_apply intro:self_reachable)
moreover
- -- {*If it is not on the stack, then its @{term l} and @{term r} fields are unchanged*}
+ \<comment> \<open>If it is not on the stack, then its @{term l} and @{term r} fields are unchanged\<close>
from i6
have "?puI6"
by (simp add:new_stack_eq)
moreover
- -- {*If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed*}
+ \<comment> \<open>If it is on the stack, then its @{term l} and @{term r} fields can be reconstructed\<close>
from stackDist i6 t_notin_stack i7
have "?puI7" by (clarsimp simp:addr_t_eq new_stack_eq)
--- a/src/HOL/Hoare/SepLogHeap.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare/SepLogHeap.thy Sat Jan 02 18:48:45 2016 +0100
@@ -12,8 +12,8 @@
type_synonym heap = "(nat \<Rightarrow> nat option)"
-text{* @{text "Some"} means allocated, @{text "None"} means
-free. Address @{text "0"} serves as the null reference. *}
+text\<open>\<open>Some\<close> means allocated, \<open>None\<close> means
+free. Address \<open>0\<close> serves as the null reference.\<close>
subsection "Paths in the heap"
--- a/src/HOL/Hoare/Separation.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare/Separation.thy Sat Jan 02 18:48:45 2016 +0100
@@ -14,7 +14,7 @@
theory Separation imports Hoare_Logic_Abort SepLogHeap begin
-text{* The semantic definition of a few connectives: *}
+text\<open>The semantic definition of a few connectives:\<close>
definition ortho :: "heap \<Rightarrow> heap \<Rightarrow> bool" (infix "\<bottom>" 55)
where "h1 \<bottom> h2 \<longleftrightarrow> dom h1 \<inter> dom h2 = {}"
@@ -31,7 +31,7 @@
definition wand:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)"
where "wand P Q = (\<lambda>h. \<forall>h'. h' \<bottom> h \<and> P h' \<longrightarrow> Q(h++h'))"
-text{*This is what assertions look like without any syntactic sugar: *}
+text\<open>This is what assertions look like without any syntactic sugar:\<close>
lemma "VARS x y z w h
{star (%h. singl h x y) (%h. singl h z w) h}
@@ -41,11 +41,11 @@
apply(auto simp:star_def ortho_def singl_def)
done
-text{* Now we add nice input syntax. To suppress the heap parameter
+text\<open>Now we add nice input syntax. To suppress the heap parameter
of the connectives, we assume it is always called H and add/remove it
upon parsing/printing. Thus every pointer program needs to have a
program variable H, and assertions should not contain any locally
-bound Hs - otherwise they may bind the implicit H. *}
+bound Hs - otherwise they may bind the implicit H.\<close>
syntax
"_emp" :: "bool" ("emp")
@@ -54,7 +54,7 @@
"_wand" :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "-*" 60)
(* FIXME does not handle "_idtdummy" *)
-ML{*
+ML\<open>
(* free_tr takes care of free vars in the scope of sep. logic connectives:
they are implicitly applied to the heap *)
fun free_tr(t as Free _) = t $ Syntax.free "H"
@@ -74,16 +74,16 @@
fun wand_tr [P, Q] = Syntax.const @{const_syntax wand} $
absfree ("H", dummyT) P $ absfree ("H", dummyT) Q $ Syntax.free "H"
| wand_tr ts = raise TERM ("wand_tr", ts);
-*}
+\<close>
-parse_translation {*
+parse_translation \<open>
[(@{syntax_const "_emp"}, K emp_tr),
(@{syntax_const "_singl"}, K singl_tr),
(@{syntax_const "_star"}, K star_tr),
(@{syntax_const "_wand"}, K wand_tr)]
-*}
+\<close>
-text{* Now it looks much better: *}
+text\<open>Now it looks much better:\<close>
lemma "VARS H x y z w
{[x\<mapsto>y] ** [z\<mapsto>w]}
@@ -101,10 +101,10 @@
apply(auto simp:star_def ortho_def is_empty_def)
done
-text{* But the output is still unreadable. Thus we also strip the heap
-parameters upon output: *}
+text\<open>But the output is still unreadable. Thus we also strip the heap
+parameters upon output:\<close>
-ML {*
+ML \<open>
local
fun strip (Abs(_,_,(t as Const("_free",_) $ Free _) $ Bound 0)) = t
@@ -125,16 +125,16 @@
fun wand_tr' [P,Q,_] = Syntax.const @{syntax_const "_wand"} $ strip P $ strip Q
end
-*}
+\<close>
-print_translation {*
+print_translation \<open>
[(@{const_syntax is_empty}, K is_empty_tr'),
(@{const_syntax singl}, K singl_tr'),
(@{const_syntax star}, K star_tr'),
(@{const_syntax wand}, K wand_tr')]
-*}
+\<close>
-text{* Now the intermediate proof states are also readable: *}
+text\<open>Now the intermediate proof states are also readable:\<close>
lemma "VARS H x y z w
{[x\<mapsto>y] ** [z\<mapsto>w]}
@@ -152,9 +152,9 @@
apply(auto simp:star_def ortho_def is_empty_def)
done
-text{* So far we have unfolded the separation logic connectives in
+text\<open>So far we have unfolded the separation logic connectives in
proofs. Here comes a simple example of a program proof that uses a law
-of separation logic instead. *}
+of separation logic instead.\<close>
(* a law of separation logic *)
lemma star_comm: "P ** Q = Q ** P"
--- a/src/HOL/Hoare_Parallel/Gar_Coll.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/Gar_Coll.thy Sat Jan 02 18:48:45 2016 +0100
@@ -18,17 +18,17 @@
subsection \<open>The Mutator\<close>
-text \<open>The mutator first redirects an arbitrary edge @{text "R"} from
+text \<open>The mutator first redirects an arbitrary edge \<open>R\<close> from
an arbitrary accessible node towards an arbitrary accessible node
-@{text "T"}. It then colors the new target @{text "T"} black.
+\<open>T\<close>. It then colors the new target \<open>T\<close> black.
We declare the arbitrarily selected node and edge as constants:\<close>
consts R :: nat T :: nat
text \<open>\noindent The following predicate states, given a list of
-nodes @{text "m"} and a list of edges @{text "e"}, the conditions
-under which the selected edge @{text "R"} and node @{text "T"} are
+nodes \<open>m\<close> and a list of edges \<open>e\<close>, the conditions
+under which the selected edge \<open>R\<close> and node \<open>T\<close> are
valid:\<close>
definition Mut_init :: "gar_coll_state \<Rightarrow> bool" where
@@ -36,7 +36,7 @@
text \<open>\noindent For the mutator we
consider two modules, one for each action. An auxiliary variable
-@{text "\<acute>z"} is set to false if the mutator has already redirected an
+\<open>\<acute>z\<close> is set to false if the mutator has already redirected an
edge but has not yet colored the new target.\<close>
definition Redirect_Edge :: "gar_coll_state ann_com" where
@@ -80,9 +80,8 @@
subsection \<open>The Collector\<close>
-text \<open>\noindent A constant @{text "M_init"} is used to give @{text "\<acute>Ma"} a
-suitable first value, defined as a list of nodes where only the @{text
-"Roots"} are black.\<close>
+text \<open>\noindent A constant \<open>M_init\<close> is used to give \<open>\<acute>Ma\<close> a
+suitable first value, defined as a list of nodes where only the \<open>Roots\<close> are black.\<close>
consts M_init :: nodes
@@ -163,7 +162,7 @@
apply force
apply force
apply force
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply clarify
apply(simp add:Proper_Edges_def Proper_Roots_def Graph6 Graph7 Graph8 Graph12)
apply (erule disjE)
@@ -189,10 +188,10 @@
apply(erule subset_psubset_trans)
apply(erule Graph11)
apply fast
---\<open>3 subgoals left\<close>
+\<comment>\<open>3 subgoals left\<close>
apply force
apply force
---\<open>last\<close>
+\<comment>\<open>last\<close>
apply clarify
apply simp
apply(subgoal_tac "ind x = length (E x)")
@@ -247,10 +246,10 @@
apply force
apply force
apply force
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
apply clarify
apply(simp add:BtoW_def Proper_Edges_def)
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply clarify
apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12)
apply (erule disjE)
@@ -287,7 +286,7 @@
apply(erule subset_psubset_trans)
apply(erule Graph11)
apply fast
---\<open>2 subgoals left\<close>
+\<comment>\<open>2 subgoals left\<close>
apply clarify
apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12)
apply (erule disjE)
@@ -304,7 +303,7 @@
apply arith
apply (simp add: BtoW_def)
apply (simp add: BtoW_def)
---\<open>last\<close>
+\<comment>\<open>last\<close>
apply clarify
apply simp
apply(subgoal_tac "ind x = length (E x)")
@@ -521,7 +520,7 @@
"interfree_aux (Some Propagate_Black, {}, Some Redirect_Edge)"
apply (unfold modules )
apply interfree_aux
---\<open>11 subgoals left\<close>
+\<comment>\<open>11 subgoals left\<close>
apply(clarify, simp add:abbrev Graph6 Graph12)
apply(clarify, simp add:abbrev Graph6 Graph12)
apply(clarify, simp add:abbrev Graph6 Graph12)
@@ -536,7 +535,7 @@
apply (force simp add:BtoW_def)
apply(erule Graph4)
apply simp+
---\<open>7 subgoals left\<close>
+\<comment>\<open>7 subgoals left\<close>
apply(clarify, simp add:abbrev Graph6 Graph12)
apply(erule conjE)+
apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym)
@@ -548,7 +547,7 @@
apply (force simp add:BtoW_def)
apply(erule Graph4)
apply simp+
---\<open>6 subgoals left\<close>
+\<comment>\<open>6 subgoals left\<close>
apply(clarify, simp add:abbrev Graph6 Graph12)
apply(erule conjE)+
apply(rule conjI)
@@ -563,9 +562,9 @@
apply simp+
apply(simp add:BtoW_def nth_list_update)
apply force
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
apply(clarify, simp add:abbrev Graph6 Graph12)
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply(clarify, simp add:abbrev Graph6 Graph12)
apply(rule conjI)
apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym)
@@ -589,9 +588,9 @@
apply simp+
apply(force simp add:BtoW_def)
apply(force simp add:BtoW_def)
---\<open>3 subgoals left\<close>
+\<comment>\<open>3 subgoals left\<close>
apply(clarify, simp add:abbrev Graph6 Graph12)
---\<open>2 subgoals left\<close>
+\<comment>\<open>2 subgoals left\<close>
apply(clarify, simp add:abbrev Graph6 Graph12)
apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym)
apply clarify
@@ -616,21 +615,21 @@
"interfree_aux (Some Propagate_Black, {}, Some Color_Target)"
apply (unfold modules )
apply interfree_aux
---\<open>11 subgoals left\<close>
+\<comment>\<open>11 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)+
apply(erule conjE)+
apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9,
case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify,
erule allE, erule impE, assumption, erule impE, assumption,
simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force)
---\<open>7 subgoals left\<close>
+\<comment>\<open>7 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
apply(erule conjE)+
apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9,
case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify,
erule allE, erule impE, assumption, erule impE, assumption,
simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force)
---\<open>6 subgoals left\<close>
+\<comment>\<open>6 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
apply clarify
apply (rule conjI)
@@ -639,9 +638,9 @@
erule allE, erule impE, assumption, erule impE, assumption,
simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force)
apply(simp add:nth_list_update)
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
apply (rule conjI)
apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9,
@@ -652,15 +651,15 @@
apply(simp add:nth_list_update)
apply(rule impI,rule impI, case_tac "M x!T=Black",rotate_tac -1, force simp add: BtoW_def Graph10,
erule subset_psubset_trans, erule Graph11, force)
---\<open>3 subgoals left\<close>
+\<comment>\<open>3 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
---\<open>2 subgoals left\<close>
+\<comment>\<open>2 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9,
case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify,
erule allE, erule impE, assumption, erule impE, assumption,
simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force)
---\<open>3 subgoals left\<close>
+\<comment>\<open>3 subgoals left\<close>
apply(simp add:abbrev)
done
@@ -675,9 +674,9 @@
"interfree_aux (Some Count, {}, Some Redirect_Edge)"
apply (unfold modules)
apply interfree_aux
---\<open>9 subgoals left\<close>
+\<comment>\<open>9 subgoals left\<close>
apply(simp_all add:abbrev Graph6 Graph12)
---\<open>6 subgoals left\<close>
+\<comment>\<open>6 subgoals left\<close>
apply(clarify, simp add:abbrev Graph6 Graph12,
erule disjE,erule disjI1,rule disjI2,rule subset_trans, erule Graph3,force,force)+
done
@@ -694,17 +693,17 @@
"interfree_aux (Some Count, {}, Some Color_Target)"
apply (unfold modules )
apply interfree_aux
---\<open>9 subgoals left\<close>
+\<comment>\<open>9 subgoals left\<close>
apply(simp_all add:abbrev Graph7 Graph8 Graph12)
---\<open>6 subgoals left\<close>
+\<comment>\<open>6 subgoals left\<close>
apply(clarify,simp add:abbrev Graph7 Graph8 Graph12,
erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9)+
---\<open>2 subgoals left\<close>
+\<comment>\<open>2 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
apply(rule conjI)
apply(erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9)
apply(simp add:nth_list_update)
---\<open>1 subgoal left\<close>
+\<comment>\<open>1 subgoal left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12,
erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9)
done
@@ -770,9 +769,9 @@
apply(simp_all add:collector_mutator_interfree)
apply(unfold modules collector_defs Mut_init_def)
apply(tactic \<open>TRYALL (interfree_aux_tac @{context})\<close>)
---\<open>32 subgoals left\<close>
+\<comment>\<open>32 subgoals left\<close>
apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
---\<open>20 subgoals left\<close>
+\<comment>\<open>20 subgoals left\<close>
apply(tactic\<open>TRYALL (clarify_tac @{context})\<close>)
apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
apply(tactic \<open>TRYALL (eresolve_tac @{context} [disjE])\<close>)
@@ -801,10 +800,10 @@
apply(simp_all add:collector_mutator_interfree)
apply(unfold modules collector_defs Mut_init_def)
apply(tactic \<open>TRYALL (interfree_aux_tac @{context})\<close>)
---\<open>64 subgoals left\<close>
+\<comment>\<open>64 subgoals left\<close>
apply(simp_all add:nth_list_update Invariants Append_to_free0)+
apply(tactic\<open>TRYALL (clarify_tac @{context})\<close>)
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply force
apply(simp add:Append_to_free2)
apply force
--- a/src/HOL/Hoare_Parallel/Graph.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/Graph.thy Sat Jan 02 18:48:45 2016 +0100
@@ -191,12 +191,12 @@
apply clarify
apply simp
apply(case_tac "\<exists>i<length path - 1. (fst(E!R),T)=(path!(Suc i),path!i)")
---\<open>the changed edge is part of the path\<close>
+\<comment>\<open>the changed edge is part of the path\<close>
apply(erule exE)
apply(drule_tac P = "\<lambda>i. i<length path - 1 \<and> (fst(E!R),T)=(path!Suc i,path!i)" in Ex_first_occurrence)
apply clarify
apply(erule disjE)
---\<open>T is NOT a root\<close>
+\<comment>\<open>T is NOT a root\<close>
apply clarify
apply(rule_tac x = "(take m path)@patha" in exI)
apply(subgoal_tac "\<not>(length path\<le>m)")
@@ -240,7 +240,7 @@
apply(subgoal_tac "Suc (i - m)=(Suc i - m)" )
prefer 2 apply arith
apply simp
---\<open>T is a root\<close>
+\<comment>\<open>T is a root\<close>
apply(case_tac "m=0")
apply force
apply(rule_tac x = "take (Suc m) path" in exI)
@@ -253,7 +253,7 @@
apply(case_tac "R=j")
apply(force simp add: nth_list_update)
apply(force simp add: nth_list_update)
---\<open>the changed edge is not part of the path\<close>
+\<comment>\<open>the changed edge is not part of the path\<close>
apply(rule_tac x = "path" in exI)
apply simp
apply clarify
@@ -276,7 +276,7 @@
apply(erule disjE)
prefer 2 apply force
apply clarify
---\<open>there exist a black node in the path to T\<close>
+\<comment>\<open>there exist a black node in the path to T\<close>
apply(case_tac "\<exists>m<length path. M!(path!m)=Black")
apply(erule exE)
apply(drule_tac P = "\<lambda>m. m<length path \<and> M!(path!m)=Black" in Ex_first_occurrence)
@@ -318,7 +318,7 @@
apply(erule disjE)
prefer 2 apply force
apply clarify
---\<open>there exist a black node in the path to T\<close>
+\<comment>\<open>there exist a black node in the path to T\<close>
apply(case_tac "\<exists>m<length path. M!(path!m)=Black")
apply(erule exE)
apply(drule_tac P = "\<lambda>m. m<length path \<and> M!(path!m)=Black" in Ex_first_occurrence)
--- a/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Sat Jan 02 18:48:45 2016 +0100
@@ -249,12 +249,12 @@
apply(unfold Mul_Propagate_Black_def)
apply annhoare
apply(simp_all add:Mul_PBInv_def mul_collector_defs Mul_Auxk_def Graph6 Graph7 Graph8 Graph12 mul_collector_defs Queue_def)
---\<open>8 subgoals left\<close>
+\<comment>\<open>8 subgoals left\<close>
apply force
apply force
apply force
apply(force simp add:BtoW_def Graph_defs)
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply clarify
apply(simp add: mul_collector_defs Graph12 Graph6 Graph7 Graph8)
apply(disjE_tac)
@@ -269,7 +269,7 @@
apply(force)
apply(force)
apply(rule disjI2, rule disjI1, erule subset_psubset_trans, erule Graph11, force)
---\<open>2 subgoals left\<close>
+\<comment>\<open>2 subgoals left\<close>
apply clarify
apply(conjI_tac)
apply(disjE_tac)
@@ -278,7 +278,7 @@
apply(erule less_SucE)
apply force
apply (simp add:BtoW_def)
---\<open>1 subgoal left\<close>
+\<comment>\<open>1 subgoal left\<close>
apply clarify
apply simp
apply(disjE_tac)
@@ -342,11 +342,11 @@
apply (unfold Mul_Count_def)
apply annhoare
apply(simp_all add:Mul_CountInv_def mul_collector_defs Mul_Auxk_def Graph6 Graph7 Graph8 Graph12 mul_collector_defs Queue_def)
---\<open>7 subgoals left\<close>
+\<comment>\<open>7 subgoals left\<close>
apply force
apply force
apply force
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply clarify
apply(conjI_tac)
apply(disjE_tac)
@@ -357,7 +357,7 @@
back
apply force
apply force
---\<open>3 subgoals left\<close>
+\<comment>\<open>3 subgoals left\<close>
apply clarify
apply(conjI_tac)
apply(disjE_tac)
@@ -369,9 +369,9 @@
apply simp
apply(rotate_tac -1)
apply (force simp add:Blacks_def)
---\<open>2 subgoals left\<close>
+\<comment>\<open>2 subgoals left\<close>
apply force
---\<open>1 subgoal left\<close>
+\<comment>\<open>1 subgoal left\<close>
apply clarify
apply(drule_tac x = "ind x" in le_imp_less_or_eq)
apply (simp_all add:Blacks_def)
@@ -566,7 +566,7 @@
apply (unfold mul_modules)
apply interfree_aux
apply(simp_all add:mul_mutator_defs mul_collector_defs Mul_PBInv_def nth_list_update Graph6)
---\<open>7 subgoals left\<close>
+\<comment>\<open>7 subgoals left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -574,7 +574,7 @@
apply(rule conjI)
apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
---\<open>6 subgoals left\<close>
+\<comment>\<open>6 subgoals left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -582,7 +582,7 @@
apply(rule conjI)
apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -606,7 +606,7 @@
apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update)
apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans)
apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update)
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -630,7 +630,7 @@
apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update)
apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans)
apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update)
---\<open>3 subgoals left\<close>
+\<comment>\<open>3 subgoals left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -686,7 +686,7 @@
apply (force simp add: nth_list_update)
apply(rule impI, (rule disjI2)+, erule le_trans)
apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
---\<open>2 subgoals left\<close>
+\<comment>\<open>2 subgoals left\<close>
apply clarify
apply(rule conjI)
apply(disjE_tac)
@@ -756,7 +756,7 @@
apply(rule disjI1, erule less_le_trans)
apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply force
---\<open>1 subgoal left\<close>
+\<comment>\<open>1 subgoal left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -795,7 +795,7 @@
apply (unfold mul_modules)
apply interfree_aux
apply(simp_all add: mul_collector_defs mul_mutator_defs)
---\<open>7 subgoals left\<close>
+\<comment>\<open>7 subgoals left\<close>
apply clarify
apply (simp add:Graph7 Graph8 Graph12)
apply(disjE_tac)
@@ -805,7 +805,7 @@
apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10)
apply((rule disjI2)+,erule subset_psubset_trans, erule Graph11, simp)
apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9)
---\<open>6 subgoals left\<close>
+\<comment>\<open>6 subgoals left\<close>
apply clarify
apply (simp add:Graph7 Graph8 Graph12)
apply(disjE_tac)
@@ -815,7 +815,7 @@
apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10)
apply((rule disjI2)+,erule subset_psubset_trans, erule Graph11, simp)
apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9)
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
apply clarify
apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12)
apply(disjE_tac)
@@ -833,7 +833,7 @@
apply(erule le_trans)
apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10)
apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp)
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply clarify
apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12)
apply(disjE_tac)
@@ -851,7 +851,7 @@
apply(erule le_trans)
apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10)
apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp)
---\<open>3 subgoals left\<close>
+\<comment>\<open>3 subgoals left\<close>
apply clarify
apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12)
apply(case_tac "M x!(T (Muts x!j))=Black")
@@ -866,7 +866,7 @@
apply(rule conjI)
apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11)
apply (force simp add:nth_list_update)
---\<open>2 subgoals left\<close>
+\<comment>\<open>2 subgoals left\<close>
apply clarify
apply(simp add:Mul_Auxk_def Graph7 Graph8 Graph12)
apply(case_tac "M x!(T (Muts x!j))=Black")
@@ -887,7 +887,7 @@
apply(rule conjI)
apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11)
apply (force simp add:nth_list_update)
---\<open>1 subgoal left\<close>
+\<comment>\<open>1 subgoal left\<close>
apply clarify
apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12)
apply(case_tac "M x!(T (Muts x!j))=Black")
@@ -914,7 +914,7 @@
interfree_aux (Some(Mul_Count n ),{},Some(Mul_Redirect_Edge j n))"
apply (unfold mul_modules)
apply interfree_aux
---\<open>9 subgoals left\<close>
+\<comment>\<open>9 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def Graph6)
apply clarify
apply disjE_tac
@@ -928,9 +928,9 @@
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(simp add:Graph6)
---\<open>8 subgoals left\<close>
+\<comment>\<open>8 subgoals left\<close>
apply(simp add:mul_mutator_defs nth_list_update)
---\<open>7 subgoals left\<close>
+\<comment>\<open>7 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs)
apply clarify
apply disjE_tac
@@ -944,7 +944,7 @@
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(simp add:Graph6)
---\<open>6 subgoals left\<close>
+\<comment>\<open>6 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def)
apply clarify
apply disjE_tac
@@ -958,7 +958,7 @@
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(simp add:Graph6)
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def)
apply clarify
apply disjE_tac
@@ -972,7 +972,7 @@
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(simp add:Graph6)
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def)
apply clarify
apply disjE_tac
@@ -986,9 +986,9 @@
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(simp add:Graph6)
---\<open>3 subgoals left\<close>
+\<comment>\<open>3 subgoals left\<close>
apply(simp add:mul_mutator_defs nth_list_update)
---\<open>2 subgoals left\<close>
+\<comment>\<open>2 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def)
apply clarify
apply disjE_tac
@@ -1002,7 +1002,7 @@
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(simp add:Graph6)
---\<open>1 subgoal left\<close>
+\<comment>\<open>1 subgoal left\<close>
apply(simp add:mul_mutator_defs nth_list_update)
done
@@ -1019,7 +1019,7 @@
apply (unfold mul_modules)
apply interfree_aux
apply(simp_all add:mul_collector_defs mul_mutator_defs Mul_CountInv_def)
---\<open>6 subgoals left\<close>
+\<comment>\<open>6 subgoals left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12)
@@ -1033,7 +1033,7 @@
apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11)
apply (simp add: Graph7 Graph8 Graph12)
apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9)
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12)
@@ -1047,7 +1047,7 @@
apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11)
apply (simp add: Graph7 Graph8 Graph12)
apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9)
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12)
@@ -1061,7 +1061,7 @@
apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11)
apply (simp add: Graph7 Graph8 Graph12)
apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9)
---\<open>3 subgoals left\<close>
+\<comment>\<open>3 subgoals left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12)
@@ -1075,7 +1075,7 @@
apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11)
apply (simp add: Graph7 Graph8 Graph12)
apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9)
---\<open>2 subgoals left\<close>
+\<comment>\<open>2 subgoals left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12 nth_list_update)
@@ -1093,7 +1093,7 @@
apply(rule conjI)
apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9)
apply (simp add: nth_list_update)
---\<open>1 subgoal left\<close>
+\<comment>\<open>1 subgoal left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12)
@@ -1171,11 +1171,11 @@
apply(simp_all add:mul_collector_mutator_interfree)
apply(unfold mul_modules mul_collector_defs mul_mutator_defs)
apply(tactic \<open>TRYALL (interfree_aux_tac @{context})\<close>)
---\<open>42 subgoals left\<close>
+\<comment>\<open>42 subgoals left\<close>
apply (clarify,simp add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)+
---\<open>24 subgoals left\<close>
+\<comment>\<open>24 subgoals left\<close>
apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
---\<open>14 subgoals left\<close>
+\<comment>\<open>14 subgoals left\<close>
apply(tactic \<open>TRYALL (clarify_tac @{context})\<close>)
apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
apply(tactic \<open>TRYALL (resolve_tac @{context} [conjI])\<close>)
@@ -1184,57 +1184,57 @@
apply(tactic \<open>TRYALL (eresolve_tac @{context} [conjE])\<close>)
apply(tactic \<open>TRYALL (eresolve_tac @{context} [disjE])\<close>)
apply(tactic \<open>TRYALL (eresolve_tac @{context} [disjE])\<close>)
---\<open>72 subgoals left\<close>
+\<comment>\<open>72 subgoals left\<close>
apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
---\<open>35 subgoals left\<close>
+\<comment>\<open>35 subgoals left\<close>
apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI1],
resolve_tac @{context} [subset_trans],
eresolve_tac @{context} @{thms Graph3},
force_tac @{context},
assume_tac @{context}])\<close>)
---\<open>28 subgoals left\<close>
+\<comment>\<open>28 subgoals left\<close>
apply(tactic \<open>TRYALL (eresolve_tac @{context} [conjE])\<close>)
apply(tactic \<open>TRYALL (eresolve_tac @{context} [disjE])\<close>)
---\<open>34 subgoals left\<close>
+\<comment>\<open>34 subgoals left\<close>
apply(rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
apply(case_tac [!] "M x!(T (Muts x ! j))=Black")
apply(simp_all add:Graph10)
---\<open>47 subgoals left\<close>
+\<comment>\<open>47 subgoals left\<close>
apply(tactic \<open>TRYALL(EVERY'[REPEAT o resolve_tac @{context} [disjI2],
eresolve_tac @{context} @{thms subset_psubset_trans},
eresolve_tac @{context} @{thms Graph11},
force_tac @{context}])\<close>)
---\<open>41 subgoals left\<close>
+\<comment>\<open>41 subgoals left\<close>
apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI2],
resolve_tac @{context} [disjI1],
eresolve_tac @{context} @{thms le_trans},
force_tac (@{context} addsimps @{thms Queue_def less_Suc_eq_le le_length_filter_update})])\<close>)
---\<open>35 subgoals left\<close>
+\<comment>\<open>35 subgoals left\<close>
apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI2],
resolve_tac @{context} [disjI1],
eresolve_tac @{context} @{thms psubset_subset_trans},
resolve_tac @{context} @{thms Graph9},
force_tac @{context}])\<close>)
---\<open>31 subgoals left\<close>
+\<comment>\<open>31 subgoals left\<close>
apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI2],
resolve_tac @{context} [disjI1],
eresolve_tac @{context} @{thms subset_psubset_trans},
eresolve_tac @{context} @{thms Graph11},
force_tac @{context}])\<close>)
---\<open>29 subgoals left\<close>
+\<comment>\<open>29 subgoals left\<close>
apply(tactic \<open>TRYALL(EVERY'[REPEAT o resolve_tac @{context} [disjI2],
eresolve_tac @{context} @{thms subset_psubset_trans},
eresolve_tac @{context} @{thms subset_psubset_trans},
eresolve_tac @{context} @{thms Graph11},
force_tac @{context}])\<close>)
---\<open>25 subgoals left\<close>
+\<comment>\<open>25 subgoals left\<close>
apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI2],
resolve_tac @{context} [disjI2],
resolve_tac @{context} [disjI1],
eresolve_tac @{context} @{thms le_trans},
force_tac (@{context} addsimps @{thms Queue_def less_Suc_eq_le le_length_filter_update})])\<close>)
---\<open>10 subgoals left\<close>
+\<comment>\<open>10 subgoals left\<close>
apply(rule disjI2,rule disjI2,rule conjI,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update, rule disjI1, rule less_imp_le, erule less_le_trans, force simp add:Queue_def less_Suc_eq_le le_length_filter_update)+
done
@@ -1247,9 +1247,9 @@
apply(simp_all add:mul_collector_mutator_interfree)
apply(unfold mul_modules mul_collector_defs mul_mutator_defs)
apply(tactic \<open>TRYALL (interfree_aux_tac @{context})\<close>)
---\<open>76 subgoals left\<close>
+\<comment>\<open>76 subgoals left\<close>
apply (clarsimp simp add: nth_list_update)+
---\<open>56 subgoals left\<close>
+\<comment>\<open>56 subgoals left\<close>
apply (clarsimp simp add: Mul_AppendInv_def Append_to_free0 nth_list_update)+
done
@@ -1269,7 +1269,7 @@
COEND
\<lbrace>False\<rbrace>"
apply oghoare
---\<open>Strengthening the precondition\<close>
+\<comment>\<open>Strengthening the precondition\<close>
apply(rule Int_greatest)
apply (case_tac n)
apply(force simp add: Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append)
@@ -1279,15 +1279,15 @@
apply(case_tac i)
apply(simp add:Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append)
apply(simp add: Mul_Mutator_def mul_mutator_defs mul_collector_defs nth_append nth_map_upt)
---\<open>Collector\<close>
+\<comment>\<open>Collector\<close>
apply(rule Mul_Collector)
---\<open>Mutator\<close>
+\<comment>\<open>Mutator\<close>
apply(erule Mul_Mutator)
---\<open>Interference freedom\<close>
+\<comment>\<open>Interference freedom\<close>
apply(simp add:Mul_interfree_Collector_Mutator)
apply(simp add:Mul_interfree_Mutator_Collector)
apply(simp add:Mul_interfree_Mutator_Mutator)
---\<open>Weakening of the postcondition\<close>
+\<comment>\<open>Weakening of the postcondition\<close>
apply(case_tac n)
apply(simp add:Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append)
apply(simp add:Mul_Mutator_def mul_mutator_defs mul_collector_defs nth_append)
--- a/src/HOL/Hoare_Parallel/OG_Com.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Com.thy Sat Jan 02 18:48:45 2016 +0100
@@ -10,8 +10,8 @@
type_synonym 'a assn = "'a set"
text \<open>The syntax of commands is defined by two mutually recursive
-datatypes: @{text "'a ann_com"} for annotated commands and @{text "'a
-com"} for non-annotated commands.\<close>
+datatypes: \<open>'a ann_com\<close> for annotated commands and \<open>'a
+com\<close> for non-annotated commands.\<close>
datatype 'a ann_com =
AnnBasic "('a assn)" "('a \<Rightarrow> 'a)"
@@ -27,7 +27,7 @@
| Cond "('a bexp)" "('a com)" "('a com)"
| While "('a bexp)" "('a assn)" "('a com)"
-text \<open>The function @{text pre} extracts the precondition of an
+text \<open>The function \<open>pre\<close> extracts the precondition of an
annotated command:\<close>
primrec pre ::"'a ann_com \<Rightarrow> 'a assn" where
--- a/src/HOL/Hoare_Parallel/OG_Examples.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Examples.thy Sat Jan 02 18:48:45 2016 +0100
@@ -41,7 +41,7 @@
COEND
\<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1 \<and> \<acute>pr2=0 \<and> \<not>\<acute>in2\<rbrace>"
apply oghoare
---\<open>104 verification conditions.\<close>
+\<comment>\<open>104 verification conditions.\<close>
apply auto
done
@@ -89,7 +89,7 @@
COEND
\<lbrace>False\<rbrace>"
apply oghoare
---\<open>122 vc\<close>
+\<comment>\<open>122 vc\<close>
apply auto
done
@@ -116,7 +116,7 @@
COEND
\<lbrace>False\<rbrace>"
apply oghoare
---\<open>38 vc\<close>
+\<comment>\<open>38 vc\<close>
apply auto
done
@@ -135,7 +135,7 @@
COEND
\<lbrace>False\<rbrace>"
apply oghoare
---\<open>20 vc\<close>
+\<comment>\<open>20 vc\<close>
apply auto
done
@@ -167,40 +167,40 @@
COEND
\<lbrace>False\<rbrace>"
apply oghoare
---\<open>35 vc\<close>
+\<comment>\<open>35 vc\<close>
apply simp_all
---\<open>16 vc\<close>
+\<comment>\<open>16 vc\<close>
apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
---\<open>11 vc\<close>
+\<comment>\<open>11 vc\<close>
apply simp_all
apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
---\<open>10 subgoals left\<close>
+\<comment>\<open>10 subgoals left\<close>
apply(erule less_SucE)
apply simp
apply simp
---\<open>9 subgoals left\<close>
+\<comment>\<open>9 subgoals left\<close>
apply(case_tac "i=k")
apply force
apply simp
apply(case_tac "i=l")
apply force
apply force
---\<open>8 subgoals left\<close>
+\<comment>\<open>8 subgoals left\<close>
prefer 8
apply force
apply force
---\<open>6 subgoals left\<close>
+\<comment>\<open>6 subgoals left\<close>
prefer 6
apply(erule_tac x=j in allE)
apply fastforce
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
prefer 5
apply(case_tac [!] "j=k")
---\<open>10 subgoals left\<close>
+\<comment>\<open>10 subgoals left\<close>
apply simp_all
apply(erule_tac x=k in allE)
apply force
---\<open>9 subgoals left\<close>
+\<comment>\<open>9 subgoals left\<close>
apply(case_tac "j=l")
apply simp
apply(erule_tac x=k in allE)
@@ -211,7 +211,7 @@
apply(erule_tac x=k in allE)
apply(erule_tac x=l in allE)
apply force
---\<open>8 subgoals left\<close>
+\<comment>\<open>8 subgoals left\<close>
apply force
apply(case_tac "j=l")
apply simp
@@ -220,21 +220,21 @@
apply force
apply force
apply force
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
apply(erule_tac x=k in allE)
apply(erule_tac x=l in allE)
apply(case_tac "j=l")
apply force
apply force
apply force
---\<open>3 subgoals left\<close>
+\<comment>\<open>3 subgoals left\<close>
apply(erule_tac x=k in allE)
apply(erule_tac x=l in allE)
apply(case_tac "j=l")
apply force
apply force
apply force
---\<open>1 subgoals left\<close>
+\<comment>\<open>1 subgoals left\<close>
apply(erule_tac x=k in allE)
apply(erule_tac x=l in allE)
apply(case_tac "j=l")
@@ -294,9 +294,9 @@
COEND
\<lbrace>f(\<acute>x)=0 \<or> f(\<acute>y)=0\<rbrace>"
apply oghoare
---\<open>98 verification conditions\<close>
+\<comment>\<open>98 verification conditions\<close>
apply auto
---\<open>auto takes about 3 minutes !!\<close>
+\<comment>\<open>auto takes about 3 minutes !!\<close>
done
text \<open>Easier Version: without AWAIT. Apt and Olderog. page 256:\<close>
@@ -327,9 +327,9 @@
COEND
\<lbrace>f(\<acute>x)=0 \<or> f(\<acute>y)=0\<rbrace>"
apply oghoare
---\<open>20 vc\<close>
+\<comment>\<open>20 vc\<close>
apply auto
---\<open>auto takes aprox. 2 minutes.\<close>
+\<comment>\<open>auto takes aprox. 2 minutes.\<close>
done
subsection \<open>Producer/Consumer\<close>
@@ -429,19 +429,19 @@
COEND
\<lbrace> \<forall>k<length a. (a ! k)=(\<acute>b ! k)\<rbrace>"
apply oghoare
---\<open>138 vc\<close>
+\<comment>\<open>138 vc\<close>
apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
---\<open>112 subgoals left\<close>
+\<comment>\<open>112 subgoals left\<close>
apply(simp_all (no_asm))
---\<open>43 subgoals left\<close>
+\<comment>\<open>43 subgoals left\<close>
apply(tactic \<open>ALLGOALS (conjI_Tac @{context} (K all_tac))\<close>)
---\<open>419 subgoals left\<close>
+\<comment>\<open>419 subgoals left\<close>
apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
---\<open>99 subgoals left\<close>
+\<comment>\<open>99 subgoals left\<close>
apply(simp_all only:length_0_conv [THEN sym])
---\<open>20 subgoals left\<close>
+\<comment>\<open>20 subgoals left\<close>
apply (simp_all del:length_0_conv length_greater_0_conv add: nth_list_update mod_lemma)
---\<open>9 subgoals left\<close>
+\<comment>\<open>9 subgoals left\<close>
apply (force simp add:less_Suc_eq)
apply(hypsubst_thin, drule sym)
apply (force simp add:less_Suc_eq)+
--- a/src/HOL/Hoare_Parallel/OG_Hoare.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy Sat Jan 02 18:48:45 2016 +0100
@@ -119,19 +119,19 @@
apply (unfold com_validity_def)
apply(rule oghoare_induct)
apply simp_all
---\<open>Basic\<close>
+\<comment>\<open>Basic\<close>
apply(simp add: SEM_def sem_def)
apply(fast dest: rtrancl_imp_UN_relpow Basic_ntran)
---\<open>Seq\<close>
+\<comment>\<open>Seq\<close>
apply(rule impI)
apply(rule subset_trans)
prefer 2 apply simp
apply(simp add: L3_5ii L3_5i)
---\<open>Cond\<close>
+\<comment>\<open>Cond\<close>
apply(simp add: L3_5iv)
---\<open>While\<close>
+\<comment>\<open>While\<close>
apply (force simp add: L3_5v dest: SEM_fwhile)
---\<open>Conseq\<close>
+\<comment>\<open>Conseq\<close>
apply(force simp add: SEM_def sem_def)
done
@@ -174,11 +174,11 @@
(\<forall>q. \<turnstile> c q \<longrightarrow> (if co' = None then t\<in>q else t \<in> pre(the co') \<and> \<turnstile> (the co') q )))"
apply(rule ann_transition_transition.induct [THEN conjunct1])
apply simp_all
---\<open>Basic\<close>
+\<comment>\<open>Basic\<close>
apply clarify
apply(frule ann_hoare_case_analysis)
apply force
---\<open>Seq\<close>
+\<comment>\<open>Seq\<close>
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
@@ -189,21 +189,21 @@
apply force
apply(rule AnnSeq,simp)
apply(fast intro: AnnConseq)
---\<open>Cond1\<close>
+\<comment>\<open>Cond1\<close>
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
---\<open>Cond2\<close>
+\<comment>\<open>Cond2\<close>
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
---\<open>While\<close>
+\<comment>\<open>While\<close>
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply force
@@ -214,7 +214,7 @@
apply simp
apply(rule AnnWhile)
apply simp_all
---\<open>Await\<close>
+\<comment>\<open>Await\<close>
apply(frule ann_hoare_case_analysis,simp)
apply clarify
apply(drule atom_hoare_sound)
@@ -348,7 +348,7 @@
prefer 11
apply(rule TrueI)
apply simp_all
---\<open>Basic\<close>
+\<comment>\<open>Basic\<close>
apply(erule_tac x = "i" in all_dupE, erule (1) notE impE)
apply(erule_tac x = "j" in allE , erule (1) notE impE)
apply(simp add: interfree_def)
@@ -365,12 +365,12 @@
apply(force intro: converse_rtrancl_into_rtrancl
simp add: com_validity_def SEM_def sem_def All_None_def)
apply(simp add:assertions_lemma)
---\<open>Seqs\<close>
+\<comment>\<open>Seqs\<close>
apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE)
apply(drule Parallel_Strong_Soundness_Seq,simp+)
apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE)
apply(drule Parallel_Strong_Soundness_Seq,simp+)
---\<open>Await\<close>
+\<comment>\<open>Await\<close>
apply(rule_tac x = "i" in allE , assumption , erule (1) notE impE)
apply(erule_tac x = "j" in allE , erule (1) notE impE)
apply(simp add: interfree_def)
@@ -397,9 +397,9 @@
else t\<in>pre(the(com(Rs ! j))) \<and> \<turnstile> the(com(Rs ! j)) post(Ts ! j))) \<and> interfree Rs"
apply(erule rtrancl_induct2)
apply clarify
---\<open>Base\<close>
+\<comment>\<open>Base\<close>
apply force
---\<open>Induction step\<close>
+\<comment>\<open>Induction step\<close>
apply clarify
apply(drule Parallel_length_post_PStar)
apply clarify
@@ -431,7 +431,7 @@
apply (unfold com_validity_def)
apply(rule oghoare_induct)
apply(rule TrueI)+
---\<open>Parallel\<close>
+\<comment>\<open>Parallel\<close>
apply(simp add: SEM_def sem_def)
apply(clarify, rename_tac x y i Ts')
apply(frule Parallel_length_post_PStar)
@@ -445,19 +445,19 @@
apply(drule_tac s = "length Rs" in sym)
apply(erule allE, erule impE, assumption)
apply(force dest: nth_mem simp add: All_None_def)
---\<open>Basic\<close>
+\<comment>\<open>Basic\<close>
apply(simp add: SEM_def sem_def)
apply(force dest: rtrancl_imp_UN_relpow Basic_ntran)
---\<open>Seq\<close>
+\<comment>\<open>Seq\<close>
apply(rule subset_trans)
prefer 2 apply assumption
apply(simp add: L3_5ii L3_5i)
---\<open>Cond\<close>
+\<comment>\<open>Cond\<close>
apply(simp add: L3_5iv)
---\<open>While\<close>
+\<comment>\<open>While\<close>
apply(simp add: L3_5v)
apply (blast dest: SEM_fwhile)
---\<open>Conseq\<close>
+\<comment>\<open>Conseq\<close>
apply(auto simp add: SEM_def sem_def)
done
--- a/src/HOL/Hoare_Parallel/OG_Syntax.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Syntax.thy Sat Jan 02 18:48:45 2016 +0100
@@ -3,7 +3,7 @@
begin
text\<open>Syntax for commands and for assertions and boolean expressions in
- commands @{text com} and annotated commands @{text ann_com}.\<close>
+ commands \<open>com\<close> and annotated commands \<open>ann_com\<close>.\<close>
abbreviation Skip :: "'a com" ("SKIP" 63)
where "SKIP \<equiv> Basic id"
--- a/src/HOL/Hoare_Parallel/OG_Tactics.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Tactics.thy Sat Jan 02 18:48:45 2016 +0100
@@ -56,10 +56,9 @@
apply assumption+
done
-text \<open>Three new proof rules for special instances of the @{text
-AnnBasic} and the @{text AnnAwait} commands when the transformation
-performed on the state is the identity, and for an @{text AnnAwait}
-command where the boolean condition is @{text "{s. True}"}:\<close>
+text \<open>Three new proof rules for special instances of the \<open>AnnBasic\<close> and the \<open>AnnAwait\<close> commands when the transformation
+performed on the state is the identity, and for an \<open>AnnAwait\<close>
+command where the boolean condition is \<open>{s. True}\<close>:\<close>
lemma AnnatomRule:
"\<lbrakk> atom_com(c); \<parallel>- r c q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnAwait r {s. True} c) q"
@@ -81,9 +80,8 @@
apply simp
done
-text \<open>Lemmata to avoid using the definition of @{text
-map_ann_hoare}, @{text interfree_aux}, @{text interfree_swap} and
-@{text interfree} by splitting it into different cases:\<close>
+text \<open>Lemmata to avoid using the definition of \<open>map_ann_hoare\<close>, \<open>interfree_aux\<close>, \<open>interfree_swap\<close> and
+\<open>interfree\<close> by splitting it into different cases:\<close>
lemma interfree_aux_rule1: "interfree_aux(co, q, None)"
by(simp add:interfree_aux_def)
@@ -286,9 +284,9 @@
addsimps @{thms ParallelConseq_list} @ @{thms my_simp_list})
\<close>
-text \<open>The following tactic applies @{text tac} to each conjunct in a
-subgoal of the form @{text "A \<Longrightarrow> a1 \<and> a2 \<and> .. \<and> an"} returning
-@{text n} subgoals, one for each conjunct:\<close>
+text \<open>The following tactic applies \<open>tac\<close> to each conjunct in a
+subgoal of the form \<open>A \<Longrightarrow> a1 \<and> a2 \<and> .. \<and> an\<close> returning
+\<open>n\<close> subgoals, one for each conjunct:\<close>
ML \<open>
fun conjI_Tac ctxt tac i st = st |>
@@ -307,18 +305,18 @@
\item[HoareRuleTac] is called at the level of parallel programs, it
uses the ParallelTac to solve parallel composition of programs.
This verification has two parts, namely, (1) all component programs are
- correct and (2) they are interference free. @{text HoareRuleTac} is
- also called at the level of atomic regions, i.e. @{text "\<langle> \<rangle>"} and
- @{text "AWAIT b THEN _ END"}, and at each interference freedom test.
+ correct and (2) they are interference free. \<open>HoareRuleTac\<close> is
+ also called at the level of atomic regions, i.e. \<open>\<langle> \<rangle>\<close> and
+ \<open>AWAIT b THEN _ END\<close>, and at each interference freedom test.
\item[AnnHoareRuleTac] is for component programs which
are annotated programs and so, there are not unknown assertions
(no need to use the parameter precond, see NOTE).
- NOTE: precond(::bool) informs if the subgoal has the form @{text "\<parallel>- ?p c q"},
+ NOTE: precond(::bool) informs if the subgoal has the form \<open>\<parallel>- ?p c q\<close>,
in this case we have precond=False and the generated verification
- condition would have the form @{text "?p \<subseteq> \<dots>"} which can be solved by
- @{text "rtac subset_refl"}, if True we proceed to simplify it using
+ condition would have the form \<open>?p \<subseteq> \<dots>\<close> which can be solved by
+ \<open>rtac subset_refl\<close>, if True we proceed to simplify it using
the simplification tactics above.
\end{description}
@@ -448,15 +446,14 @@
\<close>
-text \<open>The final tactic is given the name @{text oghoare}:\<close>
+text \<open>The final tactic is given the name \<open>oghoare\<close>:\<close>
ML \<open>
fun oghoare_tac ctxt = SUBGOAL (fn (_, i) => HoareRuleTac ctxt true i)
\<close>
-text \<open>Notice that the tactic for parallel programs @{text
-"oghoare_tac"} is initially invoked with the value @{text true} for
-the parameter @{text precond}.
+text \<open>Notice that the tactic for parallel programs \<open>oghoare_tac\<close> is initially invoked with the value \<open>true\<close> for
+the parameter \<open>precond\<close>.
Parts of the tactic can be also individually used to generate the
verification conditions for annotated sequential programs and to
--- a/src/HOL/Hoare_Parallel/RG_Com.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Com.thy Sat Jan 02 18:48:45 2016 +0100
@@ -5,8 +5,8 @@
theory RG_Com imports Main begin
text \<open>Semantics of assertions and boolean expressions (bexp) as sets
-of states. Syntax of commands @{text com} and parallel commands
-@{text par_com}.\<close>
+of states. Syntax of commands \<open>com\<close> and parallel commands
+\<open>par_com\<close>.\<close>
type_synonym 'a bexp = "'a set"
--- a/src/HOL/Hoare_Parallel/RG_Examples.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Examples.thy Sat Jan 02 18:48:45 2016 +0100
@@ -269,7 +269,7 @@
\<lbrace>\<forall>i<n. (\<acute>X i) mod n=i \<and> (\<forall>j<\<acute>X i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and>
(\<acute>Y i<m \<longrightarrow> P(B!(\<acute>Y i)) \<and> \<acute>Y i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y j \<le> \<acute>X i)\<rbrace>]"
apply(rule Parallel)
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
apply force+
apply clarify
apply simp
--- a/src/HOL/Hoare_Parallel/RG_Hoare.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy Sat Jan 02 18:48:45 2016 +0100
@@ -458,7 +458,7 @@
apply(subgoal_tac "x\<in> cp (Some(Await b P)) s")
apply(erule_tac i=i in unique_ctran_Await,force,simp_all)
apply(simp add:cp_def)
---\<open>here starts the different part.\<close>
+\<comment>\<open>here starts the different part.\<close>
apply(erule ctran.cases,simp_all)
apply(drule Star_imp_cptn)
apply clarify
@@ -740,7 +740,7 @@
apply (simp del:list.map)
apply(simp only:last_lift_not_None)
apply simp
---\<open>@{text "\<exists>i<length x. fst (x ! i) = Some Q"}\<close>
+\<comment>\<open>\<open>\<exists>i<length x. fst (x ! i) = Some Q\<close>\<close>
apply(erule exE)
apply(drule_tac n=i and P="\<lambda>i. i < length x \<and> fst (x ! i) = Some Q" in Ex_first_occurrence)
apply clarify
@@ -882,13 +882,13 @@
apply(erule cptn_mod.induct)
apply safe
apply (simp_all del:last.simps)
---\<open>5 subgoals left\<close>
+\<comment>\<open>5 subgoals left\<close>
apply(simp add:comm_def)
---\<open>4 subgoals left\<close>
+\<comment>\<open>4 subgoals left\<close>
apply(rule etran_in_comm)
apply(erule mp)
apply(erule tl_of_assum_in_assum,simp)
---\<open>While-None\<close>
+\<comment>\<open>While-None\<close>
apply(ind_cases "((Some (While b P), s), None, t) \<in> ctran" for s t)
apply(simp add:comm_def)
apply(simp add:cptn_iff_cptn_mod [THEN sym])
@@ -913,7 +913,7 @@
apply simp
apply clarify
apply (simp add:last_length)
---\<open>WhileOne\<close>
+\<comment>\<open>WhileOne\<close>
apply(thin_tac "P = While b P \<longrightarrow> Q" for Q)
apply(rule ctran_in_comm,simp)
apply(simp add:Cons_lift del:list.map)
@@ -949,23 +949,23 @@
apply(case_tac "fst(xs!i)")
apply force
apply force
---\<open>last=None\<close>
+\<comment>\<open>last=None\<close>
apply clarify
apply(subgoal_tac "(map (lift (While b P)) ((Some P, sa) # xs))\<noteq>[]")
apply(drule last_conv_nth)
apply (simp del:list.map)
apply(simp only:last_lift_not_None)
apply simp
---\<open>WhileMore\<close>
+\<comment>\<open>WhileMore\<close>
apply(thin_tac "P = While b P \<longrightarrow> Q" for Q)
apply(rule ctran_in_comm,simp del:last.simps)
---\<open>metiendo la hipotesis antes de dividir la conclusion.\<close>
+\<comment>\<open>metiendo la hipotesis antes de dividir la conclusion.\<close>
apply(subgoal_tac "(Some (While b P), snd (last ((Some P, sa) # xs))) # ys \<in> assum (pre, rely)")
apply (simp del:last.simps)
prefer 2
apply(erule assum_after_body)
apply (simp del:last.simps)+
---\<open>lo de antes.\<close>
+\<comment>\<open>lo de antes.\<close>
apply(simp add:comm_def del:list.map last.simps)
apply(rule conjI)
apply clarify
@@ -1001,7 +1001,7 @@
apply(case_tac "fst(xs!i)")
apply force
apply force
---\<open>@{text "i \<ge> length xs"}\<close>
+\<comment>\<open>\<open>i \<ge> length xs\<close>\<close>
apply(subgoal_tac "i-length xs <length ys")
prefer 2
apply arith
@@ -1012,7 +1012,7 @@
apply(erule mp)
apply(case_tac "last((Some P, sa) # xs)")
apply(simp add:lift_def del:last.simps)
---\<open>@{text "i>length xs"}\<close>
+\<comment>\<open>\<open>i>length xs\<close>\<close>
apply(case_tac "i-length xs")
apply arith
apply(simp add:nth_append del:list.map last.simps)
@@ -1021,7 +1021,7 @@
prefer 2
apply arith
apply simp
---\<open>last=None\<close>
+\<comment>\<open>last=None\<close>
apply clarify
apply(case_tac ys)
apply(simp add:Cons_lift del:list.map last.simps)
@@ -1107,16 +1107,16 @@
\<longrightarrow> (snd(clist!i!j), snd(clist!i!Suc j)) \<in> Guar(xs!i)"
apply(unfold par_cp_def)
apply (rule ccontr)
---\<open>By contradiction:\<close>
+\<comment>\<open>By contradiction:\<close>
apply simp
apply(erule exE)
---\<open>the first c-tran that does not satisfy the guarantee-condition is from @{text "\<sigma>_i"} at step @{text "m"}.\<close>
+\<comment>\<open>the first c-tran that does not satisfy the guarantee-condition is from \<open>\<sigma>_i\<close> at step \<open>m\<close>.\<close>
apply(drule_tac n=j and P="\<lambda>j. \<exists>i. H i j" for H in Ex_first_occurrence)
apply(erule exE)
apply clarify
---\<open>@{text "\<sigma>_i \<in> A(pre, rely_1)"}\<close>
+\<comment>\<open>\<open>\<sigma>_i \<in> A(pre, rely_1)\<close>\<close>
apply(subgoal_tac "take (Suc (Suc m)) (clist!i) \<in> assum(Pre(xs!i), Rely(xs!i))")
---\<open>but this contradicts @{text "\<Turnstile> \<sigma>_i sat [pre_i,rely_i,guar_i,post_i]"}\<close>
+\<comment>\<open>but this contradicts \<open>\<Turnstile> \<sigma>_i sat [pre_i,rely_i,guar_i,post_i]\<close>\<close>
apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> \<Turnstile> (J i) sat [I i,K i,M i,N i]" for H J I K M N in allE,erule impE,assumption)
apply(simp add:com_validity_def)
apply(erule_tac x=s in allE)
@@ -1142,9 +1142,9 @@
apply(simp add:conjoin_def compat_label_def)
apply clarify
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (P j) \<or> Q j" for H P Q in allE,simp)
---\<open>each etran in @{text "\<sigma>_1[0\<dots>m]"} corresponds to\<close>
+\<comment>\<open>each etran in \<open>\<sigma>_1[0\<dots>m]\<close> corresponds to\<close>
apply(erule disjE)
---\<open>a c-tran in some @{text "\<sigma>_{ib}"}\<close>
+\<comment>\<open>a c-tran in some \<open>\<sigma>_{ib}\<close>\<close>
apply clarify
apply(case_tac "i=ib",simp)
apply(erule etranE,simp)
@@ -1160,8 +1160,8 @@
apply(simp add:same_state_def)
apply(erule_tac x=i and P="\<lambda>j. (T j) \<longrightarrow> (\<forall>i. (H j i) \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in all_dupE)
apply(erule_tac x=ib and P="\<lambda>j. (T j) \<longrightarrow> (\<forall>i. (H j i) \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in allE,simp)
---\<open>or an e-tran in @{text "\<sigma>"},
-therefore it satisfies @{text "rely \<or> guar_{ib}"}\<close>
+\<comment>\<open>or an e-tran in \<open>\<sigma>\<close>,
+therefore it satisfies \<open>rely \<or> guar_{ib}\<close>\<close>
apply (force simp add:par_assum_def same_state_def)
done
--- a/src/HOL/Hoare_Parallel/RG_Tran.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Tran.thy Sat Jan 02 18:48:45 2016 +0100
@@ -178,20 +178,20 @@
\<longrightarrow> (Some a, s) # (Q, t) # xs \<in> cptn_mod"
apply(induct a)
apply simp_all
---\<open>basic\<close>
+\<comment>\<open>basic\<close>
apply clarify
apply(erule ctran.cases,simp_all)
apply(rule CptnModNone,rule Basic,simp)
apply clarify
apply(erule ctran.cases,simp_all)
---\<open>Seq1\<close>
+\<comment>\<open>Seq1\<close>
apply(rule_tac xs="[(None,ta)]" in CptnModSeq2)
apply(erule CptnModNone)
apply(rule CptnModOne)
apply simp
apply simp
apply(simp add:lift_def)
---\<open>Seq2\<close>
+\<comment>\<open>Seq2\<close>
apply(erule_tac x=sa in allE)
apply(erule_tac x="Some P2" in allE)
apply(erule allE,erule impE, assumption)
@@ -208,12 +208,12 @@
apply (simp add:last_length)
apply (simp add:last_length)
apply(simp add:lift_def)
---\<open>Cond\<close>
+\<comment>\<open>Cond\<close>
apply clarify
apply(erule ctran.cases,simp_all)
apply(force elim: CptnModCondT)
apply(force elim: CptnModCondF)
---\<open>While\<close>
+\<comment>\<open>While\<close>
apply clarify
apply(erule ctran.cases,simp_all)
apply(rule CptnModNone,erule WhileF,simp)
@@ -223,7 +223,7 @@
apply(force elim:CptnModWhile1)
apply clarify
apply(force simp add:last_length elim:CptnModWhile2)
---\<open>await\<close>
+\<comment>\<open>await\<close>
apply clarify
apply(erule ctran.cases,simp_all)
apply(rule CptnModNone,erule Await,simp+)
@@ -295,7 +295,7 @@
apply(erule CondT,simp)
apply(rule CptnComp)
apply(erule CondF,simp)
---\<open>Seq1\<close>
+\<comment>\<open>Seq1\<close>
apply(erule cptn.cases,simp_all)
apply(rule CptnOne)
apply clarify
@@ -315,7 +315,7 @@
apply(rule Seq2,simp)
apply(drule_tac P=P1 in lift_is_cptn)
apply(simp add:lift_def)
---\<open>Seq2\<close>
+\<comment>\<open>Seq2\<close>
apply(rule cptn_append_is_cptn)
apply(drule_tac P=P1 in lift_is_cptn)
apply(simp add:lift_def)
@@ -325,12 +325,12 @@
apply(rule last_fst_esp)
apply (simp add:last_length)
apply(simp add:Cons_lift lift_def split_def last_conv_nth)
---\<open>While1\<close>
+\<comment>\<open>While1\<close>
apply(rule CptnComp)
apply(rule WhileT,simp)
apply(drule_tac P="While b P" in lift_is_cptn)
apply(simp add:lift_def)
---\<open>While2\<close>
+\<comment>\<open>While2\<close>
apply(rule CptnComp)
apply(rule WhileT,simp)
apply(rule cptn_append_is_cptn)
@@ -496,7 +496,7 @@
apply clarify
apply(erule_tac x="0" and P="\<lambda>j. H j \<longrightarrow> (P j \<or> Q j)" for H P Q in all_dupE, simp)
apply(erule disjE)
---\<open>first step is a Component step\<close>
+\<comment>\<open>first step is a Component step\<close>
apply clarify
apply simp
apply(subgoal_tac "a=(xs[i:=(fst(clist!i!0))])")
@@ -516,7 +516,7 @@
apply(erule etranE,simp)
apply(rule ParCptnComp)
apply(erule ParComp,simp)
---\<open>applying the induction hypothesis\<close>
+\<comment>\<open>applying the induction hypothesis\<close>
apply(erule_tac x="xs[i := fst (clist ! i ! 0)]" in allE)
apply(erule_tac x="snd (clist ! i ! 0)" in allE)
apply(erule mp)
@@ -630,7 +630,7 @@
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
apply force
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)
---\<open>first step is an environmental step\<close>
+\<comment>\<open>first step is an environmental step\<close>
apply clarify
apply(erule par_etran.cases)
apply simp
--- a/src/HOL/MicroJava/BV/BVExample.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/BV/BVExample.thy Sat Jan 02 18:48:45 2016 +0100
@@ -152,14 +152,14 @@
text \<open>
The next definition and three proof rules implement an algorithm to
- enumarate natural numbers. The command @{text "apply (elim pc_end pc_next pc_0"}
+ enumarate natural numbers. The command \<open>apply (elim pc_end pc_next pc_0\<close>
transforms a goal of the form
@{prop [display] "pc < n \<Longrightarrow> P pc"}
into a series of goals
@{prop [display] "P 0"}
@{prop [display] "P (Suc 0)"}
- @{text "\<dots>"}
+ \<open>\<dots>\<close>
@{prop [display] "P n"}
\<close>
--- a/src/HOL/MicroJava/BV/BVSpec.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/BV/BVSpec.thy Sat Jan 02 18:48:45 2016 +0100
@@ -16,20 +16,20 @@
\<close>
definition
- -- "The program counter will always be inside the method:"
+ \<comment> "The program counter will always be inside the method:"
check_bounded :: "instr list \<Rightarrow> exception_table \<Rightarrow> bool" where
"check_bounded ins et \<longleftrightarrow>
(\<forall>pc < length ins. \<forall>pc' \<in> set (succs (ins!pc) pc). pc' < length ins) \<and>
(\<forall>e \<in> set et. fst (snd (snd e)) < length ins)"
definition
- -- "The method type only contains declared classes:"
+ \<comment> "The method type only contains declared classes:"
check_types :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state list \<Rightarrow> bool" where
"check_types G mxs mxr phi \<longleftrightarrow> set phi \<subseteq> states G mxs mxr"
definition
- -- "An instruction is welltyped if it is applicable and its effect"
- -- "is compatible with the type at all successor instructions:"
+ \<comment> "An instruction is welltyped if it is applicable and its effect"
+ \<comment> "is compatible with the type at all successor instructions:"
wt_instr :: "[instr,jvm_prog,ty,method_type,nat,p_count,
exception_table,p_count] \<Rightarrow> bool" where
"wt_instr i G rT phi mxs max_pc et pc \<longleftrightarrow>
@@ -37,16 +37,16 @@
(\<forall>(pc',s') \<in> set (eff i G pc et (phi!pc)). pc' < max_pc \<and> G \<turnstile> s' <=' phi!pc')"
definition
- -- \<open>The type at @{text "pc=0"} conforms to the method calling convention:\<close>
+ \<comment> \<open>The type at \<open>pc=0\<close> conforms to the method calling convention:\<close>
wt_start :: "[jvm_prog,cname,ty list,nat,method_type] \<Rightarrow> bool" where
"wt_start G C pTs mxl phi \<longleftrightarrow>
G \<turnstile> Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)) <=' phi!0"
definition
- -- "A method is welltyped if the body is not empty, if execution does not"
- -- "leave the body, if the method type covers all instructions and mentions"
- -- "declared classes only, if the method calling convention is respected, and"
- -- "if all instructions are welltyped."
+ \<comment> "A method is welltyped if the body is not empty, if execution does not"
+ \<comment> "leave the body, if the method type covers all instructions and mentions"
+ \<comment> "declared classes only, if the method calling convention is respected, and"
+ \<comment> "if all instructions are welltyped."
wt_method :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list,
exception_table,method_type] \<Rightarrow> bool" where
"wt_method G C pTs rT mxs mxl ins et phi \<longleftrightarrow>
@@ -59,7 +59,7 @@
(\<forall>pc. pc<max_pc \<longrightarrow> wt_instr (ins!pc) G rT phi mxs max_pc et pc))"
definition
- -- "A program is welltyped if it is wellformed and all methods are welltyped"
+ \<comment> "A program is welltyped if it is wellformed and all methods are welltyped"
wt_jvm_prog :: "[jvm_prog,prog_type] \<Rightarrow> bool" where
"wt_jvm_prog G phi \<longleftrightarrow>
wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b,et)).
--- a/src/HOL/MicroJava/BV/BVSpecTypeSafe.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/BV/BVSpecTypeSafe.thy Sat Jan 02 18:48:45 2016 +0100
@@ -54,8 +54,8 @@
text \<open>
- Relates @{text match_any} from the Bytecode Verifier with
- @{text match_exception_table} from the operational semantics:
+ Relates \<open>match_any\<close> from the Bytecode Verifier with
+ \<open>match_exception_table\<close> from the operational semantics:
\<close>
lemma in_match_any:
"match_exception_table G xcpt pc et = Some pc' \<Longrightarrow>
@@ -125,7 +125,7 @@
text \<open>
We can prove separately that the recursive search for exception
- handlers (@{text find_handler}) in the frame stack results in
+ handlers (\<open>find_handler\<close>) in the frame stack results in
a conforming state (if there was no matching exception handler
in the current frame). We require that the exception is a valid
heap address, and that the state before the exception occured
@@ -137,22 +137,22 @@
\<Longrightarrow> G,phi \<turnstile>JVM (find_handler G (Some xcp) hp frs)\<surd>"
(is "\<And>f. \<lbrakk> ?wt; ?adr; ?hp; ?correct (None, hp, f#frs) \<rbrakk> \<Longrightarrow> ?correct (?find frs)")
proof (induct frs)
- -- "the base case is trivial, as it should be"
+ \<comment> "the base case is trivial, as it should be"
show "?correct (?find [])" by (simp add: correct_state_def)
- -- "we will need both forms @{text wt_jvm_prog} and @{text wf_prog} later"
+ \<comment> "we will need both forms \<open>wt_jvm_prog\<close> and \<open>wf_prog\<close> later"
assume wt: ?wt
then obtain mb where wf: "wf_prog mb G" by (simp add: wt_jvm_prog_def)
- -- "these two don't change in the induction:"
+ \<comment> "these two don't change in the induction:"
assume adr: ?adr
assume hp: ?hp
- -- "the assumption for the cons case:"
+ \<comment> "the assumption for the cons case:"
fix f f' frs'
assume cr: "?correct (None, hp, f#f'#frs')"
- -- "the induction hypothesis as produced by Isabelle, immediatly simplified
+ \<comment> "the induction hypothesis as produced by Isabelle, immediatly simplified
with the fixed assumptions above"
assume "\<And>f. \<lbrakk> ?wt; ?adr; ?hp; ?correct (None, hp, f#frs') \<rbrakk> \<Longrightarrow> ?correct (?find frs')"
with wt adr hp
@@ -245,7 +245,7 @@
declare raise_if_def [simp]
text \<open>
- The requirement of lemma @{text uncaught_xcpt_correct} (that
+ The requirement of lemma \<open>uncaught_xcpt_correct\<close> (that
the exception is a valid reference on the heap) is always met
for welltyped instructions and conformant states:
\<close>
@@ -356,7 +356,7 @@
phi_pc': "phi C sig ! handler = Some (ST', LT')" and
frame': "correct_frame G hp (ST',LT') maxl ins ?f'"
proof (cases "ins!pc")
- case Return -- "can't generate exceptions:"
+ case Return \<comment> "can't generate exceptions:"
with xp' have False by (simp add: split_beta split: split_if_asm)
thus ?thesis ..
next
@@ -571,7 +571,7 @@
}
ultimately
show ?thesis by (rule that)
- qed (insert xp', auto) -- "the other instructions don't generate exceptions"
+ qed (insert xp', auto) \<comment> "the other instructions don't generate exceptions"
from state' meth hp_ok "class" frames phi_pc' frame' prehp
have ?thesis by (unfold correct_state_def) simp
--- a/src/HOL/MicroJava/BV/Effect.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/BV/Effect.thy Sat Jan 02 18:48:45 2016 +0100
@@ -52,9 +52,9 @@
= (PrimT Integer#ST,LT)" |
"eff' (Ifcmpeq b, G, (ts1#ts2#ST,LT)) = (ST,LT)" |
"eff' (Goto b, G, s) = s" |
- -- "Return has no successor instruction in the same method"
+ \<comment> "Return has no successor instruction in the same method"
"eff' (Return, G, s) = s" |
- -- "Throw always terminates abruptly"
+ \<comment> "Throw always terminates abruptly"
"eff' (Throw, G, s) = s" |
"eff' (Invoke C mn fpTs, G, (ST,LT)) = (let ST' = drop (length fpTs) ST
in (fst (snd (the (method (G,C) (mn,fpTs))))#(tl ST'),LT))"
--- a/src/HOL/MicroJava/BV/JVMType.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/BV/JVMType.thy Sat Jan 02 18:48:45 2016 +0100
@@ -12,8 +12,8 @@
type_synonym locvars_type = "ty err list"
type_synonym opstack_type = "ty list"
type_synonym state_type = "opstack_type \<times> locvars_type"
-type_synonym state = "state_type option err" -- "for Kildall"
-type_synonym method_type = "state_type option list" -- "for BVSpec"
+type_synonym state = "state_type option err" \<comment> "for Kildall"
+type_synonym method_type = "state_type option list" \<comment> "for BVSpec"
type_synonym class_type = "sig \<Rightarrow> method_type"
type_synonym prog_type = "cname \<Rightarrow> class_type"
--- a/src/HOL/MicroJava/Comp/CorrCompTp.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy Sat Jan 02 18:48:45 2016 +0100
@@ -1269,19 +1269,19 @@
apply (intro strip)
apply (rule conjI)
- -- "app"
+ \<comment> "app"
apply (rule Call_app [THEN app_mono_mxs])
apply assumption+
apply (rule HOL.refl)
apply assumption
apply (simp add: max_ssize_def max_of_list_elem ssize_sto_def)
- -- \<open>@{text "<=s"}\<close>
+ \<comment> \<open>\<open><=s\<close>\<close>
apply (frule max_spec2mheads, (erule exE)+, (erule conjE)+)
apply (simp add: wf_prog_ws_prog [THEN comp_method])
apply (simp add: max_spec_preserves_length [symmetric])
- -- "@{text check_type}"
+ \<comment> "\<open>check_type\<close>"
apply (simp add: max_ssize_def ssize_sto_def)
apply (simp add: max_of_list_def)
apply (subgoal_tac "(max (length pTsa + length ST) (length ST)) = (length pTsa + length ST)")
--- a/src/HOL/MicroJava/Comp/LemmasComp.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/Comp/LemmasComp.thy Sat Jan 02 18:48:45 2016 +0100
@@ -349,7 +349,7 @@
apply (simp add: map_of_map2)
apply (simp (no_asm_simp) add: compMethod_def split_beta)
- -- "remaining subgoals"
+ \<comment> "remaining subgoals"
apply (auto intro: inv_f_eq simp add: inj_on_def is_class_def)
done
--- a/src/HOL/MicroJava/DFA/Kildall.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/DFA/Kildall.thy Sat Jan 02 18:48:45 2016 +0100
@@ -348,10 +348,10 @@
r = "{(ss',ss) . ss <[r] ss'} <*lex*> finite_psubset"
in while_rule)
--- "Invariant holds initially:"
+\<comment> "Invariant holds initially:"
apply (simp add:stables_def)
--- "Invariant is preserved:"
+\<comment> "Invariant is preserved:"
apply(simp add: stables_def split_paired_all)
apply(rename_tac ss w)
apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
@@ -393,16 +393,16 @@
apply (blast dest!: boundedD)
--- "Postcondition holds upon termination:"
+\<comment> "Postcondition holds upon termination:"
apply(clarsimp simp add: stables_def split_paired_all)
--- "Well-foundedness of the termination relation:"
+\<comment> "Well-foundedness of the termination relation:"
apply (rule wf_lex_prod)
apply (insert orderI [THEN acc_le_listI])
apply (simp add: acc_def lesssub_def wfP_wf_eq [symmetric])
apply (rule wf_finite_psubset)
--- "Loop decreases along termination relation:"
+\<comment> "Loop decreases along termination relation:"
apply(simp add: stables_def split_paired_all)
apply(rename_tac ss w)
apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
--- a/src/HOL/MicroJava/J/Conform.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/Conform.thy Sat Jan 02 18:48:45 2016 +0100
@@ -7,7 +7,7 @@
theory Conform imports State WellType Exceptions begin
-type_synonym 'c env' = "'c prog \<times> (vname \<rightharpoonup> ty)" -- "same as @{text env} of @{text WellType.thy}"
+type_synonym 'c env' = "'c prog \<times> (vname \<rightharpoonup> ty)" \<comment> "same as \<open>env\<close> of \<open>WellType.thy\<close>"
definition hext :: "aheap => aheap => bool" ("_ \<le>| _" [51,51] 50) where
"h\<le>|h' == \<forall>a C fs. h a = Some(C,fs) --> (\<exists>fs'. h' a = Some(C,fs'))"
--- a/src/HOL/MicroJava/J/Decl.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/Decl.thy Sat Jan 02 18:48:45 2016 +0100
@@ -8,23 +8,23 @@
theory Decl imports Type begin
type_synonym
- fdecl = "vname \<times> ty" -- "field declaration, cf. 8.3 (, 9.3)"
+ fdecl = "vname \<times> ty" \<comment> "field declaration, cf. 8.3 (, 9.3)"
type_synonym
- sig = "mname \<times> ty list" -- "signature of a method, cf. 8.4.2"
+ sig = "mname \<times> ty list" \<comment> "signature of a method, cf. 8.4.2"
type_synonym
- 'c mdecl = "sig \<times> ty \<times> 'c" -- "method declaration in a class"
+ 'c mdecl = "sig \<times> ty \<times> 'c" \<comment> "method declaration in a class"
type_synonym
'c "class" = "cname \<times> fdecl list \<times> 'c mdecl list"
- -- "class = superclass, fields, methods"
+ \<comment> "class = superclass, fields, methods"
type_synonym
- 'c cdecl = "cname \<times> 'c class" -- "class declaration, cf. 8.1"
+ 'c cdecl = "cname \<times> 'c class" \<comment> "class declaration, cf. 8.1"
type_synonym
- 'c prog = "'c cdecl list" -- "program"
+ 'c prog = "'c cdecl list" \<comment> "program"
translations
--- a/src/HOL/MicroJava/J/Eval.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/Eval.thy Sat Jan 02 18:48:45 2016 +0100
@@ -8,7 +8,7 @@
theory Eval imports State WellType begin
- -- "Auxiliary notions"
+ \<comment> "Auxiliary notions"
definition fits :: "java_mb prog \<Rightarrow> state \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60) where
"G,s\<turnstile>a' fits T \<equiv> case T of PrimT T' \<Rightarrow> False | RefT T' \<Rightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T"
@@ -23,7 +23,7 @@
"new_xcpt_var vn \<equiv> \<lambda>(x,s). Norm (lupd(vn\<mapsto>the x) s)"
- -- "Evaluation relations"
+ \<comment> "Evaluation relations"
inductive
eval :: "[java_mb prog,xstate,expr,val,xstate] => bool "
@@ -36,21 +36,21 @@
for G :: "java_mb prog"
where
- -- "evaluation of expressions"
+ \<comment> "evaluation of expressions"
- XcptE:"G\<turnstile>(Some xc,s) -e\<succ>undefined-> (Some xc,s)" -- "cf. 15.5"
+ XcptE:"G\<turnstile>(Some xc,s) -e\<succ>undefined-> (Some xc,s)" \<comment> "cf. 15.5"
- -- "cf. 15.8.1"
+ \<comment> "cf. 15.8.1"
| NewC: "[| h = heap s; (a,x) = new_Addr h;
h'= h(a\<mapsto>(C,init_vars (fields (G,C)))) |] ==>
G\<turnstile>Norm s -NewC C\<succ>Addr a-> c_hupd h' (x,s)"
- -- "cf. 15.15"
+ \<comment> "cf. 15.15"
| Cast: "[| G\<turnstile>Norm s0 -e\<succ>v-> (x1,s1);
x2 = raise_if (\<not> cast_ok G C (heap s1) v) ClassCast x1 |] ==>
G\<turnstile>Norm s0 -Cast C e\<succ>v-> (x2,s1)"
- -- "cf. 15.7.1"
+ \<comment> "cf. 15.7.1"
| Lit: "G\<turnstile>Norm s -Lit v\<succ>v-> Norm s"
| BinOp:"[| G\<turnstile>Norm s -e1\<succ>v1-> s1;
@@ -59,27 +59,27 @@
| Add => Intg (the_Intg v1 + the_Intg v2)) |] ==>
G\<turnstile>Norm s -BinOp bop e1 e2\<succ>v-> s2"
- -- "cf. 15.13.1, 15.2"
+ \<comment> "cf. 15.13.1, 15.2"
| LAcc: "G\<turnstile>Norm s -LAcc v\<succ>the (locals s v)-> Norm s"
- -- "cf. 15.25.1"
+ \<comment> "cf. 15.25.1"
| LAss: "[| G\<turnstile>Norm s -e\<succ>v-> (x,(h,l));
l' = (if x = None then l(va\<mapsto>v) else l) |] ==>
G\<turnstile>Norm s -va::=e\<succ>v-> (x,(h,l'))"
- -- "cf. 15.10.1, 15.2"
+ \<comment> "cf. 15.10.1, 15.2"
| FAcc: "[| G\<turnstile>Norm s0 -e\<succ>a'-> (x1,s1);
v = the (snd (the (heap s1 (the_Addr a'))) (fn,T)) |] ==>
G\<turnstile>Norm s0 -{T}e..fn\<succ>v-> (np a' x1,s1)"
- -- "cf. 15.25.1"
+ \<comment> "cf. 15.25.1"
| FAss: "[| G\<turnstile> Norm s0 -e1\<succ>a'-> (x1,s1); a = the_Addr a';
G\<turnstile>(np a' x1,s1) -e2\<succ>v -> (x2,s2);
h = heap s2; (c,fs) = the (h a);
h' = h(a\<mapsto>(c,(fs((fn,T)\<mapsto>v)))) |] ==>
G\<turnstile>Norm s0 -{T}e1..fn:=e2\<succ>v-> c_hupd h' (x2,s2)"
- -- "cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15"
+ \<comment> "cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15"
| Call: "[| G\<turnstile>Norm s0 -e\<succ>a'-> s1; a = the_Addr a';
G\<turnstile>s1 -ps[\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a));
(md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs));
@@ -88,43 +88,43 @@
G\<turnstile>Norm s0 -{C}e..mn({pTs}ps)\<succ>v-> (x4,(heap s4,l))"
- -- "evaluation of expression lists"
+ \<comment> "evaluation of expression lists"
- -- "cf. 15.5"
+ \<comment> "cf. 15.5"
| XcptEs:"G\<turnstile>(Some xc,s) -e[\<succ>]undefined-> (Some xc,s)"
- -- "cf. 15.11.???"
+ \<comment> "cf. 15.11.???"
| Nil: "G\<turnstile>Norm s0 -[][\<succ>][]-> Norm s0"
- -- "cf. 15.6.4"
+ \<comment> "cf. 15.6.4"
| Cons: "[| G\<turnstile>Norm s0 -e \<succ> v -> s1;
G\<turnstile> s1 -es[\<succ>]vs-> s2 |] ==>
G\<turnstile>Norm s0 -e#es[\<succ>]v#vs-> s2"
- -- "execution of statements"
+ \<comment> "execution of statements"
- -- "cf. 14.1"
+ \<comment> "cf. 14.1"
| XcptS:"G\<turnstile>(Some xc,s) -c-> (Some xc,s)"
- -- "cf. 14.5"
+ \<comment> "cf. 14.5"
| Skip: "G\<turnstile>Norm s -Skip-> Norm s"
- -- "cf. 14.7"
+ \<comment> "cf. 14.7"
| Expr: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1 |] ==>
G\<turnstile>Norm s0 -Expr e-> s1"
- -- "cf. 14.2"
+ \<comment> "cf. 14.2"
| Comp: "[| G\<turnstile>Norm s0 -c1-> s1;
G\<turnstile> s1 -c2-> s2|] ==>
G\<turnstile>Norm s0 -c1;; c2-> s2"
- -- "cf. 14.8.2"
+ \<comment> "cf. 14.8.2"
| Cond: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1;
G\<turnstile> s1 -(if the_Bool v then c1 else c2)-> s2|] ==>
G\<turnstile>Norm s0 -If(e) c1 Else c2-> s2"
- -- "cf. 14.10, 14.10.1"
+ \<comment> "cf. 14.10, 14.10.1"
| LoopF:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; \<not>the_Bool v |] ==>
G\<turnstile>Norm s0 -While(e) c-> s1"
| LoopT:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; the_Bool v;
--- a/src/HOL/MicroJava/J/Example.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/Example.thy Sat Jan 02 18:48:45 2016 +0100
@@ -40,8 +40,8 @@
datatype vnam' = vee' | x' | e'
text \<open>
- @{text cnam'} and @{text vnam'} are intended to be isomorphic
- to @{text cnam} and @{text vnam}
+ \<open>cnam'\<close> and \<open>vnam'\<close> are intended to be isomorphic
+ to \<open>cnam\<close> and \<open>vnam\<close>
\<close>
axiomatization cnam' :: "cnam' => cname"
@@ -133,7 +133,7 @@
lemma map_of_Cons2 [simp]: "aa\<noteq>k ==> map_of ((k,bb)#ps) aa = map_of ps aa"
apply (simp (no_asm_simp))
done
-declare map_of_Cons [simp del] -- "sic!"
+declare map_of_Cons [simp del] \<comment> "sic!"
lemma class_tprg_Object [simp]: "class tprg Object = Some (undefined, [], [])"
apply (unfold ObjectC_def class_def)
@@ -375,25 +375,25 @@
lemmas t = ty_expr_ty_exprs_wt_stmt.intros
schematic_goal wt_test: "(tprg, empty(e\<mapsto>Class Base))\<turnstile>
Expr(e::=NewC Ext);; Expr({Base}LAcc e..foo({?pTs'}[Lit Null]))\<surd>"
-apply (rule ty_expr_ty_exprs_wt_stmt.intros) -- ";;"
-apply (rule t) -- "Expr"
-apply (rule t) -- "LAss"
-apply simp -- \<open>@{text "e \<noteq> This"}\<close>
-apply (rule t) -- "LAcc"
+apply (rule ty_expr_ty_exprs_wt_stmt.intros) \<comment> ";;"
+apply (rule t) \<comment> "Expr"
+apply (rule t) \<comment> "LAss"
+apply simp \<comment> \<open>\<open>e \<noteq> This\<close>\<close>
+apply (rule t) \<comment> "LAcc"
apply (simp (no_asm))
apply (simp (no_asm))
-apply (rule t) -- "NewC"
+apply (rule t) \<comment> "NewC"
apply (simp (no_asm))
apply (simp (no_asm))
-apply (rule t) -- "Expr"
-apply (rule t) -- "Call"
-apply (rule t) -- "LAcc"
+apply (rule t) \<comment> "Expr"
+apply (rule t) \<comment> "Call"
+apply (rule t) \<comment> "LAcc"
apply (simp (no_asm))
apply (simp (no_asm))
-apply (rule t) -- "Cons"
-apply (rule t) -- "Lit"
+apply (rule t) \<comment> "Cons"
+apply (rule t) \<comment> "Lit"
apply (simp (no_asm))
-apply (rule t) -- "Nil"
+apply (rule t) \<comment> "Nil"
apply (simp (no_asm))
apply (rule max_spec_foo_Base)
done
@@ -406,38 +406,38 @@
" [|new_Addr (heap (snd s0)) = (a, None)|] ==>
tprg\<turnstile>s0 -test-> ?s"
apply (unfold test_def)
--- "?s = s3 "
-apply (rule e) -- ";;"
-apply (rule e) -- "Expr"
-apply (rule e) -- "LAss"
-apply (rule e) -- "NewC"
+\<comment> "?s = s3 "
+apply (rule e) \<comment> ";;"
+apply (rule e) \<comment> "Expr"
+apply (rule e) \<comment> "LAss"
+apply (rule e) \<comment> "NewC"
apply force
apply force
apply (simp (no_asm))
apply (erule thin_rl)
-apply (rule e) -- "Expr"
-apply (rule e) -- "Call"
-apply (rule e) -- "LAcc"
+apply (rule e) \<comment> "Expr"
+apply (rule e) \<comment> "Call"
+apply (rule e) \<comment> "LAcc"
apply force
-apply (rule e) -- "Cons"
-apply (rule e) -- "Lit"
-apply (rule e) -- "Nil"
+apply (rule e) \<comment> "Cons"
+apply (rule e) \<comment> "Lit"
+apply (rule e) \<comment> "Nil"
apply (simp (no_asm))
apply (force simp add: foo_Ext_def)
apply (simp (no_asm))
-apply (rule e) -- "Expr"
-apply (rule e) -- "FAss"
-apply (rule e) -- "Cast"
-apply (rule e) -- "LAcc"
+apply (rule e) \<comment> "Expr"
+apply (rule e) \<comment> "FAss"
+apply (rule e) \<comment> "Cast"
+apply (rule e) \<comment> "LAcc"
apply (simp (no_asm))
apply (simp (no_asm))
apply (simp (no_asm))
-apply (rule e) -- "XcptE"
+apply (rule e) \<comment> "XcptE"
apply (simp (no_asm))
apply (rule surjective_pairing [symmetric, THEN[2]trans], subst prod.inject, force)
apply (simp (no_asm))
apply (simp (no_asm))
-apply (rule e) -- "XcptE"
+apply (rule e) \<comment> "XcptE"
done
end
--- a/src/HOL/MicroJava/J/JTypeSafe.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/JTypeSafe.thy Sat Jan 02 18:48:45 2016 +0100
@@ -198,15 +198,15 @@
apply( rule eval_evals_exec_induct)
apply( unfold c_hupd_def)
--- "several simplifications, XcptE, XcptEs, XcptS, Skip, Nil??"
+\<comment> "several simplifications, XcptE, XcptEs, XcptS, Skip, Nil??"
apply( simp_all)
apply( tactic "ALLGOALS (REPEAT o resolve_tac @{context} [impI, allI])")
apply( tactic \<open>ALLGOALS (eresolve_tac @{context} [@{thm ty_expr.cases}, @{thm ty_exprs.cases}, @{thm wt_stmt.cases}]
THEN_ALL_NEW (full_simp_tac (put_simpset (simpset_of @{theory_context Conform}) @{context})))\<close>)
apply(tactic "ALLGOALS (EVERY' [REPEAT o (eresolve_tac @{context} [conjE]), REPEAT o hyp_subst_tac @{context}])")
--- "Level 7"
--- "15 NewC"
+\<comment> "Level 7"
+\<comment> "15 NewC"
apply (drule sym)
apply( drule new_AddrD)
apply( erule disjE)
@@ -221,13 +221,13 @@
apply( rule_tac [2] rtrancl.rtrancl_refl)
apply( simp (no_asm))
--- "for Cast"
+\<comment> "for Cast"
defer 1
--- "14 Lit"
+\<comment> "14 Lit"
apply( erule conf_litval)
--- "13 BinOp"
+\<comment> "13 BinOp"
apply (tactic "forward_hyp_tac @{context}")
apply (tactic "forward_hyp_tac @{context}")
apply( rule conjI, erule (1) hext_trans)
@@ -236,34 +236,34 @@
apply( drule eval_no_xcpt)
apply( simp split add: binop.split)
--- "12 LAcc"
+\<comment> "12 LAcc"
apply simp
apply( fast elim: conforms_localD [THEN lconfD])
--- "for FAss"
+\<comment> "for FAss"
apply( tactic \<open>EVERY'[eresolve_tac @{context} [@{thm ty_expr.cases}, @{thm ty_exprs.cases}, @{thm wt_stmt.cases}]
THEN_ALL_NEW (full_simp_tac @{context}), REPEAT o (eresolve_tac @{context} [conjE]), hyp_subst_tac @{context}] 3\<close>)
--- "for if"
+\<comment> "for if"
apply( tactic \<open>(Induct_Tacs.case_tac @{context} "the_Bool v" [] NONE THEN_ALL_NEW
(asm_full_simp_tac @{context})) 7\<close>)
apply (tactic "forward_hyp_tac @{context}")
--- "11+1 if"
+\<comment> "11+1 if"
prefer 7
apply( fast intro: hext_trans)
prefer 7
apply( fast intro: hext_trans)
--- "10 Expr"
+\<comment> "10 Expr"
prefer 6
apply( fast)
--- "9 ???"
+\<comment> "9 ???"
apply( simp_all)
--- "8 Cast"
+\<comment> "8 Cast"
prefer 8
apply (rule conjI)
apply (fast intro: conforms_xcpt_change xconf_raise_if)
@@ -275,7 +275,7 @@
apply assumption+
--- "7 LAss"
+\<comment> "7 LAss"
apply (fold fun_upd_def)
apply( tactic \<open>(eresolve_tac @{context} [@{thm ty_expr.cases}, @{thm ty_exprs.cases}, @{thm wt_stmt.cases}]
THEN_ALL_NEW (full_simp_tac @{context})) 1\<close>)
@@ -284,13 +284,13 @@
apply (simp)
apply( blast intro: conforms_upd_local conf_widen)
--- "6 FAcc"
+\<comment> "6 FAcc"
apply (rule conjI)
apply (simp add: np_def)
apply (fast intro: conforms_xcpt_change xconf_raise_if)
apply( fast elim!: FAcc_type_sound)
--- "5 While"
+\<comment> "5 While"
prefer 5
apply(erule_tac V = "a \<longrightarrow> b" for a b in thin_rl)
apply(drule (1) ty_expr_ty_exprs_wt_stmt.Loop)
@@ -298,7 +298,7 @@
apply (tactic "forward_hyp_tac @{context}")
--- "4 Cond"
+\<comment> "4 Cond"
prefer 4
apply (case_tac "the_Bool v")
apply simp
@@ -306,31 +306,31 @@
apply simp
apply( fast dest: evals_no_xcpt intro: conf_hext hext_trans)
--- "3 ;;"
+\<comment> "3 ;;"
prefer 3
apply( fast dest: evals_no_xcpt intro: conf_hext hext_trans)
--- "2 FAss"
+\<comment> "2 FAss"
apply (subgoal_tac "(np a' x1, aa, ba) ::\<preceq> (G, lT)")
prefer 2
apply (simp add: np_def)
apply (fast intro: conforms_xcpt_change xconf_raise_if)
apply( case_tac "x2")
- -- "x2 = None"
+ \<comment> "x2 = None"
apply (simp)
apply (tactic "forward_hyp_tac @{context}", clarify)
apply( drule eval_no_xcpt)
apply( erule FAss_type_sound, rule HOL.refl, assumption+)
- -- "x2 = Some a"
+ \<comment> "x2 = Some a"
apply ( simp (no_asm_simp))
apply( fast intro: hext_trans)
apply( tactic "prune_params_tac @{context}")
--- "Level 52"
+\<comment> "Level 52"
--- "1 Call"
+\<comment> "1 Call"
apply( case_tac "x")
prefer 2
apply( clarsimp)
--- a/src/HOL/MicroJava/J/State.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/State.thy Sat Jan 02 18:48:45 2016 +0100
@@ -10,10 +10,10 @@
begin
type_synonym
- fields' = "(vname \<times> cname \<rightharpoonup> val)" -- "field name, defining class, value"
+ fields' = "(vname \<times> cname \<rightharpoonup> val)" \<comment> "field name, defining class, value"
type_synonym
- obj = "cname \<times> fields'" -- "class instance with class name and fields"
+ obj = "cname \<times> fields'" \<comment> "class instance with class name and fields"
definition obj_ty :: "obj => ty" where
"obj_ty obj == Class (fst obj)"
@@ -21,11 +21,11 @@
definition init_vars :: "('a \<times> ty) list => ('a \<rightharpoonup> val)" where
"init_vars == map_of o map (\<lambda>(n,T). (n,default_val T))"
-type_synonym aheap = "loc \<rightharpoonup> obj" -- \<open>"@{text heap}" used in a translation below\<close>
-type_synonym locals = "vname \<rightharpoonup> val" -- "simple state, i.e. variable contents"
+type_synonym aheap = "loc \<rightharpoonup> obj" \<comment> \<open>"\<open>heap\<close>" used in a translation below\<close>
+type_synonym locals = "vname \<rightharpoonup> val" \<comment> "simple state, i.e. variable contents"
-type_synonym state = "aheap \<times> locals" -- "heap, local parameter including This"
-type_synonym xstate = "val option \<times> state" -- "state including exception information"
+type_synonym state = "aheap \<times> locals" \<comment> "heap, local parameter including This"
+type_synonym xstate = "val option \<times> state" \<comment> "state including exception information"
abbreviation (input)
heap :: "state => aheap"
@@ -52,7 +52,7 @@
definition raise_if :: "bool \<Rightarrow> xcpt \<Rightarrow> val option \<Rightarrow> val option" where
"raise_if b x xo \<equiv> if b \<and> (xo = None) then Some (Addr (XcptRef x)) else xo"
-text \<open>Make @{text new_Addr} completely specified (at least for the code generator)\<close>
+text \<open>Make \<open>new_Addr\<close> completely specified (at least for the code generator)\<close>
(*
definition new_Addr :: "aheap => loc \<times> val option" where
"new_Addr h \<equiv> SOME (a,x). (h a = None \<and> x = None) | x = Some (Addr (XcptRef OutOfMemory))"
--- a/src/HOL/MicroJava/J/SystemClasses.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/SystemClasses.thy Sat Jan 02 18:48:45 2016 +0100
@@ -8,7 +8,7 @@
theory SystemClasses imports Decl begin
text \<open>
- This theory provides definitions for the @{text Object} class,
+ This theory provides definitions for the \<open>Object\<close> class,
and the system exceptions.
\<close>
--- a/src/HOL/MicroJava/J/Term.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/Term.thy Sat Jan 02 18:48:45 2016 +0100
@@ -6,26 +6,26 @@
theory Term imports Value begin
-datatype binop = Eq | Add -- "function codes for binary operation"
+datatype binop = Eq | Add \<comment> "function codes for binary operation"
datatype expr
- = NewC cname -- "class instance creation"
- | Cast cname expr -- "type cast"
- | Lit val -- "literal value, also references"
- | BinOp binop expr expr -- "binary operation"
- | LAcc vname -- "local (incl. parameter) access"
- | LAss vname expr ("_::=_" [90,90]90) -- "local assign"
- | FAcc cname expr vname ("{_}_.._" [10,90,99]90) -- "field access"
+ = NewC cname \<comment> "class instance creation"
+ | Cast cname expr \<comment> "type cast"
+ | Lit val \<comment> "literal value, also references"
+ | BinOp binop expr expr \<comment> "binary operation"
+ | LAcc vname \<comment> "local (incl. parameter) access"
+ | LAss vname expr ("_::=_" [90,90]90) \<comment> "local assign"
+ | FAcc cname expr vname ("{_}_.._" [10,90,99]90) \<comment> "field access"
| FAss cname expr vname
- expr ("{_}_.._:=_" [10,90,99,90]90) -- "field ass."
+ expr ("{_}_.._:=_" [10,90,99,90]90) \<comment> "field ass."
| Call cname expr mname
- "ty list" "expr list" ("{_}_.._'( {_}_')" [10,90,99,10,10] 90) -- "method call"
+ "ty list" "expr list" ("{_}_.._'( {_}_')" [10,90,99,10,10] 90) \<comment> "method call"
datatype_compat expr
datatype stmt
- = Skip -- "empty statement"
- | Expr expr -- "expression statement"
+ = Skip \<comment> "empty statement"
+ | Expr expr \<comment> "expression statement"
| Comp stmt stmt ("_;; _" [61,60]60)
| Cond expr stmt stmt ("If '(_') _ Else _" [80,79,79]70)
| Loop expr stmt ("While '(_') _" [80,79]70)
--- a/src/HOL/MicroJava/J/Type.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/Type.thy Sat Jan 02 18:48:45 2016 +0100
@@ -17,7 +17,7 @@
end
-text \<open>These instantiations only ensure that the merge in theory @{text "MicroJava"} succeeds. FIXME\<close>
+text \<open>These instantiations only ensure that the merge in theory \<open>MicroJava\<close> succeeds. FIXME\<close>
instantiation cnam :: typerep
begin
@@ -44,20 +44,20 @@
end
- -- "exceptions"
+ \<comment> "exceptions"
datatype
xcpt
= NullPointer
| ClassCast
| OutOfMemory
--- "class names"
+\<comment> "class names"
datatype cname
= Object
| Xcpt xcpt
| Cname cnam
-typedecl vnam -- "variable or field name"
+typedecl vnam \<comment> "variable or field name"
instantiation vnam :: equal
begin
@@ -92,7 +92,7 @@
end
-typedecl mname -- "method name"
+typedecl mname \<comment> "method name"
instantiation mname :: equal
begin
@@ -127,26 +127,26 @@
end
--- "names for @{text This} pointer and local/field variables"
+\<comment> "names for \<open>This\<close> pointer and local/field variables"
datatype vname
= This
| VName vnam
--- "primitive type, cf. 4.2"
+\<comment> "primitive type, cf. 4.2"
datatype prim_ty
- = Void -- "'result type' of void methods"
+ = Void \<comment> "'result type' of void methods"
| Boolean
| Integer
--- "reference type, cf. 4.3"
+\<comment> "reference type, cf. 4.3"
datatype ref_ty
- = NullT -- "null type, cf. 4.1"
- | ClassT cname -- "class type"
+ = NullT \<comment> "null type, cf. 4.1"
+ | ClassT cname \<comment> "class type"
--- "any type, cf. 4.1"
+\<comment> "any type, cf. 4.1"
datatype ty
- = PrimT prim_ty -- "primitive type"
- | RefT ref_ty -- "reference type"
+ = PrimT prim_ty \<comment> "primitive type"
+ | RefT ref_ty \<comment> "reference type"
abbreviation NT :: ty
where "NT == RefT NullT"
--- a/src/HOL/MicroJava/J/TypeRel.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/TypeRel.thy Sat Jan 02 18:48:45 2016 +0100
@@ -8,7 +8,7 @@
imports Decl
begin
--- "direct subclass, cf. 8.1.3"
+\<comment> "direct subclass, cf. 8.1.3"
inductive_set
subcls1 :: "'c prog => (cname \<times> cname) set"
@@ -181,7 +181,7 @@
field :: "'c prog \<times> cname => ( vname \<rightharpoonup> cname \<times> ty )" (* ###curry *)
fields :: "'c prog \<times> cname => ((vname \<times> cname) \<times> ty) list" (* ###curry *)
--- "methods of a class, with inheritance, overriding and hiding, cf. 8.4.6"
+\<comment> "methods of a class, with inheritance, overriding and hiding, cf. 8.4.6"
defs method_def [code]: "method \<equiv> \<lambda>(G,C). class_rec G C empty (\<lambda>C fs ms ts.
ts ++ map_of (map (\<lambda>(s,m). (s,(C,m))) ms))"
@@ -195,7 +195,7 @@
done
--- "list of fields of a class, including inherited and hidden ones"
+\<comment> "list of fields of a class, including inherited and hidden ones"
defs fields_def [code]: "fields \<equiv> \<lambda>(G,C). class_rec G C [] (\<lambda>C fs ms ts.
map (\<lambda>(fn,ft). ((fn,C),ft)) fs @ ts)"
@@ -219,12 +219,12 @@
done
--- "widening, viz. method invocation conversion,cf. 5.3 i.e. sort of syntactic subtyping"
+\<comment> "widening, viz. method invocation conversion,cf. 5.3 i.e. sort of syntactic subtyping"
inductive
widen :: "'c prog => [ty , ty ] => bool" ("_ \<turnstile> _ \<preceq> _" [71,71,71] 70)
for G :: "'c prog"
where
- refl [intro!, simp]: "G\<turnstile> T \<preceq> T" -- "identity conv., cf. 5.1.1"
+ refl [intro!, simp]: "G\<turnstile> T \<preceq> T" \<comment> "identity conv., cf. 5.1.1"
| subcls : "G\<turnstile>C\<preceq>C D ==> G\<turnstile>Class C \<preceq> Class D"
| null [intro!]: "G\<turnstile> NT \<preceq> RefT R"
@@ -232,8 +232,8 @@
lemmas refl = HOL.refl
--- "casting conversion, cf. 5.5 / 5.1.5"
--- "left out casts on primitve types"
+\<comment> "casting conversion, cf. 5.5 / 5.1.5"
+\<comment> "left out casts on primitve types"
inductive
cast :: "'c prog => [ty , ty ] => bool" ("_ \<turnstile> _ \<preceq>? _" [71,71,71] 70)
for G :: "'c prog"
--- a/src/HOL/MicroJava/J/Value.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/Value.thy Sat Jan 02 18:48:45 2016 +0100
@@ -6,19 +6,19 @@
theory Value imports Type begin
-typedecl loc' -- "locations, i.e. abstract references on objects"
+typedecl loc' \<comment> "locations, i.e. abstract references on objects"
datatype loc
- = XcptRef xcpt -- "special locations for pre-allocated system exceptions"
- | Loc loc' -- "usual locations (references on objects)"
+ = XcptRef xcpt \<comment> "special locations for pre-allocated system exceptions"
+ | Loc loc' \<comment> "usual locations (references on objects)"
datatype val
- = Unit -- "dummy result value of void methods"
- | Null -- "null reference"
- | Bool bool -- "Boolean value"
- | Intg int -- "integer value, name Intg instead of Int because
+ = Unit \<comment> "dummy result value of void methods"
+ | Null \<comment> "null reference"
+ | Bool bool \<comment> "Boolean value"
+ | Intg int \<comment> "integer value, name Intg instead of Int because
of clash with HOL/Set.thy"
- | Addr loc -- "addresses, i.e. locations of objects "
+ | Addr loc \<comment> "addresses, i.e. locations of objects "
primrec the_Bool :: "val => bool" where
"the_Bool (Bool b) = b"
@@ -29,12 +29,12 @@
primrec the_Addr :: "val => loc" where
"the_Addr (Addr a) = a"
-primrec defpval :: "prim_ty => val" -- "default value for primitive types" where
+primrec defpval :: "prim_ty => val" \<comment> "default value for primitive types" where
"defpval Void = Unit"
| "defpval Boolean = Bool False"
| "defpval Integer = Intg 0"
-primrec default_val :: "ty => val" -- "default value for all types" where
+primrec default_val :: "ty => val" \<comment> "default value for all types" where
"default_val (PrimT pt) = defpval pt"
| "default_val (RefT r ) = Null"
--- a/src/HOL/MicroJava/J/WellType.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/J/WellType.thy Sat Jan 02 18:48:45 2016 +0100
@@ -43,13 +43,13 @@
more_spec_def: "more_spec G == \<lambda>((d,h),pTs). \<lambda>((d',h'),pTs'). G\<turnstile>d\<preceq>d' \<and>
list_all2 (\<lambda>T T'. G\<turnstile>T\<preceq>T') pTs pTs'"
- -- "applicable methods, cf. 15.11.2.1"
+ \<comment> "applicable methods, cf. 15.11.2.1"
appl_methds_def: "appl_methds G C == \<lambda>(mn, pTs).
{((Class md,rT),pTs') |md rT mb pTs'.
method (G,C) (mn, pTs') = Some (md,rT,mb) \<and>
list_all2 (\<lambda>T T'. G\<turnstile>T\<preceq>T') pTs pTs'}"
- -- "maximally specific methods, cf. 15.11.2.2"
+ \<comment> "maximally specific methods, cf. 15.11.2.2"
max_spec_def: "max_spec G C sig == {m. m \<in>appl_methds G C sig \<and>
(\<forall>m'\<in>appl_methds G C sig.
more_spec G m' m --> m' = m)}"
@@ -100,8 +100,8 @@
type_synonym
java_mb = "vname list \<times> (vname \<times> ty) list \<times> stmt \<times> expr"
--- "method body with parameter names, local variables, block, result expression."
--- "local variables might include This, which is hidden anyway"
+\<comment> "method body with parameter names, local variables, block, result expression."
+\<comment> "local variables might include This, which is hidden anyway"
inductive
ty_expr :: "'c env => expr => ty => bool" ("_ \<turnstile> _ :: _" [51, 51, 51] 50)
@@ -110,19 +110,19 @@
where
NewC: "[| is_class (prg E) C |] ==>
- E\<turnstile>NewC C::Class C" -- "cf. 15.8"
+ E\<turnstile>NewC C::Class C" \<comment> "cf. 15.8"
- -- "cf. 15.15"
+ \<comment> "cf. 15.15"
| Cast: "[| E\<turnstile>e::C; is_class (prg E) D;
prg E\<turnstile>C\<preceq>? Class D |] ==>
E\<turnstile>Cast D e:: Class D"
- -- "cf. 15.7.1"
+ \<comment> "cf. 15.7.1"
| Lit: "[| typeof (\<lambda>v. None) x = Some T |] ==>
E\<turnstile>Lit x::T"
- -- "cf. 15.13.1"
+ \<comment> "cf. 15.13.1"
| LAcc: "[| localT E v = Some T; is_type (prg E) T |] ==>
E\<turnstile>LAcc v::T"
@@ -132,42 +132,42 @@
else T' = T \<and> T = PrimT Integer|] ==>
E\<turnstile>BinOp bop e1 e2::T'"
- -- "cf. 15.25, 15.25.1"
+ \<comment> "cf. 15.25, 15.25.1"
| LAss: "[| v ~= This;
E\<turnstile>LAcc v::T;
E\<turnstile>e::T';
prg E\<turnstile>T'\<preceq>T |] ==>
E\<turnstile>v::=e::T'"
- -- "cf. 15.10.1"
+ \<comment> "cf. 15.10.1"
| FAcc: "[| E\<turnstile>a::Class C;
field (prg E,C) fn = Some (fd,fT) |] ==>
E\<turnstile>{fd}a..fn::fT"
- -- "cf. 15.25, 15.25.1"
+ \<comment> "cf. 15.25, 15.25.1"
| FAss: "[| E\<turnstile>{fd}a..fn::T;
E\<turnstile>v ::T';
prg E\<turnstile>T'\<preceq>T |] ==>
E\<turnstile>{fd}a..fn:=v::T'"
- -- "cf. 15.11.1, 15.11.2, 15.11.3"
+ \<comment> "cf. 15.11.1, 15.11.2, 15.11.3"
| Call: "[| E\<turnstile>a::Class C;
E\<turnstile>ps[::]pTs;
max_spec (prg E) C (mn, pTs) = {((md,rT),pTs')} |] ==>
E\<turnstile>{C}a..mn({pTs'}ps)::rT"
--- "well-typed expression lists"
+\<comment> "well-typed expression lists"
- -- "cf. 15.11.???"
+ \<comment> "cf. 15.11.???"
| Nil: "E\<turnstile>[][::][]"
- -- "cf. 15.11.???"
+ \<comment> "cf. 15.11.???"
| Cons:"[| E\<turnstile>e::T;
E\<turnstile>es[::]Ts |] ==>
E\<turnstile>e#es[::]T#Ts"
--- "well-typed statements"
+\<comment> "well-typed statements"
| Skip:"E\<turnstile>Skip\<surd>"
@@ -178,13 +178,13 @@
E\<turnstile>s2\<surd> |] ==>
E\<turnstile>s1;; s2\<surd>"
- -- "cf. 14.8"
+ \<comment> "cf. 14.8"
| Cond:"[| E\<turnstile>e::PrimT Boolean;
E\<turnstile>s1\<surd>;
E\<turnstile>s2\<surd> |] ==>
E\<turnstile>If(e) s1 Else s2\<surd>"
- -- "cf. 14.10"
+ \<comment> "cf. 14.10"
| Loop:"[| E\<turnstile>e::PrimT Boolean;
E\<turnstile>s\<surd> |] ==>
E\<turnstile>While(e) s\<surd>"
--- a/src/HOL/MicroJava/JVM/JVMExec.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/JVM/JVMExec.thy Sat Jan 02 18:48:45 2016 +0100
@@ -10,7 +10,7 @@
fun
exec :: "jvm_prog \<times> jvm_state => jvm_state option"
--- "exec is not recursive. fun is just used for pattern matching"
+\<comment> "exec is not recursive. fun is just used for pattern matching"
where
"exec (G, xp, hp, []) = None"
@@ -30,8 +30,8 @@
text \<open>
The start configuration of the JVM: in the start heap, we call a
- method @{text m} of class @{text C} in program @{text G}. The
- @{text this} pointer of the frame is set to @{text Null} to simulate
+ method \<open>m\<close> of class \<open>C\<close> in program \<open>G\<close>. The
+ \<open>this\<close> pointer of the frame is set to \<open>Null\<close> to simulate
a static method invokation.
\<close>
definition start_state :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> jvm_state" where
--- a/src/HOL/MicroJava/JVM/JVMExecInstr.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/JVM/JVMExecInstr.thy Sat Jan 02 18:48:45 2016 +0100
@@ -67,8 +67,8 @@
else []
in
(xp', hp, frs'@(stk, vars, Cl, sig, pc)#frs))" |
- -- "Because exception handling needs the pc of the Invoke instruction,"
- -- "Invoke doesn't change stk and pc yet (@{text Return} does that)."
+ \<comment> "Because exception handling needs the pc of the Invoke instruction,"
+ \<comment> "Invoke doesn't change stk and pc yet (\<open>Return\<close> does that)."
"exec_instr Return G hp stk0 vars Cl sig0 pc frs =
(if frs=[] then
@@ -78,8 +78,8 @@
(mn,pt) = sig0; n = length pt
in
(None, hp, (val#(drop (n+1) stk),loc,C,sig,pc+1)#tl frs))"
- -- "Return drops arguments from the caller's stack and increases"
- -- "the program counter in the caller" |
+ \<comment> "Return drops arguments from the caller's stack and increases"
+ \<comment> "the program counter in the caller" |
"exec_instr Pop G hp stk vars Cl sig pc frs =
(None, hp, (tl stk, vars, Cl, sig, pc+1)#frs)" |
--- a/src/HOL/MicroJava/JVM/JVMInstructions.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/JVM/JVMInstructions.thy Sat Jan 02 18:48:45 2016 +0100
@@ -9,35 +9,35 @@
datatype
- instr = Load nat -- "load from local variable"
- | Store nat -- "store into local variable"
- | LitPush val -- "push a literal (constant)"
- | New cname -- "create object"
- | Getfield vname cname -- "Fetch field from object"
- | Putfield vname cname -- "Set field in object "
- | Checkcast cname -- "Check whether object is of given type"
- | Invoke cname mname "(ty list)" -- "inv. instance meth of an object"
- | Return -- "return from method"
- | Pop -- "pop top element from opstack"
- | Dup -- "duplicate top element of opstack"
- | Dup_x1 -- "duplicate top element and push 2 down"
- | Dup_x2 -- "duplicate top element and push 3 down"
- | Swap -- "swap top and next to top element"
- | IAdd -- "integer addition"
- | Goto int -- "goto relative address"
- | Ifcmpeq int -- "branch if int/ref comparison succeeds"
- | Throw -- "throw top of stack as exception"
+ instr = Load nat \<comment> "load from local variable"
+ | Store nat \<comment> "store into local variable"
+ | LitPush val \<comment> "push a literal (constant)"
+ | New cname \<comment> "create object"
+ | Getfield vname cname \<comment> "Fetch field from object"
+ | Putfield vname cname \<comment> "Set field in object "
+ | Checkcast cname \<comment> "Check whether object is of given type"
+ | Invoke cname mname "(ty list)" \<comment> "inv. instance meth of an object"
+ | Return \<comment> "return from method"
+ | Pop \<comment> "pop top element from opstack"
+ | Dup \<comment> "duplicate top element of opstack"
+ | Dup_x1 \<comment> "duplicate top element and push 2 down"
+ | Dup_x2 \<comment> "duplicate top element and push 3 down"
+ | Swap \<comment> "swap top and next to top element"
+ | IAdd \<comment> "integer addition"
+ | Goto int \<comment> "goto relative address"
+ | Ifcmpeq int \<comment> "branch if int/ref comparison succeeds"
+ | Throw \<comment> "throw top of stack as exception"
type_synonym
bytecode = "instr list"
type_synonym
exception_entry = "p_count \<times> p_count \<times> p_count \<times> cname"
- -- "start-pc, end-pc, handler-pc, exception type"
+ \<comment> "start-pc, end-pc, handler-pc, exception type"
type_synonym
exception_table = "exception_entry list"
type_synonym
jvm_method = "nat \<times> nat \<times> bytecode \<times> exception_table"
- -- "max stacksize, size of register set, instruction sequence, handler table"
+ \<comment> "max stacksize, size of register set, instruction sequence, handler table"
type_synonym
jvm_prog = "jvm_method prog"
--- a/src/HOL/MicroJava/JVM/JVMListExample.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/JVM/JVMListExample.thy Sat Jan 02 18:48:45 2016 +0100
@@ -9,7 +9,7 @@
begin
text \<open>
- Since the types @{typ cnam}, @{text vnam}, and @{text mname} are
+ Since the types @{typ cnam}, \<open>vnam\<close>, and \<open>mname\<close> are
anonymous, we describe distinctness of names in the example by axioms:
\<close>
axiomatization list_nam test_nam :: cnam
--- a/src/HOL/MicroJava/JVM/JVMState.thy Sat Jan 02 18:46:36 2016 +0100
+++ b/src/HOL/MicroJava/JVM/JVMState.thy Sat Jan 02 18:48:45 2016 +0100
@@ -22,11 +22,11 @@
sig \<times>
p_count"
- -- "operand stack"
- -- "local variables (including this pointer and method parameters)"
- -- "name of class where current method is defined"
- -- "method name + parameter types"
- -- "program counter within frame"
+ \<comment> "operand stack"
+ \<comment> "local variables (including this pointer and method parameters)"
+ \<comment> "name of class where current method is defined"
+ \<comment> "method name + parameter types"
+ \<comment> "program counter within frame"
subsection \<open>Exceptions\<close>
@@ -35,7 +35,7 @@
subsection \<open>Runtime State\<close>
type_synonym
- jvm_state = "val option \<times> aheap \<times> frame list" -- "exception flag, heap, frames"
+ jvm_state = "val option \<times> aheap \<times> frame list" \<comment> "exception flag, heap, frames"
subsection \<open>Lemmas\<close>